feat: Add properties of Lorentz group

This commit is contained in:
jstoobysmith 2024-05-13 07:42:55 -04:00
parent b9c20861d1
commit db0f6de59c
2 changed files with 173 additions and 21 deletions

View file

@ -50,16 +50,94 @@ namespace spaceTime
def asSmoothManifold : ModelWithCorners spaceTime spaceTime := 𝓘(, spaceTime)
@[simp]
def stdBasis : Basis (Fin 4) spaceTime := Pi.basisFun (Fin 4)
lemma stdBasis_apply (μ ν : Fin 4) : stdBasis μ ν = if μ = ν then 1 else 0 := by
erw [stdBasis, Pi.basisFun_apply, LinearMap.stdBasis_apply']
lemma stdBasis_not_eq {μ ν : Fin 4} (h : μ ≠ ν) : stdBasis μ ν = 0 := by
rw [stdBasis_apply]
exact if_neg h
lemma stdBasis_0 : stdBasis 0 = ![1, 0, 0, 0] := by
funext i
fin_cases i
<;> simp [stdBasis_apply]
lemma stdBasis_1 : stdBasis 1 = ![0, 1, 0, 0] := by
funext i
fin_cases i
<;> simp [stdBasis_apply]
lemma stdBasis_2 : stdBasis 2 = ![0, 0, 1, 0] := by
funext i
fin_cases i
<;> simp [stdBasis_apply]
lemma stdBasis_3 : stdBasis 3 = ![0, 0, 0, 1] := by
funext i
fin_cases i
<;> simp [stdBasis_apply]
lemma stdBasis_mulVec (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ) :
(Λ *ᵥ stdBasis μ) ν = Λ ν μ := by
rw [mulVec, dotProduct, Fintype.sum_eq_single μ, stdBasis_apply]
simp
intro x h
rw [stdBasis_apply, if_neg (Ne.symm h)]
simp
lemma explicit (x : spaceTime) : x = ![x 0, x 1, x 2, x 3] := by
funext i
fin_cases i <;> rfl
def η : Matrix (Fin 4) (Fin 4) :=
![![1, 0, 0, 0], ![0, -1, 0, 0], ![0, 0, -1, 0], ![0, 0, 0, -1]]
lemma η_symmetric (μ ν : Fin 4) : η μ ν = η ν μ := by
lemma η_off_diagonal {μ ν : Fin 4} (h : μ ≠ ν) : η μ ν = 0 := by
fin_cases μ <;>
fin_cases ν <;>
simp only [η, Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one, Matrix.cons_val_one,
simp_all [η, Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one, Matrix.cons_val_one,
Matrix.cons_val_succ', Matrix.cons_val_zero, Matrix.empty_val', Matrix.head_cons,
Matrix.head_fin_const, Matrix.head_cons, Matrix.vecCons_const, Fin.mk_one, Fin.mk_one]
Matrix.head_fin_const, Matrix.head_cons, Matrix.vecCons_const, Fin.mk_one, Fin.mk_one,
vecHead, vecTail, Function.comp_apply]
lemma η_symmetric (μ ν : Fin 4) : η μ ν = η ν μ := by
by_cases h : μ = ν
rw [h]
rw [η_off_diagonal h]
refine (η_off_diagonal ?_).symm
exact fun a => h (id a.symm)
lemma η_transpose : η.transpose = η := by
funext μ ν
rw [transpose_apply, η_symmetric]
lemma η_sq : η * η = 1 := by
funext μ ν
rw [mul_apply, Fin.sum_univ_four]
fin_cases μ <;> fin_cases ν <;>
simp [η, Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one, Matrix.cons_val_one,
Matrix.cons_val_succ', Matrix.cons_val_zero, Matrix.empty_val', Matrix.head_cons,
Matrix.head_fin_const, Matrix.head_cons, Matrix.vecCons_const, Fin.mk_one, Fin.mk_one,
vecHead, vecTail, Function.comp_apply]
lemma η_mulVec (x : spaceTime) : η *ᵥ x = ![x 0, -x 1, -x 2, -x 3] := by
rw [explicit x]
rw [η]
funext i
rw [mulVec, dotProduct, Fin.sum_univ_four]
fin_cases i <;>
simp [η, Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one, Matrix.cons_val_one,
Matrix.cons_val_succ', Matrix.cons_val_zero, Matrix.empty_val', Matrix.head_cons,
Matrix.head_fin_const, Matrix.head_cons, Matrix.vecCons_const, Fin.mk_one, Fin.mk_one,
vecHead, vecTail, Function.comp_apply]
@[simps!]
def linearMapForSpaceTime (x : spaceTime) : spaceTime →ₗ[] where
@ -72,6 +150,7 @@ def linearMapForSpaceTime (x : spaceTime) : spaceTime →ₗ[] where
rw [mulVec_smul, dotProduct_smul]
rfl
@[simps!]
def ηLin : LinearMap.BilinForm spaceTime where
toFun x := linearMapForSpaceTime x
map_add' x y := by
@ -86,13 +165,74 @@ def ηLin : LinearMap.BilinForm spaceTime where
rw [smul_dotProduct]
rfl
lemma ηLin_expand (x y : spaceTime) : ηLin x y = x 0 * y 0 - x 1 * y 1 - x 2 * y 2 - x 3 * y 3 := by
rw [ηLin]
simp only [LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply, Fin.isValue]
erw [η_mulVec]
nth_rewrite 1 [explicit x]
simp only [dotProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, Fin.sum_univ_four,
cons_val_zero, cons_val_one, head_cons, mul_neg, cons_val_two, tail_cons, cons_val_three]
ring
lemma ηLin_symm (x y : spaceTime) : ηLin x y = ηLin y x := by
rw [ηLin_expand, ηLin_expand]
ring
lemma ηLin_stdBasis_apply (μ : Fin 4) (x : spaceTime) : ηLin (stdBasis μ) x = η μ μ * x μ := by
rw [ηLin_expand]
fin_cases μ
<;> simp [stdBasis_0, stdBasis_1, stdBasis_2, stdBasis_3, η]
lemma ηLin_η_stdBasis (μ ν : Fin 4) : ηLin (stdBasis μ) (stdBasis ν) = η μ ν := by
rw [ηLin_stdBasis_apply]
by_cases h : μ = ν
rw [stdBasis_apply]
subst h
simp
rw [stdBasis_not_eq, η_off_diagonal h]
simp
exact fun a => h (id a.symm)
lemma ηLin_mulVec_left (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
ηLin (Λ *ᵥ x) y = ηLin x ((η * Λᵀ * η) *ᵥ y) := by
simp only [ηLin_apply_apply, mulVec_mulVec]
rw [(vecMul_transpose Λ x).symm, ← dotProduct_mulVec, mulVec_mulVec]
rw [← mul_assoc, ← mul_assoc, η_sq, one_mul]
lemma ηLin_mulVec_right (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ) :
ηLin x (Λ *ᵥ y) = ηLin ((η * Λᵀ * η) *ᵥ x) y := by
rw [ηLin_symm, ηLin_symm ((η * Λᵀ * η) *ᵥ x) _ ]
exact ηLin_mulVec_left y x Λ
lemma ηLin_matrix_stdBasis (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ) :
ηLin (stdBasis ν) (Λ *ᵥ stdBasis μ) = η ν ν * Λ ν μ := by
rw [ηLin_stdBasis_apply, stdBasis_mulVec]
lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ) :
Λ = 1 ↔ ∀ (x y : spaceTime), ηLin x y = ηLin x (Λ *ᵥ y) := by
apply Iff.intro
intro h
subst h
simp
intro h
funext μ ν
have h1 := h (stdBasis μ) (stdBasis ν)
rw [ηLin_matrix_stdBasis, ηLin_η_stdBasis] at h1
fin_cases μ <;> fin_cases ν <;>
simp_all [η, Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one, Matrix.cons_val_one,
Matrix.cons_val_succ', Matrix.cons_val_zero, Matrix.empty_val', Matrix.head_cons,
Matrix.head_fin_const, Matrix.head_cons, Matrix.vecCons_const, Fin.mk_one, Fin.mk_one,
vecHead, vecTail, Function.comp_apply]
def quadraticForm : QuadraticForm spaceTime := ηLin.toQuadraticForm
def lorentzGroup : Subgroup (GeneralLinearGroup (Fin 4) ) where
carrier := {Λ | ∀ (x y : spaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y}
mul_mem' {a b} := by
intros ha hb x y
simp
simp only [Units.val_mul, mulVec_mulVec]
rw [← mulVec_mulVec, ← mulVec_mulVec, ha, hb]
one_mem' := by
intros x y
@ -105,6 +245,31 @@ def lorentzGroup : Subgroup (GeneralLinearGroup (Fin 4) ) where
simp [hx]
lemma mem_lorentzGroup_iff (Λ : GeneralLinearGroup (Fin 4) ) :
Λ ∈ lorentzGroup ↔ ∀ (x y : spaceTime), ηLin (Λ *ᵥ x) (Λ *ᵥ y) = ηLin x y := by
rfl
lemma mem_lorentzGroup_iff' (Λ : GeneralLinearGroup (Fin 4) ) :
Λ ∈ lorentzGroup ↔ ∀ (x y : spaceTime), ηLin (x) ((η * Λ.1ᵀ * η * Λ.1) *ᵥ y) = ηLin x y := by
rw [mem_lorentzGroup_iff]
apply Iff.intro
intro h
intro x y
have h1 := h x y
rw [ηLin_mulVec_left, mulVec_mulVec] at h1
exact h1
intro h
intro x y
rw [ηLin_mulVec_left, mulVec_mulVec]
exact h x y
lemma mem_lorentzGroup_iff'' (Λ : GeneralLinearGroup (Fin 4) ) :
Λ ∈ lorentzGroup ↔ η * Λ.1ᵀ * η * Λ.1 = 1 := by
rw [mem_lorentzGroup_iff', ηLin_matrix_eq_identity_iff (η * Λ.1ᵀ * η * Λ.1)]
apply Iff.intro
· simp_all only [ηLin_apply_apply, implies_true, iff_true, one_mulVec]
· simp_all only [ηLin_apply_apply, mulVec_mulVec, implies_true]
abbrev cliffordAlgebra := CliffordAlgebra quadraticForm
end spaceTime

View file

@ -13,7 +13,7 @@ This file defines the Gamma matrices.
- Prove that the algebra generated by the gamma matrices is ismorphic to the
Clifford algebra assocaited with spacetime.
-
- Include relations for gamma matrices.
-/
namespace StandardModel
@ -41,22 +41,9 @@ def γ : Fin 4 → Matrix (Fin 4) (Fin 4) := ![γ0, γ1, γ2, γ3]
namespace γ
variable (μ : Fin 4)
/-- The trace of the gamma matrices is zero. -/
lemma trace_eq_zero (μ : Fin 4) : Matrix.trace (γ μ) = 0 := by
fin_cases μ
<;> simp [γ, γ0, γ1, γ2, γ3]
<;> rw [Matrix.trace, Fin.sum_univ_four]
<;> simp
any_goals rfl
change 0 + 0 = 0
simp [add_zero]
open spaceTime
variable (μ ν : Fin 4)
@[simp]
def γSet : Set (Matrix (Fin 4) (Fin 4) ) := {γ i | i : Fin 4}