feat: Add informal def for SM gauge group.

This commit is contained in:
jstoobysmith 2024-09-18 07:37:33 -04:00
parent 62f1153088
commit 3aa00ff1e7
3 changed files with 82 additions and 17 deletions

View file

@ -6,6 +6,7 @@ Authors: Joseph Tooby-Smith
import Mathlib.Data.Complex.Exponential
import Mathlib.Geometry.Manifold.Instances.Real
import Mathlib.LinearAlgebra.Matrix.ToLin
import HepLean.Meta.Informal
/-!
# The Standard Model
@ -21,8 +22,65 @@ open Matrix
open Complex
open ComplexConjugate
/-- The global gauge group of the standard model. -/
abbrev GaugeGroup : Type :=
/-- The global gauge group of the Standard Model with no discrete quotients.
The `I` in the Name is an indication of the statement that this has no discrete quotients. -/
abbrev GaugeGroupI : Type :=
specialUnitaryGroup (Fin 3) × specialUnitaryGroup (Fin 2) × unitary
informal_definition gaugeGroup₆SubGroup where
physics :≈ "The subgroup of the un-quotiented gauge group which acts trivially on
all particles in the standard model. "
math :≈ "The ℤ₆-subgroup of ``GaugeGroupI with elements (α^2 * I₃, α^(-3) * I₂, α), where `α`
is a sixth complex root of unity."
ref :≈ "https://math.ucr.edu/home/baez/guts.pdf"
deps :≈ [``GaugeGroupI]
informal_definition GaugeGroup₆ where
physics :≈ "The smallest possible gauge group of the Standard Model."
math :≈ "The quotient of ``GaugeGroupI by the ℤ₆-subgroup `gaugeGroup₆SubGroup`."
ref :≈ "https://math.ucr.edu/home/baez/guts.pdf"
deps :≈ [``GaugeGroupI, ``StandardModel.gaugeGroup₆SubGroup]
informal_definition gaugeGroup₂SubGroup where
physics :≈ "The ₂subgroup of the un-quotiented gauge group which acts trivially on
all particles in the standard model. "
math :≈ "The ℤ₂-subgroup of ``GaugeGroupI derived from the ℤ₂ subgroup of `gaugeGroup₆SubGroup`."
ref :≈ "https://math.ucr.edu/home/baez/guts.pdf"
deps :≈ [``GaugeGroupI, ``StandardModel.gaugeGroup₆SubGroup]
informal_definition GaugeGroup₂ where
physics :≈ "The guage group of the Standard Model with a ℤ₂ quotient."
math :≈ "The quotient of ``GaugeGroupI by the ℤ₂-subgroup `gaugeGroup₂SubGroup`."
ref :≈ "https://math.ucr.edu/home/baez/guts.pdf"
deps :≈ [``GaugeGroupI, ``StandardModel.gaugeGroup₂SubGroup]
informal_definition gaugeGroup₃SubGroup where
physics :≈ "The ℤ₃-subgroup of the un-quotiented gauge group which acts trivially on
all particles in the standard model. "
math :≈ "The ℤ₃-subgroup of ``GaugeGroupI derived from the ℤ₃ subgroup of `gaugeGroup₆SubGroup`."
ref :≈ "https://math.ucr.edu/home/baez/guts.pdf"
deps :≈ [``GaugeGroupI, ``StandardModel.gaugeGroup₆SubGroup]
informal_definition GaugeGroup₃ where
physics :≈ "The guage group of the Standard Model with a ℤ₃-quotient."
math :≈ "The quotient of ``GaugeGroupI by the ℤ₃-subgroup `gaugeGroup₃SubGroup`."
ref :≈ "https://math.ucr.edu/home/baez/guts.pdf"
deps :≈ [``GaugeGroupI, ``StandardModel.gaugeGroup₃SubGroup]
/-- Specifies the allowed quotients of `SU(3) x SU(2) x U(1)` which give a valid
gauge group of the Standard Model. -/
inductive GaugeGroupQuot : Type
| ℤ₆ : GaugeGroupQuot
| ℤ₂ : GaugeGroupQuot
| ℤ₃ : GaugeGroupQuot
| I : GaugeGroupQuot
informal_definition GaugeGroup where
physics :≈ "The (global) gauge group of the Standard Model given a choice of quotient."
math :≈ "The map from `GaugeGroupQuot` to `Type` which gives the gauge group of the Standard Model
for a given choice of quotient."
ref :≈ "https://math.ucr.edu/home/baez/guts.pdf"
deps :≈ [``GaugeGroupI, ``gaugeGroup₆SubGroup, ``gaugeGroup₂SubGroup, ``gaugeGroup₃SubGroup,
``GaugeGroupQuot]
end StandardModel