docs: For SO(3)

This commit is contained in:
jstoobysmith 2024-11-26 09:33:24 +00:00
parent 8324c95451
commit fe0f2c26c7
3 changed files with 18 additions and 0 deletions

View file

@ -71,9 +71,16 @@ informal_definition GaugeGroup₃ where
/-- 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
/-- The element of `GaugeGroupQuot` corresponding to the quotient of the full SM gauge group
by the sub-group `ℤ₆`. -/
| ℤ₆ : GaugeGroupQuot
/-- The element of `GaugeGroupQuot` corresponding to the quotient of the full SM gauge group
by the sub-group `ℤ₂`. -/
| ℤ₂ : GaugeGroupQuot
/-- The element of `GaugeGroupQuot` corresponding to the quotient of the full SM gauge group
by the sub-group `ℤ₃`. -/
| ℤ₃ : GaugeGroupQuot
/-- The element of `GaugeGroupQuot` corresponding to the full SM gauge group. -/
| I : GaugeGroupQuot
informal_definition GaugeGroup where

View file

@ -103,6 +103,8 @@ def HiggsVec.toField (φ : HiggsVec) : HiggsField where
rw [Bundle.contMDiffAt_section]
exact smoothAt_const
/-- For all spacetime points, the constant Higgs field defined by a Higgs vector,
returns that Higgs Vector. -/
@[simp]
lemma HiggsVec.toField_apply (φ : HiggsVec) (x : SpaceTime) : (φ.toField x) = φ := rfl