feat: Add contr of basis
This commit is contained in:
parent
74d4b2c2c0
commit
865164ca81
6 changed files with 173 additions and 6 deletions
|
@ -58,6 +58,12 @@ lemma leftBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
|||
change (M.1 *ᵥ (Pi.single j 1)) i = _
|
||||
simp only [mulVec_single, mul_one]
|
||||
|
||||
@[simp]
|
||||
lemma leftBasis_toFin2ℂ (i : Fin 2) : (leftBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [leftBasis, Basis.coe_ofEquivFun]
|
||||
rw [LeftHandedModule.toFin2ℂ]
|
||||
rfl
|
||||
|
||||
/-- The vector space ℂ^2 carrying the representation of SL(2,C) given by
|
||||
M → (M⁻¹)ᵀ. In index notation corresponds to a Weyl fermion with indices ψ_a. -/
|
||||
def altLeftHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
||||
|
@ -85,6 +91,12 @@ def altLeftHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
|||
def altLeftBasis : Basis (Fin 2) ℂ altLeftHanded := Basis.ofEquivFun
|
||||
(Equiv.linearEquiv ℂ AltLeftHandedModule.toFin2ℂFun)
|
||||
|
||||
@[simp]
|
||||
lemma altLeftBasis_toFin2ℂ (i : Fin 2) : (altLeftBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [altLeftBasis, Basis.coe_ofEquivFun]
|
||||
rw [AltLeftHandedModule.toFin2ℂ]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma altLeftBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
||||
(LinearMap.toMatrix altLeftBasis altLeftBasis) (altLeftHanded.ρ M) i j = (M.1⁻¹)ᵀ i j := by
|
||||
|
@ -117,6 +129,12 @@ def rightHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
|||
def rightBasis : Basis (Fin 2) ℂ rightHanded := Basis.ofEquivFun
|
||||
(Equiv.linearEquiv ℂ RightHandedModule.toFin2ℂFun)
|
||||
|
||||
@[simp]
|
||||
lemma rightBasis_toFin2ℂ (i : Fin 2) : (rightBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [rightBasis, Basis.coe_ofEquivFun]
|
||||
rw [RightHandedModule.toFin2ℂ]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma rightBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
||||
(LinearMap.toMatrix rightBasis rightBasis) (rightHanded.ρ M) i j = (M.1.map star) i j := by
|
||||
|
@ -153,6 +171,12 @@ def altRightHanded : Rep ℂ SL(2,ℂ) := Rep.of {
|
|||
def altRightBasis : Basis (Fin 2) ℂ altRightHanded := Basis.ofEquivFun
|
||||
(Equiv.linearEquiv ℂ AltRightHandedModule.toFin2ℂFun)
|
||||
|
||||
@[simp]
|
||||
lemma altRightBasis_toFin2ℂ (i : Fin 2) : (altRightBasis i).toFin2ℂ = Pi.single i 1 := by
|
||||
simp only [altRightBasis, Basis.coe_ofEquivFun]
|
||||
rw [AltRightHandedModule.toFin2ℂ]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma altRightBasis_ρ_apply (M : SL(2,ℂ)) (i j : Fin 2) :
|
||||
(LinearMap.toMatrix altRightBasis altRightBasis) (altRightHanded.ρ M) i j =
|
||||
|
|
|
@ -137,6 +137,16 @@ lemma leftAltContraction_hom_tmul (ψ : leftHanded) (φ : altLeftHanded) :
|
|||
leftAltContraction.hom (ψ ⊗ₜ φ) = ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma leftAltContraction_basis (i j : Fin 2) :
|
||||
leftAltContraction.hom (leftBasis i ⊗ₜ altLeftBasis j) = if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [leftAltContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, leftBasis_toFin2ℂ, altLeftBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/-- The linear map from altLeftHandedWeyl ⊗ leftHandedWeyl to ℂ given by
|
||||
summing over components of altLeftHandedWeyl and leftHandedWeyl in the
|
||||
standard basis (i.e. the dot product).
|
||||
|
@ -153,6 +163,16 @@ lemma altLeftContraction_hom_tmul (φ : altLeftHanded) (ψ : leftHanded) :
|
|||
altLeftContraction.hom (φ ⊗ₜ ψ) = φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma altLeftContraction_basis (i j : Fin 2) :
|
||||
altLeftContraction.hom (altLeftBasis i ⊗ₜ leftBasis j) = if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [altLeftContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, leftBasis_toFin2ℂ, altLeftBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/--
|
||||
The linear map from rightHandedWeyl ⊗ altRightHandedWeyl to ℂ given by
|
||||
summing over components of rightHandedWeyl and altRightHandedWeyl in the
|
||||
|
@ -182,6 +202,16 @@ lemma rightAltContraction_hom_tmul (ψ : rightHanded) (φ : altRightHanded) :
|
|||
rightAltContraction.hom (ψ ⊗ₜ φ) = ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma rightAltContraction_basis (i j : Fin 2) :
|
||||
rightAltContraction.hom (rightBasis i ⊗ₜ altRightBasis j) = if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [rightAltContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, rightBasis_toFin2ℂ, altRightBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/--
|
||||
The linear map from altRightHandedWeyl ⊗ rightHandedWeyl to ℂ given by
|
||||
summing over components of altRightHandedWeyl and rightHandedWeyl in the
|
||||
|
@ -211,6 +241,16 @@ lemma altRightContraction_hom_tmul (φ : altRightHanded) (ψ : rightHanded) :
|
|||
altRightContraction.hom (φ ⊗ₜ ψ) = φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ := by
|
||||
rfl
|
||||
|
||||
lemma altRightContraction_basis (i j : Fin 2) :
|
||||
altRightContraction.hom (altRightBasis i ⊗ₜ rightBasis j) = if i.1 = j.1 then (1 : ℂ) else 0 := by
|
||||
rw [altRightContraction_hom_tmul]
|
||||
simp only [Action.instMonoidalCategory_tensorUnit_V, rightBasis_toFin2ℂ, altRightBasis_toFin2ℂ,
|
||||
dotProduct_single, mul_one]
|
||||
rw [Pi.single_apply]
|
||||
simp only [Fin.ext_iff]
|
||||
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
|
||||
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
|
||||
|
||||
/-!
|
||||
|
||||
## Symmetry properties
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue