refactor: Lint

This commit is contained in:
jstoobysmith 2024-11-11 07:22:36 +00:00
parent a8243f4e79
commit 081955c993
6 changed files with 35 additions and 34 deletions

View file

@ -718,12 +718,10 @@ open Lorentz
lemma altLeftAltRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) )
(hv : IsSelfAdjoint v) (M : SL(2,)) :
TensorProduct.map (altLeftHanded.ρ M) (altRightHanded.ρ M) (altLeftAltRightToMatrix.symm v) =
altLeftAltRightToMatrix.symm
(SL2C.toLinearMapSelfAdjointMatrix (M.transpose⁻¹) ⟨v, hv⟩) := by
altLeftAltRightToMatrix.symm (SL2C.toSelfAdjointMap (M.transpose⁻¹) ⟨v, hv⟩) := by
rw [altLeftAltRightToMatrix_ρ_symm]
apply congrArg
simp only [MonoidHom.coe_mk, OneHom.coe_mk,
SL2C.toLinearMapSelfAdjointMatrix_apply_coe, SpecialLinearGroup.coe_inv,
simp only [SL2C.toSelfAdjointMap_apply_coe, SpecialLinearGroup.coe_inv,
SpecialLinearGroup.coe_transpose]
congr
· rw [SL2C.inverse_coe]
@ -737,8 +735,7 @@ lemma altLeftAltRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2)
lemma leftRightToMatrix_ρ_symm_selfAdjoint (v : Matrix (Fin 2) (Fin 2) )
(hv : IsSelfAdjoint v) (M : SL(2,)) :
TensorProduct.map (leftHanded.ρ M) (rightHanded.ρ M) (leftRightToMatrix.symm v) =
leftRightToMatrix.symm
(SL2C.toLinearMapSelfAdjointMatrix M ⟨v, hv⟩) := by
leftRightToMatrix.symm (SL2C.toSelfAdjointMap M ⟨v, hv⟩) := by
rw [leftRightToMatrix_ρ_symm]
rfl