feat: Add contr of basis

This commit is contained in:
jstoobysmith 2024-10-23 08:01:23 +00:00
parent 74d4b2c2c0
commit 865164ca81
6 changed files with 173 additions and 6 deletions

View file

@ -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