refactor: Lint

This commit is contained in:
jstoobysmith 2024-10-15 11:39:40 +00:00
parent 89d1b1a50b
commit f72d69e2ba
6 changed files with 84 additions and 69 deletions

View file

@ -13,7 +13,7 @@ import HepLean.SpaceTime.WeylFermion.Two
We define the metrics for Weyl fermions, often denoted `ε` in the literature.
These allow us to go from left-handed to alt-left-handed Weyl fermions and back,
and from right-handed to alt-right-handed Weyl fermions and back.
and from right-handed to alt-right-handed Weyl fermions and back.
-/
@ -49,7 +49,7 @@ lemma metricRaw_comm (M : SL(2,)) : metricRaw * M.1 = (M.1⁻¹)ᵀ * metricR
mul_one, smul_empty, tail_cons, neg_smul, mul_neg, neg_cons, neg_neg, neg_zero, neg_empty,
empty_vecMul, add_cons, empty_add_empty, empty_mul, Equiv.symm_apply_apply]
lemma star_comm_metricRaw (M : SL(2,)) : M.1.map star * metricRaw = metricRaw * ((M.1)⁻¹)ᴴ := by
lemma star_comm_metricRaw (M : SL(2,)) : M.1.map star * metricRaw = metricRaw * ((M.1)⁻¹)ᴴ := by
rw [metricRaw]
rw [SpaceTime.SL2C.inverse_coe, eta_fin_two M.1]
rw [SpecialLinearGroup.coe_inv, Matrix.adjugate_fin_two,
@ -83,11 +83,13 @@ def leftMetric : 𝟙_ (Rep SL(2,)) ⟶ leftHanded ⊗ leftHanded where
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • leftMetricVal =
(TensorProduct.map (leftHanded.ρ M) (leftHanded.ρ M)) (x' • leftMetricVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [leftMetricVal]
erw [leftLeftToMatrix_ρ_symm]
@ -114,11 +116,13 @@ def altLeftMetric : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ altLeftHande
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • altLeftMetricVal =
(TensorProduct.map (altLeftHanded.ρ M) (altLeftHanded.ρ M)) (x' • altLeftMetricVal)
simp
simp only [Action.instMonoidalCategory_tensorObj_V, _root_.map_smul]
apply congrArg
simp [altLeftMetricVal]
erw [altLeftaltLeftToMatrix_ρ_symm]
@ -127,7 +131,6 @@ def altLeftMetric : 𝟙_ (Rep SL(2,)) ⟶ altLeftHanded ⊗ altLeftHande
simp only [SpecialLinearGroup.det_coe, isUnit_iff_ne_zero, ne_eq, one_ne_zero,
not_false_eq_true, mul_nonsing_inv, mul_one]
/-- The metric `ε_{dot a}_{dot a}` as an element of `(rightHanded ⊗ rightHanded).V`. -/
def rightMetricVal : (rightHanded ⊗ rightHanded).V :=
rightRightToMatrix.symm (- metricRaw)
@ -146,7 +149,9 @@ def rightMetric : 𝟙_ (Rep SL(2,)) ⟶ rightHanded ⊗ rightHanded wher
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • rightMetricVal =
(TensorProduct.map (rightHanded.ρ M) (rightHanded.ρ M)) (x' • rightMetricVal)
@ -186,7 +191,9 @@ def altRightMetric : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHa
rfl}
comm M := by
ext x : 2
simp
simp only [Action.instMonoidalCategory_tensorObj_V, Action.instMonoidalCategory_tensorUnit_V,
Action.tensorUnit_ρ', CategoryTheory.Category.id_comp, Action.tensor_ρ', ModuleCat.coe_comp,
Function.comp_apply]
let x' : := x
change x' • altRightMetricVal =
(TensorProduct.map (altRightHanded.ρ M) (altRightHanded.ρ M)) (x' • altRightMetricVal)
@ -201,7 +208,7 @@ def altRightMetric : 𝟙_ (Rep SL(2,)) ⟶ altRightHanded ⊗ altRightHa
have h1 : ((M.1).map star * (M.1)⁻¹ᴴᵀ) = 1 := by
refine transpose_eq_one.mp ?_
rw [@transpose_mul]
simp
simp only [transpose_transpose, RCLike.star_def]
change (M.1)⁻¹ᴴ * (M.1)ᴴ = 1
rw [← @conjTranspose_mul]
simp