PhysLean/HepLean/StandardModel/Basic.lean

128 lines
4.6 KiB
Text
Raw Normal View History

2024-05-03 06:12:59 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
2024-07-12 16:39:44 -04:00
Released under Apache 2.0 license as described in the file LICENSE.
2024-05-03 06:12:59 -04:00
Authors: Joseph Tooby-Smith
-/
import Mathlib.Geometry.Manifold.Instances.Real
import HepLean.SpaceTime.Basic
import HepLean.Meta.Informal.Basic
2024-05-06 11:09:37 -04:00
/-!
# The Standard Model
2024-05-03 06:12:59 -04:00
2024-05-06 11:09:37 -04:00
This file defines the basic properties of the standard model in particle physics.
-/
2025-01-22 10:32:39 +00:00
TODO "Redefine the gauge group as a quotient of SU(3) x SU(2) x U(1) by a subgroup of ℤ₆."
2024-05-03 06:12:59 -04:00
universe v u
namespace StandardModel
open Manifold
open Matrix
open Complex
open ComplexConjugate
/-- 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 :=
2024-05-09 15:09:14 -04:00
specialUnitaryGroup (Fin 3) × specialUnitaryGroup (Fin 2) × unitary
2024-05-09 08:16:23 -04:00
/-- The subgroup of the un-quotiented gauge group which acts trivially on all particles in the
standard model, i.e., the ℤ₆-subgroup of `GaugeGroupI` with elements `(α^2 * I₃, α^(-3) * I₂, α)`,
where `α` is a sixth complex root of unity.
See https://math.ucr.edu/home/baez/guts.pdf
-/
informal_definition gaugeGroup₆SubGroup where
deps := [``GaugeGroupI]
/-- The smallest possible gauge group of the Standard Model, i.e., the quotient of `GaugeGroupI` by
the ℤ₆-subgroup `gaugeGroup₆SubGroup`.
See https://math.ucr.edu/home/baez/guts.pdf
-/
informal_definition GaugeGroup₆ where
deps := [``GaugeGroupI, ``StandardModel.gaugeGroup₆SubGroup]
/-- The ₂subgroup of the un-quotiented gauge group which acts trivially on all particles in the
standard model, i.e., the ℤ₂-subgroup of `GaugeGroupI` derived from the ℤ₂ subgroup of
`gaugeGroup₆SubGroup`.
See https://math.ucr.edu/home/baez/guts.pdf
-/
informal_definition gaugeGroup₂SubGroup where
deps := [``GaugeGroupI, ``StandardModel.gaugeGroup₆SubGroup]
/-- The guage group of the Standard Model with a ℤ₂ quotient, i.e., the quotient of `GaugeGroupI` by
the ℤ₂-subgroup `gaugeGroup₂SubGroup`.
See https://math.ucr.edu/home/baez/guts.pdf
-/
informal_definition GaugeGroup₂ where
deps := [``GaugeGroupI, ``StandardModel.gaugeGroup₂SubGroup]
/-- The ℤ₃-subgroup of the un-quotiented gauge group which acts trivially on all particles in the
standard model, i.e., the ℤ₃-subgroup of `GaugeGroupI` derived from the ℤ₃ subgroup of
`gaugeGroup₆SubGroup`.
See https://math.ucr.edu/home/baez/guts.pdf
-/
informal_definition gaugeGroup₃SubGroup where
deps := [``GaugeGroupI, ``StandardModel.gaugeGroup₆SubGroup]
/-- The guage group of the Standard Model with a ℤ₃-quotient, i.e., the quotient of `GaugeGroupI` by
the ℤ₃-subgroup `gaugeGroup₃SubGroup`.
See https://math.ucr.edu/home/baez/guts.pdf
-/
informal_definition GaugeGroup₃ where
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
2024-11-26 09:33:24 +00:00
/-- The element of `GaugeGroupQuot` corresponding to the quotient of the full SM gauge group
by the sub-group `ℤ₆`. -/
| ℤ₆ : GaugeGroupQuot
2024-11-26 09:33:24 +00:00
/-- The element of `GaugeGroupQuot` corresponding to the quotient of the full SM gauge group
by the sub-group `ℤ₂`. -/
| ℤ₂ : GaugeGroupQuot
2024-11-26 09:33:24 +00:00
/-- The element of `GaugeGroupQuot` corresponding to the quotient of the full SM gauge group
by the sub-group `ℤ₃`. -/
| ℤ₃ : GaugeGroupQuot
2024-11-26 09:33:24 +00:00
/-- The element of `GaugeGroupQuot` corresponding to the full SM gauge group. -/
| I : GaugeGroupQuot
/-- The (global) gauge group of the Standard Model given a choice of quotient, i.e., the map from
`GaugeGroupQuot` to `Type` which gives the gauge group of the Standard Model for a given choice of
quotient.
See https://math.ucr.edu/home/baez/guts.pdf
-/
informal_definition GaugeGroup where
deps := [``GaugeGroupI, ``gaugeGroup₆SubGroup, ``gaugeGroup₂SubGroup, ``gaugeGroup₃SubGroup,
``GaugeGroupQuot]
/-!
## Smoothness structure on the gauge group.
-/
/-- The gauge group `GaugeGroupI` is a Lie group. -/
informal_lemma gaugeGroupI_lie where
deps := [``GaugeGroupI]
/-- For every `q` in `GaugeGroupQuot` the group `GaugeGroup q` is a Lie group. -/
informal_lemma gaugeGroup_lie where
deps := [``GaugeGroup]
/-- The trivial principal bundle over SpaceTime with structure group `GaugeGroupI`. -/
informal_definition gaugeBundleI where
deps := [``GaugeGroupI, ``SpaceTime]
/-- A global section of `gaugeBundleI`. -/
informal_definition gaugeTransformI where
deps := [``gaugeBundleI]
2024-05-03 06:12:59 -04:00
end StandardModel