PhysLean/HepLean/StandardModel/HiggsBoson/Basic.lean

187 lines
5.7 KiB
Text
Raw Normal View History

2024-05-09 15:09:14 -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-09 15:09:14 -04:00
Authors: Joseph Tooby-Smith
-/
2024-06-25 07:06:32 -04:00
import HepLean.SpaceTime.Basic
2024-05-09 15:09:14 -04:00
import Mathlib.Data.Complex.Exponential
import Mathlib.Tactic.Polyrith
import Mathlib.Geometry.Manifold.VectorBundle.Basic
import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection
import Mathlib.Geometry.Manifold.Instances.Real
import Mathlib.Analysis.InnerProductSpace.Basic
import Mathlib.Geometry.Manifold.ContMDiff.Product
import HepLean.Meta.Informal.Basic
2024-05-09 15:09:14 -04:00
/-!
2024-07-10 11:34:34 -04:00
2024-05-09 15:09:14 -04:00
# The Higgs field
2024-07-10 11:34:34 -04:00
This file defines the basic properties for the Higgs field in the standard model.
2024-05-09 15:09:14 -04:00
## References
- We use conventions given in: [Review of Particle Physics, PDG][ParticleDataGroup:2018ovx]
2024-05-09 15:09:14 -04:00
-/
2024-07-10 11:34:34 -04:00
2024-05-09 15:09:14 -04:00
namespace StandardModel
noncomputable section
open Manifold
open Matrix
open Complex
open ComplexConjugate
open SpaceTime
2024-07-10 11:34:34 -04:00
/-!
## The definition of the Higgs vector space.
In other words, the target space of the Higgs field.
-/
/-- The complex vector space in which the Higgs field takes values. -/
abbrev HiggsVec := EuclideanSpace (Fin 2)
namespace HiggsVec
/-- The continuous linear map from the vector space `HiggsVec` to `(Fin 2 → )` achieved by
casting vectors. -/
def toFin2 : HiggsVec →L[] (Fin 2 → ) where
toFun x := x
2024-08-20 14:04:58 +02:00
map_add' x y := rfl
map_smul' a x := rfl
2024-07-10 11:34:34 -04:00
/-- The map `toFin2` is smooth. -/
lemma smooth_toFin2 : Smooth 𝓘(, HiggsVec) 𝓘(, Fin 2 → ) toFin2 :=
ContinuousLinearMap.smooth toFin2
/-- An orthonormal basis of `HiggsVec`. -/
def orthonormBasis : OrthonormalBasis (Fin 2) HiggsVec :=
EuclideanSpace.basisFun (Fin 2)
2024-09-06 05:43:52 -04:00
/-- Generating a Higgs vector from a real number, such that the norm-squared of that Higgs vector
is the given real number. -/
def ofReal (a : ) : HiggsVec :=
![Real.sqrt a, 0]
@[simp]
lemma ofReal_normSq {a : } (ha : 0 ≤ a) : ‖ofReal a‖ ^ 2 = a := by
simp only [ofReal]
rw [PiLp.norm_sq_eq_of_L2]
rw [@Fin.sum_univ_two]
simp only [Fin.isValue, cons_val_zero, norm_real, Real.norm_eq_abs, _root_.sq_abs, cons_val_one,
head_cons, norm_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, add_zero]
exact Real.sq_sqrt ha
2024-07-10 11:34:34 -04:00
end HiggsVec
2024-07-09 09:27:28 -04:00
/-!
2024-07-10 11:34:34 -04:00
## Definition of the Higgs Field
2024-07-09 09:27:28 -04:00
2024-07-10 11:34:34 -04:00
We also define the Higgs bundle.
2024-07-09 09:27:28 -04:00
-/
2024-05-09 15:09:14 -04:00
2024-07-09 19:22:16 -04:00
/-! TODO: Make `HiggsBundle` an associated bundle. -/
/-- The trivial vector bundle 𝓡² × ℂ². -/
abbrev HiggsBundle := Bundle.Trivial SpaceTime HiggsVec
2024-05-09 15:09:14 -04:00
2024-11-11 11:58:48 +00:00
/-- The instance of a smooth vector bundle with total space `HiggsBundle` and fiber `HiggsVec`. -/
instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
2024-11-02 08:50:17 +00:00
Bundle.Trivial.smoothVectorBundle HiggsVec
2024-05-09 15:09:14 -04:00
2024-07-10 11:34:34 -04:00
/-- A Higgs field is a smooth section of the Higgs bundle. -/
abbrev HiggsField : Type := SmoothSection SpaceTime.asSmoothManifold HiggsVec HiggsBundle
2024-05-09 15:09:14 -04:00
2024-07-10 11:34:34 -04:00
/-- Given a vector in `HiggsVec` the constant Higgs field with value equal to that
2024-05-09 15:09:14 -04:00
section. -/
2024-07-10 11:34:34 -04:00
def HiggsVec.toField (φ : HiggsVec) : HiggsField where
2024-06-08 03:47:16 +02:00
toFun := fun _ ↦ φ
2024-05-09 15:09:14 -04:00
contMDiff_toFun := by
intro x
rw [Bundle.contMDiffAt_section]
exact smoothAt_const
2024-11-26 09:33:24 +00:00
/-- 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
namespace HiggsField
2024-07-10 11:34:34 -04:00
open HiggsVec
2024-05-09 15:09:14 -04:00
2024-07-09 09:27:28 -04:00
/-!
## Relation to `HiggsVec`
-/
2024-07-10 11:34:34 -04:00
/-- Given a `HiggsField`, the corresponding map from `SpaceTime` to `HiggsVec`. -/
def toHiggsVec (φ : HiggsField) : SpaceTime → HiggsVec := φ
2024-05-09 15:09:14 -04:00
lemma toHiggsVec_smooth (φ : HiggsField) : Smooth 𝓘(, SpaceTime) 𝓘(, HiggsVec) φ.toHiggsVec := by
2024-05-09 15:09:14 -04:00
intro x0
have h1 := φ.contMDiff x0
rw [Bundle.contMDiffAt_section] at h1
exact h1
lemma toField_toHiggsVec_apply (φ : HiggsField) (x : SpaceTime) :
2024-06-08 03:56:15 +02:00
(φ.toHiggsVec x).toField x = φ x := rfl
2024-05-09 15:09:14 -04:00
2024-07-10 11:34:34 -04:00
lemma toFin2_comp_toHiggsVec (φ : HiggsField) :
toFin2 ∘ φ.toHiggsVec = φ := rfl
2024-05-09 15:09:14 -04:00
2024-07-09 09:27:28 -04:00
/-!
2024-05-09 15:09:14 -04:00
2024-07-10 11:34:34 -04:00
## Smoothness properties of components
2024-07-09 09:27:28 -04:00
-/
lemma toVec_smooth (φ : HiggsField) : Smooth 𝓘(, SpaceTime) 𝓘(, Fin 2 → ) φ :=
2024-07-10 11:34:34 -04:00
smooth_toFin2.comp φ.toHiggsVec_smooth
2024-07-09 09:27:28 -04:00
lemma apply_smooth (φ : HiggsField) :
∀ i, Smooth 𝓘(, SpaceTime) 𝓘(, ) (fun (x : SpaceTime) => (φ x i)) :=
(smooth_pi_space).mp (φ.toVec_smooth)
2024-07-10 11:34:34 -04:00
lemma apply_re_smooth (φ : HiggsField) (i : Fin 2) :
Smooth 𝓘(, SpaceTime) 𝓘(, ) (reCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
2024-07-09 09:27:28 -04:00
(ContinuousLinearMap.smooth reCLM).comp (φ.apply_smooth i)
2024-07-10 11:34:34 -04:00
lemma apply_im_smooth (φ : HiggsField) (i : Fin 2) :
Smooth 𝓘(, SpaceTime) 𝓘(, ) (imCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
2024-07-09 09:27:28 -04:00
(ContinuousLinearMap.smooth imCLM).comp (φ.apply_smooth i)
/-!
2024-07-10 11:34:34 -04:00
## Constant Higgs fields
2024-07-09 09:27:28 -04:00
-/
2024-05-09 15:09:14 -04:00
2024-07-10 11:34:34 -04:00
/-- A Higgs field is constant if it is equal for all `x` `y` in `spaceTime`. -/
def IsConst (Φ : HiggsField) : Prop := ∀ x y, Φ x = Φ y
2024-05-09 15:09:14 -04:00
2024-08-20 14:04:58 +02:00
lemma isConst_of_higgsVec (φ : HiggsVec) : φ.toField.IsConst := fun _ => congrFun rfl
2024-05-09 15:09:14 -04:00
lemma isConst_iff_of_higgsVec (Φ : HiggsField) : Φ.IsConst ↔ ∃ (φ : HiggsVec), Φ = φ.toField :=
2024-06-08 03:47:16 +02:00
Iff.intro (fun h ↦ ⟨Φ 0, by ext x y; rw [← h x 0]; rfl⟩) (fun ⟨φ, hφ⟩ x y ↦ by subst hφ; rfl)
2024-05-09 15:09:14 -04:00
2024-09-06 05:43:52 -04:00
/-- Generating a constant Higgs field from a real number, such that the norm-squared of that Higgs
vector is the given real number. -/
def ofReal (a : ) : HiggsField := (HiggsVec.ofReal a).toField
2024-09-12 11:07:17 -04:00
/-- The higgs field which is all zero. -/
def zero : HiggsField := ofReal 0
2024-09-13 10:46:30 -04:00
informal_lemma zero_is_zero_section where
2024-09-15 10:14:34 -04:00
physics :≈ "The zero Higgs field is the zero section of the Higgs bundle."
math :≈ "The HiggsField `zero` defined by `ofReal 0`
2024-09-13 10:46:30 -04:00
is the constant zero-section of the bundle `HiggsBundle`."
2024-09-16 10:07:40 -04:00
deps :≈ [`StandardModel.HiggsField.zero]
2024-09-13 10:46:30 -04:00
end HiggsField
2024-07-10 11:34:34 -04:00
2024-05-09 15:09:14 -04:00
end
end StandardModel