PhysLean/HepLean/SpaceTime/CliffordAlgebra.lean
2024-07-12 16:39:44 -04:00

70 lines
2.1 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Analysis.Complex.Basic
/-!
# The Clifford Algebra
This file defines the Gamma matrices.
-/
/-! TODO: Prove algebra generated by gamma matrices is isomorphic to Clifford algebra. -/
/-! TODO: Define relations between the gamma matrices. -/
namespace spaceTime
open Complex
noncomputable section diracRepresentation
/-- The γ⁰ gamma matrix in the Dirac representation. -/
def γ0 : Matrix (Fin 4) (Fin 4) :=
![![1, 0, 0, 0], ![0, 1, 0, 0], ![0, 0, -1, 0], ![0, 0, 0, -1]]
/-- The γ¹ gamma matrix in the Dirac representation. -/
def γ1 : Matrix (Fin 4) (Fin 4) :=
![![0, 0, 0, 1], ![0, 0, 1, 0], ![0, -1, 0, 0], ![-1, 0, 0, 0]]
/-- The γ² gamma matrix in the Dirac representation. -/
def γ2 : Matrix (Fin 4) (Fin 4) :=
![![0, 0, 0, - I], ![0, 0, I, 0], ![0, I, 0, 0], ![-I, 0, 0, 0]]
/-- The γ³ gamma matrix in the Dirac representation. -/
def γ3 : Matrix (Fin 4) (Fin 4) :=
![![0, 0, 1, 0], ![0, 0, 0, -1], ![-1, 0, 0, 0], ![0, 1, 0, 0]]
/-- The γ⁵ gamma matrix in the Dirac representation. -/
def γ5 : Matrix (Fin 4) (Fin 4) := I • (γ0 * γ1 * γ2 * γ3)
/-- The γ gamma matrices in the Dirac representation. -/
@[simp]
def γ : Fin 4 → Matrix (Fin 4) (Fin 4) := ![γ0, γ1, γ2, γ3]
namespace γ
open spaceTime
variable (μ ν : Fin 4)
/-- The subset of `Matrix (Fin 4) (Fin 4) ` formed by the gamma matrices in the Dirac
representation. -/
@[simp]
def γSet : Set (Matrix (Fin 4) (Fin 4) ) := {γ i | i : Fin 4}
lemma γ_in_γSet (μ : Fin 4) : γ μ ∈ γSet := by
simp [γSet]
/-- The algebra generated by the gamma matrices in the Dirac representation. -/
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 spaceTime