PhysLean/HepLean/SpaceTime/WeylFermion/Unit.lean
2024-10-15 11:29:18 +00:00

159 lines
4.9 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.WeylFermion.Basic
import HepLean.SpaceTime.WeylFermion.Contraction
import Mathlib.LinearAlgebra.TensorProduct.Matrix
import HepLean.SpaceTime.WeylFermion.Two
/-!
# Units of Weyl fermions
We define the units for Weyl fermions, often denoted `δ` in the literature.
-/
namespace Fermion
noncomputable section
open Matrix
open MatrixGroups
open Complex
open TensorProduct
open CategoryTheory.MonoidalCategory
/-- The left-alt-left unit `δₐᵃ` as an element of `(leftHanded ⊗ altLeftHanded).V`. -/
def leftAltLeftUnitVal : (leftHanded ⊗ altLeftHanded).V :=
leftAltLeftToMatrix.symm 1
/-- The left-alt-left unit `δₐᵃ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded `,
manifesting the invariance under the `SL(2,)` action. -/
def leftAltLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ altLeftHanded where
hom := {
toFun := fun a =>
let a' : := a
a' • leftAltLeftUnitVal,
map_add' := fun x y => by
simp only [add_smul]
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp
let x' : := x
change x' • leftAltLeftUnitVal =
(TensorProduct.map (leftHanded.ρ M) (altLeftHanded.ρ M)) (x' • leftAltLeftUnitVal)
simp
apply congrArg
simp [leftAltLeftUnitVal]
erw [leftAltLeftToMatrix_ρ_symm]
apply congrArg
simp
/-- The alt-left-left unit `δᵃₐ` as an element of `(altLeftHanded ⊗ leftHanded).V`. -/
def altLeftLeftUnitVal : (altLeftHanded ⊗ leftHanded).V :=
altLeftLeftToMatrix.symm 1
/-- The alt-left-left unit `δᵃₐ` as a morphism `𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded `,
manifesting the invariance under the `SL(2,)` action. -/
def altLeftLeftUnit : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ leftHanded where
hom := {
toFun := fun a =>
let a' : := a
a' • altLeftLeftUnitVal,
map_add' := fun x y => by
simp only [add_smul]
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp
let x' : := x
change x' • altLeftLeftUnitVal =
(TensorProduct.map (altLeftHanded.ρ M) (leftHanded.ρ M)) (x' • altLeftLeftUnitVal)
simp
apply congrArg
simp [altLeftLeftUnitVal]
erw [altLeftLeftToMatrix_ρ_symm]
apply congrArg
simp only [mul_one, ← transpose_mul, SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq,
one_ne_zero, not_false_eq_true, mul_nonsing_inv, transpose_one]
/-- The right-alt-right unit `δ_{dot a}^{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
def rightAltRightUnitVal : (rightHanded ⊗ altRightHanded).V :=
rightAltRightToMatrix.symm 1
/-- The right-alt-right unit `δ_{dot a}^{dot a}` as a morphism
`𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHanded`, manifesting
the invariance under the `SL(2,)` action. -/
def rightAltRightUnit : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ altRightHanded where
hom := {
toFun := fun a =>
let a' : := a
a' • rightAltRightUnitVal,
map_add' := fun x y => by
simp only [add_smul]
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp
let x' : := x
change x' • rightAltRightUnitVal =
(TensorProduct.map (rightHanded.ρ M) (altRightHanded.ρ M)) (x' • rightAltRightUnitVal)
simp
apply congrArg
simp [rightAltRightUnitVal]
erw [rightAltRightToMatrix_ρ_symm]
apply congrArg
simp
symm
refine transpose_eq_one.mp ?h.h.h.a
simp
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [@conjTranspose_nonsing_inv]
simp
/-- The alt-right-right unit `δ^{dot a}_{dot a}` as an element of
`(rightHanded ⊗ altRightHanded).V`. -/
def altRightRightUnitVal : (altRightHanded ⊗ rightHanded).V :=
altRightRightToMatrix.symm 1
/-- The alt-right-right unit `δ^{dot a}_{dot a}` as a morphism
`𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHanded`, manifesting
the invariance under the `SL(2,)` action. -/
def altRightRightUnit : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ rightHanded where
hom := {
toFun := fun a =>
let a' : := a
a' • altRightRightUnitVal,
map_add' := fun x y => by
simp only [add_smul]
map_smul' := fun m x => by
simp only [smul_smul]
rfl}
comm M := by
ext x : 2
simp
let x' : := x
change x' • altRightRightUnitVal =
(TensorProduct.map (altRightHanded.ρ M) (rightHanded.ρ M)) (x' • altRightRightUnitVal)
simp
apply congrArg
simp [altRightRightUnitVal]
erw [altRightRightToMatrix_ρ_symm]
apply congrArg
simp
symm
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [@conjTranspose_nonsing_inv]
simp
end
end Fermion