2024-09-19 06:07:27 -04:00
|
|
|
|
/-
|
|
|
|
|
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 HepLean.StandardModel.Basic
|
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
# The Pati-Salam Model
|
|
|
|
|
|
2024-09-20 17:38:48 -04:00
|
|
|
|
The Pati-Salam model is a petite unified theory that unifies the Standard Model gauge group into
|
2024-09-19 06:07:27 -04:00
|
|
|
|
`SU(4) x SU(2) x SU(2)`.
|
|
|
|
|
|
2024-09-24 08:55:30 +00:00
|
|
|
|
This file currently contains informal-results about the Pati-Salam group.
|
2024-09-19 06:07:27 -04:00
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
namespace PatiSalam
|
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
## The Pati-Salam gauge group.
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
/-- The gauge group of the Pati-Salam model (unquotiented by ℤ₂), i.e., `SU(4) × SU(2) × SU(2)`. -/
|
2024-09-19 06:07:27 -04:00
|
|
|
|
informal_definition GaugeGroupI where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := []
|
2024-09-19 06:07:27 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
/-- The homomorphism of the Standard Model gauge group into the Pati-Salam gauge group, i.e., the
|
|
|
|
|
group homomorphism `SU(3) × SU(2) × U(1) → SU(4) × SU(2) × SU(2)` taking `(h, g, α)` to
|
|
|
|
|
`(blockdiag (α h, α ^ (-3)), g, diag (α ^ 3, α ^(-3))`.
|
|
|
|
|
|
|
|
|
|
See page 54 of https://math.ucr.edu/home/baez/guts.pdf
|
|
|
|
|
-/
|
2024-09-20 17:38:48 -04:00
|
|
|
|
informal_definition inclSM where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``GaugeGroupI, ``StandardModel.GaugeGroupI]
|
|
|
|
|
|
|
|
|
|
/-- The kernel of the map `inclSM` is equal to the subgroup `StandardModel.gaugeGroupℤ₃SubGroup`.
|
2024-09-19 06:07:27 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
See footnote 10 of https://arxiv.org/pdf/2201.07245
|
|
|
|
|
-/
|
2024-09-20 17:38:48 -04:00
|
|
|
|
informal_lemma inclSM_ker where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``inclSM, ``StandardModel.gaugeGroupℤ₃SubGroup]
|
2024-09-20 17:38:48 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
/-- The group embedding from `StandardModel.GaugeGroupℤ₃` to `GaugeGroupI` induced by `inclSM` by
|
2025-02-10 10:51:44 +00:00
|
|
|
|
quotienting by the kernel `inclSM_ker`.
|
2025-02-02 03:17:17 +08:00
|
|
|
|
-/
|
2024-09-20 17:38:48 -04:00
|
|
|
|
informal_definition embedSMℤ₃ where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``inclSM, ``StandardModel.GaugeGroupℤ₃, ``GaugeGroupI, ``inclSM_ker]
|
2024-09-20 17:38:48 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
/-- The equivalence between `GaugeGroupI` and `Spin(6) × Spin(4)`. -/
|
2024-09-19 06:07:27 -04:00
|
|
|
|
informal_definition gaugeGroupISpinEquiv where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``GaugeGroupI]
|
|
|
|
|
|
|
|
|
|
/-- 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 the non-trivial element `(-1, -1, -1)`.
|
2024-09-19 06:07:27 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
See https://math.ucr.edu/home/baez/guts.pdf
|
|
|
|
|
-/
|
2024-09-19 06:07:27 -04:00
|
|
|
|
informal_definition gaugeGroupℤ₂SubGroup where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``GaugeGroupI]
|
|
|
|
|
|
|
|
|
|
/-- The gauge group of the Pati-Salam model with a ℤ₂ quotient, i.e., the quotient of `GaugeGroupI`
|
|
|
|
|
by the ℤ₂-subgroup `gaugeGroupℤ₂SubGroup`.
|
2024-09-19 06:07:27 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
See https://math.ucr.edu/home/baez/guts.pdf
|
|
|
|
|
-/
|
2024-09-19 06:07:27 -04:00
|
|
|
|
informal_definition GaugeGroupℤ₂ where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``GaugeGroupI, ``gaugeGroupℤ₂SubGroup]
|
2024-09-19 06:07:27 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
/-- The group `StandardModel.gaugeGroupℤ₆SubGroup` under the homomorphism `embedSM` factors through
|
|
|
|
|
the subgroup `gaugeGroupℤ₂SubGroup`.
|
|
|
|
|
-/
|
2024-09-19 06:07:27 -04:00
|
|
|
|
informal_lemma sm_ℤ₆_factor_through_gaugeGroupℤ₂SubGroup where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``inclSM, ``StandardModel.gaugeGroupℤ₆SubGroup, ``gaugeGroupℤ₂SubGroup]
|
2024-09-19 06:07:27 -04:00
|
|
|
|
|
2025-02-02 03:17:17 +08:00
|
|
|
|
/-- The group homomorphism from `StandardModel.GaugeGroupℤ₆` to `GaugeGroupℤ₂` induced by `embedSM`.
|
|
|
|
|
-/
|
2024-09-19 06:07:27 -04:00
|
|
|
|
informal_definition embedSMℤ₆Toℤ₂ where
|
2025-02-02 03:17:17 +08:00
|
|
|
|
deps := [``inclSM, ``StandardModel.GaugeGroupℤ₆, ``GaugeGroupℤ₂,
|
2024-09-19 06:07:27 -04:00
|
|
|
|
``sm_ℤ₆_factor_through_gaugeGroupℤ₂SubGroup]
|
|
|
|
|
|
|
|
|
|
end PatiSalam
|