feat: add properties of spacetime

This commit is contained in:
jstoobysmith 2024-05-10 16:57:45 -04:00
parent 4ff3d9bc0b
commit b9c20861d1
2 changed files with 160 additions and 2 deletions

View file

@ -0,0 +1,79 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.StandardModel.Basic
/-!
# The Clifford Algebra
This file defines the Gamma matrices.
## TODO
- Prove that the algebra generated by the gamma matrices is ismorphic to the
Clifford algebra assocaited with spacetime.
-
-/
namespace StandardModel
open Complex
noncomputable section diracRepresentation
def γ0 : Matrix (Fin 4) (Fin 4) :=
![![1, 0, 0, 0], ![0, 1, 0, 0], ![0, 0, -1, 0], ![0, 0, 0, -1]]
def γ1 : Matrix (Fin 4) (Fin 4) :=
![![0, 0, 0, 1], ![0, 0, 1, 0], ![0, -1, 0, 0], ![-1, 0, 0, 0]]
def γ2 : Matrix (Fin 4) (Fin 4) :=
![![0, 0, 0, - I], ![0, 0, I, 0], ![0, I, 0, 0], ![-I, 0, 0, 0]]
def γ3 : Matrix (Fin 4) (Fin 4) :=
![![0, 0, 1, 0], ![0, 0, 0, -1], ![-1, 0, 0, 0], ![0, 1, 0, 0]]
def γ5 : Matrix (Fin 4) (Fin 4) := I • (γ0 * γ1 * γ2 * γ3)
@[simp]
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]
@[simp]
def γSet : Set (Matrix (Fin 4) (Fin 4) ) := {γ i | i : Fin 4}
lemma γ_in_γSet (μ : Fin 4) : γ μ ∈ γSet := by
simp [γSet]
def diracAlgebra : Subalgebra (Matrix (Fin 4) (Fin 4) ) :=
Algebra.adjoin γSet
lemma γSet_subset_diracAlgebra : γSet ⊆ diracAlgebra :=
Algebra.subset_adjoin
lemma γ_in_diracAlgebra (μ : Fin 4) : γ μ ∈ diracAlgebra :=
γSet_subset_diracAlgebra (γ_in_γSet μ)
end γ
end diracRepresentation
end StandardModel