refactor: LInt

This commit is contained in:
jstoobysmith 2024-11-09 18:12:05 +00:00
parent e963be5ef8
commit d058f41689
11 changed files with 38 additions and 50 deletions

View file

@ -86,7 +86,6 @@ lemma toFin1d_eq_val (ψ : ContrMod d) : ψ.toFin1d = ψ.val := by rfl
@[simps!]
def stdBasis : Basis (Fin 1 ⊕ Fin d) (ContrMod d) := Basis.ofEquivFun toFin1dEquiv
@[simp]
lemma stdBasis_toFin1dEquiv_apply_same (μ : Fin 1 ⊕ Fin d) :
toFin1dEquiv (stdBasis μ) μ = 1 := by
@ -121,7 +120,7 @@ lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ =
/-- Decomposition of a contrvariant Lorentz vector into the standard basis. -/
lemma stdBasis_decomp (v : ContrMod d) : v = ∑ i, v.toFin1d i • stdBasis i := by
apply toFin1dEquiv.injective
simp only [map_sum, _root_.map_smul]
simp only [map_sum, _root_.map_smul]
funext μ
rw [Fintype.sum_apply μ fun c => toFin1dEquiv v c • toFin1dEquiv (stdBasis c)]
change _ = ∑ x : Fin 1 ⊕ Fin d, toFin1dEquiv v x • (toFin1dEquiv (stdBasis x) μ)
@ -189,7 +188,6 @@ def toSpace (v : ContrMod d) : EuclideanSpace (Fin d) := v.val ∘ Sum.inr
-/
/-- The representation of the Lorentz group acting on `ContrModule d`. -/
def rep : Representation (LorentzGroup d) (ContrMod d) where
toFun g := Matrix.toLinAlgEquiv stdBasis g
@ -237,7 +235,6 @@ lemma toSelfAdjoint_apply_coe (x : ContrMod 3) : (toSelfAdjoint x).1 =
rw [toSelfAdjoint_apply]
rfl
lemma toSelfAdjoint_stdBasis (i : Fin 1 ⊕ Fin 3) :
toSelfAdjoint (stdBasis i) = PauliMatrix.σSAL i := by
rw [toSelfAdjoint_apply]
@ -352,11 +349,10 @@ lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ =
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
/-- Decomposition of a covariant Lorentz vector into the standard basis. -/
lemma stdBasis_decomp (v : CoMod d) : v = ∑ i, v.toFin1d i • stdBasis i := by
apply toFin1dEquiv.injective
simp only [map_sum, _root_.map_smul]
simp only [map_sum, _root_.map_smul]
funext μ
rw [Fintype.sum_apply μ fun c => toFin1dEquiv v c • toFin1dEquiv (stdBasis c)]
change _ = ∑ x : Fin 1 ⊕ Fin d, toFin1dEquiv v x • (toFin1dEquiv (stdBasis x) μ)