refactor: Simps

This commit is contained in:
jstoobysmith 2024-11-09 18:06:48 +00:00
parent a7142ef99b
commit e963be5ef8
5 changed files with 20 additions and 18 deletions

View file

@ -98,7 +98,7 @@ lemma genBoost_mul_one_plus_contr (u v : FuturePointing d) (x : Contr d) :
· congr 1
· congr 1
rw [genBoostAux₁]
simp
simp only [LinearMap.coe_mk, AddHom.coe_mk]
rw [smul_smul]
congr 1
ring
@ -198,7 +198,7 @@ lemma toMatrix_in_lorentzGroup (u v : FuturePointing d) : (toMatrix u v) ∈ Lor
smul_tmul, tmul_smul, map_sub, map_smul, smul_eq_mul, contr_self, mul_one]
rw [contrContrContractField.symm v.1.1 u, contrContrContractField.symm v.1.1 x,
contrContrContractField.symm u.1.1 x]
simp
simp only [smul_eq_mul]
ring
have hn (a : ) {b c : } (h : a ≠ 0) (hab : a * b = a * c ) : b = c := by
simp_all only [smul_eq_mul, ne_eq, mul_eq_mul_left_iff, or_false]

View file

@ -268,7 +268,7 @@ lemma self_parity_eq_zero_iff : ⟪y, (Contr d).ρ LorentzGroup.parity y⟫ₘ =
have hn := Fintype.sum_eq_zero_iff_of_nonneg (f := fun i => y.val i * y.val i) (fun i => by
simpa using mul_self_nonneg (y.val i))
rw [h] at hn
simp at hn
simp only [true_iff] at hn
apply ContrMod.ext
funext i
simpa using congrFun hn i
@ -398,7 +398,9 @@ lemma basis_left {v : Contr d} (μ : Fin 1 ⊕ Fin d) :
rw [as_sum]
rcases μ with μ | μ
· fin_cases μ
simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal]
simp only [Fin.zero_eta, Fin.isValue, ContrMod.stdBasis_apply_same, one_mul,
ContrMod.stdBasis_inl_apply_inr, zero_mul, Finset.sum_const_zero, sub_zero, minkowskiMatrix,
LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_apply_eq, Sum.elim_inl]
rfl
· simp only [Action.instMonoidalCategory_tensorUnit_V, Fin.isValue, ContrMod.stdBasis_apply,
reduceCtorEq, ↓reduceIte, zero_mul, Sum.inr.injEq, ite_mul, one_mul, Finset.sum_ite_eq,
@ -406,9 +408,9 @@ lemma basis_left {v : Contr d} (μ : Fin 1 ⊕ Fin d) :
diagonal_apply_eq, Sum.elim_inr, neg_mul, neg_inj]
rfl
lemma on_basis_mulVec (μ ν : Fin 1 ⊕ Fin d) : ⟪ContrMod.stdBasis μ, Λ *ᵥ ContrMod.stdBasis ν⟫ₘ = η μ μ * Λ μ ν := by
rw [basis_left]
rw [@ContrMod.mulVec_toFin1d]
lemma on_basis_mulVec (μ ν : Fin 1 ⊕ Fin d) :
⟪ContrMod.stdBasis μ, Λ *ᵥ ContrMod.stdBasis ν⟫ₘ = η μ μ * Λ μ ν := by
rw [basis_left, ContrMod.mulVec_toFin1d]
simp [basis_left, mulVec, dotProduct, ContrMod.stdBasis_apply, ContrMod.toFin1d_eq_val]

View file

@ -112,9 +112,9 @@ lemma stdBasis_inl_apply_inr (i : Fin d) : (stdBasis (Sum.inl 0)).val (Sum.inr i
simp
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ = ν then 1 else 0 := by
simp [stdBasis, Pi.basisFun_apply, Pi.single_apply]
simp only [stdBasis, Basis.coe_ofEquivFun]
change Pi.single μ 1 ν = _
simp [Pi.single_apply]
simp only [Pi.single_apply]
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)
@ -346,9 +346,9 @@ lemma stdBasis_toFin1dEquiv_apply_ne {μ ν : Fin 1 ⊕ Fin d} (h : μ ≠ ν
exact Pi.single_eq_of_ne' h 1
lemma stdBasis_apply (μ ν : Fin 1 ⊕ Fin d) : (stdBasis μ).val ν = if μ = ν then 1 else 0 := by
simp [stdBasis, Pi.basisFun_apply, Pi.single_apply]
simp only [stdBasis, Basis.coe_ofEquivFun]
change Pi.single μ 1 ν = _
simp [Pi.single_apply]
simp only [Pi.single_apply]
refine ite_congr ?h₁ (congrFun rfl) (congrFun rfl)
exact Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a)

View file

@ -290,7 +290,7 @@ noncomputable def pathFromTime (u : FuturePointing d) : Path timeVecNormOneFutur
funext i
match i with
| Sum.inl 0 =>
simp [Set.Icc.coe_one, one_pow, one_mul, Fin.isValue]
simp only [Set.Icc.coe_one, one_pow, one_mul, Fin.isValue]
exact (inl_eq_sqrt u).symm
| Sum.inr i =>
simp only [Set.Icc.coe_one, one_pow, one_mul, Fin.isValue]

View file

@ -97,14 +97,15 @@ open Lorentz in
lemma toMatrix_apply_contrMod (M : SL(2, )) (v : ContrMod 3) :
(toMatrix M) *ᵥ v = ContrMod.toSelfAdjoint.symm
((toLinearMapSelfAdjointMatrix M) (ContrMod.toSelfAdjoint v)) := by
simp [toMatrix, LinearMap.toMatrix_apply, ContrMod.mulVec]
simp only [ContrMod.mulVec, toMatrix, MonoidHom.coe_mk, OneHom.coe_mk]
obtain ⟨a, ha⟩ := ContrMod.toSelfAdjoint.symm.surjective v
subst ha
rw [LinearEquiv.apply_symm_apply]
simp [ContrMod.toSelfAdjoint]
change ContrMod.toFin1dEquiv.symm ((
((LinearMap.toMatrix PauliMatrix.σSAL PauliMatrix.σSAL) (toLinearMapSelfAdjointMatrix M)))
*ᵥ (((Finsupp.linearEquivFunOnFinite (Fin 1 ⊕ Fin 3)) (PauliMatrix.σSAL.repr a)))) = _
simp only [ContrMod.toSelfAdjoint, LinearEquiv.trans_symm, LinearEquiv.symm_symm,
LinearEquiv.trans_apply]
change ContrMod.toFin1dEquiv.symm ((
((LinearMap.toMatrix PauliMatrix.σSAL PauliMatrix.σSAL) (toLinearMapSelfAdjointMatrix M)))
*ᵥ (((Finsupp.linearEquivFunOnFinite (Fin 1 ⊕ Fin 3)) (PauliMatrix.σSAL.repr a)))) = _
apply congrArg
erw [LinearMap.toMatrix_mulVec_repr]
rfl
@ -136,7 +137,6 @@ lemma toLorentzGroup_eq_σSAL (M : SL(2, )) :
PauliMatrix.σSAL PauliMatrix.σSAL (toLinearMapSelfAdjointMatrix M) := by
rfl
lemma toLinearMapSelfAdjointMatrix_basis (i : Fin 1 ⊕ Fin 3) :
toLinearMapSelfAdjointMatrix M (PauliMatrix.σSAL i) =
∑ j, (toLorentzGroup M).1 j i •