refactor: Lint
This commit is contained in:
parent
89d1b1a50b
commit
f72d69e2ba
6 changed files with 84 additions and 69 deletions
|
@ -21,7 +21,6 @@ open MatrixGroups
|
|||
open Complex
|
||||
open TensorProduct
|
||||
|
||||
|
||||
/-!
|
||||
|
||||
## Contraction of Weyl fermions.
|
||||
|
@ -76,7 +75,6 @@ def altLeftBi : altLeftHanded →ₗ[ℂ] leftHanded →ₗ[ℂ] ℂ where
|
|||
simp only [_root_.map_smul, smul_dotProduct, vec2_dotProduct, Fin.isValue, smul_eq_mul,
|
||||
LinearMap.coe_mk, AddHom.coe_mk, RingHom.id_apply, LinearMap.smul_apply]
|
||||
|
||||
|
||||
/-- The bi-linear map corresponding to contraction of a right-handed Weyl fermion with a
|
||||
alt-right-handed Weyl fermion. -/
|
||||
def rightAltBi : rightHanded →ₗ[ℂ] altRightHanded →ₗ[ℂ] ℂ where
|
||||
|
@ -170,7 +168,8 @@ The linear map from rightHandedWeyl ⊗ altRightHandedWeyl to ℂ given by
|
|||
def rightAltContraction : rightHanded ⊗ altRightHanded ⟶ 𝟙_ (Rep ℂ SL(2,ℂ)) where
|
||||
hom := TensorProduct.lift rightAltBi
|
||||
comm M := TensorProduct.ext' fun ψ φ => by
|
||||
change (M.1.map star *ᵥ ψ.toFin2ℂ) ⬝ᵥ (M.1⁻¹.conjTranspose *ᵥ φ.toFin2ℂ) = ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ
|
||||
change (M.1.map star *ᵥ ψ.toFin2ℂ) ⬝ᵥ (M.1⁻¹.conjTranspose *ᵥ φ.toFin2ℂ) =
|
||||
ψ.toFin2ℂ ⬝ᵥ φ.toFin2ℂ
|
||||
have h1 : (M.1)⁻¹ᴴ = ((M.1)⁻¹.map star)ᵀ := by rfl
|
||||
rw [dotProduct_mulVec, h1, vecMul_transpose, mulVec_mulVec]
|
||||
have h2 : ((M.1)⁻¹.map star * (M.1).map star) = 1 := by
|
||||
|
@ -193,8 +192,9 @@ def rightAltContraction : rightHanded ⊗ altRightHanded ⟶ 𝟙_ (Rep ℂ SL(2
|
|||
-/
|
||||
def altRightContraction : altRightHanded ⊗ rightHanded ⟶ 𝟙_ (Rep ℂ SL(2,ℂ)) where
|
||||
hom := TensorProduct.lift altRightBi
|
||||
comm M := TensorProduct.ext' fun φ ψ => by
|
||||
change (M.1⁻¹.conjTranspose *ᵥ φ.toFin2ℂ) ⬝ᵥ (M.1.map star *ᵥ ψ.toFin2ℂ) = φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ
|
||||
comm M := TensorProduct.ext' fun φ ψ => by
|
||||
change (M.1⁻¹.conjTranspose *ᵥ φ.toFin2ℂ) ⬝ᵥ (M.1.map star *ᵥ ψ.toFin2ℂ) =
|
||||
φ.toFin2ℂ ⬝ᵥ ψ.toFin2ℂ
|
||||
have h1 : (M.1)⁻¹ᴴ = ((M.1)⁻¹.map star)ᵀ := by rfl
|
||||
rw [dotProduct_mulVec, h1, mulVec_transpose, vecMul_vecMul]
|
||||
have h2 : ((M.1)⁻¹.map star * (M.1).map star) = 1 := by
|
||||
|
@ -251,7 +251,6 @@ informal_lemma altLeftWeylContraction_invariant where
|
|||
the action of SL(2,C) on leftHandedWeyl and altLeftHandedWeyl."
|
||||
deps :≈ [``altLeftContraction]
|
||||
|
||||
|
||||
informal_lemma rightAltWeylContraction_invariant where
|
||||
math :≈ "The contraction rightAltWeylContraction is invariant with respect to
|
||||
the action of SL(2,C) on rightHandedWeyl and altRightHandedWeyl."
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue