refactor: Higgs physics
This commit is contained in:
parent
9eff5dc9bf
commit
9e88549bbe
7 changed files with 605 additions and 521 deletions
|
@ -77,5 +77,7 @@ import HepLean.SpaceTime.MinkowskiMetric
|
|||
import HepLean.SpaceTime.SL2C.Basic
|
||||
import HepLean.StandardModel.Basic
|
||||
import HepLean.StandardModel.HiggsBoson.Basic
|
||||
import HepLean.StandardModel.HiggsBoson.TargetSpace
|
||||
import HepLean.StandardModel.HiggsBoson.GaugeAction
|
||||
import HepLean.StandardModel.HiggsBoson.PointwiseInnerProd
|
||||
import HepLean.StandardModel.HiggsBoson.Potential
|
||||
import HepLean.StandardModel.Representations
|
||||
|
|
|
@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.StandardModel.HiggsBoson.Basic
|
||||
import HepLean.StandardModel.HiggsBoson.Potential
|
||||
/-!
|
||||
|
||||
# The Two Higgs Doublet Model
|
||||
|
@ -61,7 +61,7 @@ lemma swap_fields :
|
|||
|
||||
/-- If `Φ₂` is zero the potential reduces to the Higgs potential on `Φ₁`. -/
|
||||
lemma right_zero : potential Φ1 0 m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃ 𝓵₄ m₁₂2 𝓵₅ 𝓵₆ 𝓵₇ =
|
||||
StandardModel.HiggsField.potential Φ1 (- m₁₁2) (𝓵₁/2) := by
|
||||
StandardModel.HiggsField.potential (- m₁₁2) (𝓵₁/2) Φ1 := by
|
||||
funext x
|
||||
simp only [potential, normSq, ContMDiffSection.coe_zero, Pi.zero_apply, norm_zero, ne_eq,
|
||||
OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, mul_zero, add_zero, innerProd_right_zero,
|
||||
|
@ -73,7 +73,7 @@ lemma right_zero : potential Φ1 0 m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃
|
|||
|
||||
/-- If `Φ₁` is zero the potential reduces to the Higgs potential on `Φ₂`. -/
|
||||
lemma left_zero : potential 0 Φ2 m₁₁2 m₂₂2 𝓵₁ 𝓵₂ 𝓵₃ 𝓵₄ m₁₂2 𝓵₅ 𝓵₆ 𝓵₇ =
|
||||
StandardModel.HiggsField.potential Φ2 (- m₂₂2) (𝓵₂/2) := by
|
||||
StandardModel.HiggsField.potential (- m₂₂2) (𝓵₂/2) Φ2 := by
|
||||
rw [swap_fields, right_zero]
|
||||
|
||||
/-!
|
||||
|
|
|
@ -5,7 +5,7 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import HepLean.SpaceTime.Basic
|
||||
import HepLean.StandardModel.Basic
|
||||
import HepLean.StandardModel.HiggsBoson.TargetSpace
|
||||
import HepLean.StandardModel.Representations
|
||||
import Mathlib.Data.Complex.Exponential
|
||||
import Mathlib.Tactic.Polyrith
|
||||
import Mathlib.Geometry.Manifold.VectorBundle.Basic
|
||||
|
@ -17,16 +17,17 @@ import Mathlib.Analysis.InnerProductSpace.Adjoint
|
|||
import Mathlib.Geometry.Manifold.ContMDiff.Product
|
||||
import Mathlib.Algebra.QuadraticDiscriminant
|
||||
/-!
|
||||
|
||||
# The Higgs field
|
||||
|
||||
This file defines the basic properties for the higgs field in the standard model.
|
||||
This file defines the basic properties for the Higgs field in the standard model.
|
||||
|
||||
## References
|
||||
|
||||
- We use conventions given in: [Review of Particle Physics, PDG][ParticleDataGroup:2018ovx]
|
||||
|
||||
-/
|
||||
universe v u
|
||||
|
||||
namespace StandardModel
|
||||
noncomputable section
|
||||
|
||||
|
@ -35,10 +36,41 @@ open Matrix
|
|||
open Complex
|
||||
open ComplexConjugate
|
||||
open SpaceTime
|
||||
|
||||
/-!
|
||||
|
||||
## Definition of the Higgs bundle
|
||||
## 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
|
||||
map_add' x y := by simp
|
||||
map_smul' a x := by simp
|
||||
|
||||
/-- 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) ℂ
|
||||
|
||||
end HiggsVec
|
||||
|
||||
/-!
|
||||
|
||||
## Definition of the Higgs Field
|
||||
|
||||
We also define the Higgs bundle.
|
||||
-/
|
||||
|
||||
/-! TODO: Make `HiggsBundle` an associated bundle. -/
|
||||
|
@ -48,15 +80,15 @@ abbrev HiggsBundle := Bundle.Trivial SpaceTime HiggsVec
|
|||
instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
|
||||
Bundle.Trivial.smoothVectorBundle HiggsVec 𝓘(ℝ, SpaceTime)
|
||||
|
||||
/-- A higgs field is a smooth section of the higgs bundle. -/
|
||||
/-- A Higgs field is a smooth section of the Higgs bundle. -/
|
||||
abbrev HiggsField : Type := SmoothSection SpaceTime.asSmoothManifold HiggsVec HiggsBundle
|
||||
|
||||
instance : NormedAddCommGroup (Fin 2 → ℂ) := by
|
||||
exact Pi.normedAddCommGroup
|
||||
|
||||
/-- Given a vector `ℂ²` the constant higgs field with value equal to that
|
||||
/-- Given a vector in `HiggsVec` the constant Higgs field with value equal to that
|
||||
section. -/
|
||||
noncomputable def HiggsVec.toField (φ : HiggsVec) : HiggsField where
|
||||
def HiggsVec.toField (φ : HiggsVec) : HiggsField where
|
||||
toFun := fun _ ↦ φ
|
||||
contMDiff_toFun := by
|
||||
intro x
|
||||
|
@ -64,7 +96,7 @@ noncomputable def HiggsVec.toField (φ : HiggsVec) : HiggsField where
|
|||
exact smoothAt_const
|
||||
|
||||
namespace HiggsField
|
||||
open Complex Real
|
||||
open HiggsVec
|
||||
|
||||
/-!
|
||||
|
||||
|
@ -72,7 +104,7 @@ open Complex Real
|
|||
|
||||
-/
|
||||
|
||||
/-- Given a `higgsField`, the corresponding map from `spaceTime` to `higgsVec`. -/
|
||||
/-- Given a `HiggsField`, the corresponding map from `SpaceTime` to `HiggsVec`. -/
|
||||
def toHiggsVec (φ : HiggsField) : SpaceTime → HiggsVec := φ
|
||||
|
||||
lemma toHiggsVec_smooth (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, HiggsVec) φ.toHiggsVec := by
|
||||
|
@ -89,139 +121,37 @@ lemma toHiggsVec_smooth (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ
|
|||
lemma toField_toHiggsVec_apply (φ : HiggsField) (x : SpaceTime) :
|
||||
(φ.toHiggsVec x).toField x = φ x := rfl
|
||||
|
||||
lemma higgsVecToFin2ℂ_toHiggsVec (φ : HiggsField) :
|
||||
higgsVecToFin2ℂ ∘ φ.toHiggsVec = φ := rfl
|
||||
lemma toFin2ℂ_comp_toHiggsVec (φ : HiggsField) :
|
||||
toFin2ℂ ∘ φ.toHiggsVec = φ := rfl
|
||||
|
||||
/-!
|
||||
|
||||
## The inner product and norm of Higgs fields
|
||||
|
||||
-/
|
||||
|
||||
/-- Given two `HiggsField`, the map `spaceTime → ℂ` obtained by taking their inner product. -/
|
||||
def innerProd (φ1 φ2 : HiggsField) : SpaceTime → ℂ := fun x => ⟪φ1 x, φ2 x⟫_ℂ
|
||||
|
||||
/-- Notation for the inner product of two Higgs fields. -/
|
||||
scoped[StandardModel.HiggsField] notation "⟪" φ1 "," φ2 "⟫_H" => innerProd φ1 φ2
|
||||
|
||||
@[simp]
|
||||
lemma innerProd_left_zero (φ : HiggsField) : ⟪0, φ⟫_H = 0 := by
|
||||
funext x
|
||||
simp [innerProd]
|
||||
|
||||
@[simp]
|
||||
lemma innerProd_right_zero (φ : HiggsField) : ⟪φ, 0⟫_H = 0 := by
|
||||
funext x
|
||||
simp [innerProd]
|
||||
|
||||
lemma innerProd_expand (φ1 φ2 : HiggsField) :
|
||||
⟪φ1, φ2⟫_H = fun x => (conj (φ1 x 0) * (φ2 x 0) + conj (φ1 x 1) * (φ2 x 1) ) := by
|
||||
funext x
|
||||
simp only [innerProd, PiLp.inner_apply, RCLike.inner_apply, Fin.sum_univ_two]
|
||||
|
||||
/-- Given a `higgsField`, the map `spaceTime → ℝ` obtained by taking the square norm of the
|
||||
higgs vector. -/
|
||||
@[simp]
|
||||
def normSq (φ : HiggsField) : SpaceTime → ℝ := fun x => ( ‖φ x‖ ^ 2)
|
||||
|
||||
/-- Notation for the norm squared of a Higgs field. -/
|
||||
scoped[StandardModel.HiggsField] notation "‖" φ1 "‖_H ^ 2" => normSq φ1
|
||||
|
||||
lemma innerProd_self_eq_normSq (φ : HiggsField) (x : SpaceTime) :
|
||||
⟪φ, φ⟫_H x = ‖φ‖_H ^ 2 x := by
|
||||
erw [normSq, @PiLp.norm_sq_eq_of_L2, Fin.sum_univ_two]
|
||||
simp only [ ofReal_add, ofReal_pow, innerProd, PiLp.inner_apply,
|
||||
RCLike.inner_apply, Fin.sum_univ_two, conj_mul']
|
||||
|
||||
lemma normSq_eq_innerProd_self (φ : HiggsField) (x : SpaceTime) :
|
||||
‖φ x‖ ^ 2 = (⟪φ, φ⟫_H x).re := by
|
||||
rw [innerProd_self_eq_normSq]
|
||||
rfl
|
||||
|
||||
lemma toHiggsVec_norm (φ : HiggsField) (x : SpaceTime) :
|
||||
‖φ x‖ = ‖φ.toHiggsVec x‖ := rfl
|
||||
|
||||
lemma normSq_expand (φ : HiggsField) :
|
||||
φ.normSq = fun x => (conj (φ x 0) * (φ x 0) + conj (φ x 1) * (φ x 1) ).re := by
|
||||
funext x
|
||||
simp [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add, @norm_sq_eq_inner ℂ]
|
||||
|
||||
lemma normSq_nonneg (φ : HiggsField) (x : SpaceTime) : 0 ≤ φ.normSq x := by
|
||||
simp [normSq, ge_iff_le, norm_nonneg, pow_nonneg]
|
||||
|
||||
lemma normSq_zero (φ : HiggsField) (x : SpaceTime) : φ.normSq x = 0 ↔ φ x = 0 := by
|
||||
simp [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero]
|
||||
|
||||
/-!
|
||||
|
||||
## The Higgs potential
|
||||
|
||||
-/
|
||||
|
||||
/-- The Higgs potential of the form `- μ² * |φ|² + λ * |φ|⁴`. -/
|
||||
@[simp]
|
||||
def potential (φ : HiggsField) (μSq lambda : ℝ ) (x : SpaceTime) : ℝ :=
|
||||
- μSq * φ.normSq x + lambda * φ.normSq x * φ.normSq x
|
||||
|
||||
lemma potential_apply (φ : HiggsField) (μSq lambda : ℝ) (x : SpaceTime) :
|
||||
(φ.potential μSq lambda) x = HiggsVec.potential μSq lambda (φ.toHiggsVec x) := by
|
||||
simp [HiggsVec.potential, toHiggsVec_norm]
|
||||
ring
|
||||
|
||||
/-!
|
||||
|
||||
## Smoothness
|
||||
## Smoothness properties of components
|
||||
|
||||
-/
|
||||
|
||||
lemma toVec_smooth (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, Fin 2 → ℂ) φ :=
|
||||
smooth_higgsVecToFin2ℂ.comp φ.toHiggsVec_smooth
|
||||
smooth_toFin2ℂ.comp φ.toHiggsVec_smooth
|
||||
|
||||
lemma apply_smooth (φ : HiggsField) :
|
||||
∀ i, Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℂ) (fun (x : SpaceTime) => (φ x i)) :=
|
||||
(smooth_pi_space).mp (φ.toVec_smooth)
|
||||
|
||||
lemma apply_re_smooth (φ : HiggsField) (i : Fin 2):
|
||||
lemma apply_re_smooth (φ : HiggsField) (i : Fin 2) :
|
||||
Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (reCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
|
||||
(ContinuousLinearMap.smooth reCLM).comp (φ.apply_smooth i)
|
||||
|
||||
lemma apply_im_smooth (φ : HiggsField) (i : Fin 2):
|
||||
lemma apply_im_smooth (φ : HiggsField) (i : Fin 2) :
|
||||
Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (imCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
|
||||
(ContinuousLinearMap.smooth imCLM).comp (φ.apply_smooth i)
|
||||
|
||||
lemma normSq_smooth (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) φ.normSq := by
|
||||
rw [normSq_expand]
|
||||
refine Smooth.add ?_ ?_
|
||||
simp only [mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add]
|
||||
refine Smooth.add ?_ ?_
|
||||
refine Smooth.smul ?_ ?_
|
||||
exact φ.apply_re_smooth 0
|
||||
exact φ.apply_re_smooth 0
|
||||
refine Smooth.smul ?_ ?_
|
||||
exact φ.apply_im_smooth 0
|
||||
exact φ.apply_im_smooth 0
|
||||
simp only [mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add]
|
||||
refine Smooth.add ?_ ?_
|
||||
refine Smooth.smul ?_ ?_
|
||||
exact φ.apply_re_smooth 1
|
||||
exact φ.apply_re_smooth 1
|
||||
refine Smooth.smul ?_ ?_
|
||||
exact φ.apply_im_smooth 1
|
||||
exact φ.apply_im_smooth 1
|
||||
|
||||
lemma potential_smooth (φ : HiggsField) (μSq lambda : ℝ) :
|
||||
Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (fun x => φ.potential μSq lambda x) := by
|
||||
simp only [potential, normSq, neg_mul]
|
||||
exact (smooth_const.smul φ.normSq_smooth).neg.add
|
||||
((smooth_const.smul φ.normSq_smooth).smul φ.normSq_smooth)
|
||||
|
||||
/-!
|
||||
|
||||
## Constant higgs fields
|
||||
## Constant Higgs fields
|
||||
|
||||
-/
|
||||
|
||||
/-- A higgs field is constant if it is equal for all `x` `y` in `spaceTime`. -/
|
||||
/-- A Higgs field is constant if it is equal for all `x` `y` in `spaceTime`. -/
|
||||
def IsConst (Φ : HiggsField) : Prop := ∀ x y, Φ x = Φ y
|
||||
|
||||
lemma isConst_of_higgsVec (φ : HiggsVec) : φ.toField.IsConst := by
|
||||
|
@ -231,17 +161,7 @@ lemma isConst_of_higgsVec (φ : HiggsVec) : φ.toField.IsConst := by
|
|||
lemma isConst_iff_of_higgsVec (Φ : HiggsField) : Φ.IsConst ↔ ∃ (φ : HiggsVec), Φ = φ.toField :=
|
||||
Iff.intro (fun h ↦ ⟨Φ 0, by ext x y; rw [← h x 0]; rfl⟩) (fun ⟨φ, hφ⟩ x y ↦ by subst hφ; rfl)
|
||||
|
||||
lemma normSq_of_higgsVec (φ : HiggsVec) : φ.toField.normSq = fun x => (norm φ) ^ 2 := by
|
||||
funext x
|
||||
simp [normSq, HiggsVec.toField]
|
||||
|
||||
lemma potential_of_higgsVec (φ : HiggsVec) (μSq lambda : ℝ) :
|
||||
φ.toField.potential μSq lambda = fun _ => HiggsVec.potential μSq lambda φ := by
|
||||
simp [HiggsVec.potential]
|
||||
unfold potential
|
||||
rw [normSq_of_higgsVec]
|
||||
ring_nf
|
||||
|
||||
end HiggsField
|
||||
|
||||
end
|
||||
end StandardModel
|
||||
|
|
169
HepLean/StandardModel/HiggsBoson/GaugeAction.lean
Normal file
169
HepLean/StandardModel/HiggsBoson/GaugeAction.lean
Normal file
|
@ -0,0 +1,169 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.StandardModel.HiggsBoson.Basic
|
||||
/-!
|
||||
|
||||
# The action of the gauge group on the Higgs field
|
||||
|
||||
-/
|
||||
/-! TODO: Currently this only contains the action of the global gauge group. Generalize. -/
|
||||
noncomputable section
|
||||
|
||||
namespace StandardModel
|
||||
|
||||
namespace HiggsVec
|
||||
open Manifold
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
|
||||
/-!
|
||||
|
||||
## The representation of the gauge group on the Higgs vector space
|
||||
|
||||
-/
|
||||
|
||||
/-- The Higgs representation as a homomorphism from the gauge group to unitary `2×2`-matrices. -/
|
||||
@[simps!]
|
||||
noncomputable def higgsRepUnitary : GaugeGroup →* unitaryGroup (Fin 2) ℂ where
|
||||
toFun g := repU1 g.2.2 * fundamentalSU2 g.2.1
|
||||
map_mul' := by
|
||||
intro ⟨_, a2, a3⟩ ⟨_, b2, b3⟩
|
||||
change repU1 (a3 * b3) * fundamentalSU2 (a2 * b2) = _
|
||||
rw [repU1.map_mul, fundamentalSU2.map_mul, mul_assoc, mul_assoc,
|
||||
← mul_assoc (repU1 b3) _ _, repU1_fundamentalSU2_commute]
|
||||
repeat rw [mul_assoc]
|
||||
map_one' := by simp
|
||||
|
||||
/-- Takes in a `2×2`-matrix and returns a linear map of `higgsVec`. -/
|
||||
noncomputable def matrixToLin : Matrix (Fin 2) (Fin 2) ℂ →* (HiggsVec →L[ℂ] HiggsVec) where
|
||||
toFun g := LinearMap.toContinuousLinearMap
|
||||
$ Matrix.toLin orthonormBasis.toBasis orthonormBasis.toBasis g
|
||||
map_mul' g h := ContinuousLinearMap.coe_inj.mp $
|
||||
Matrix.toLin_mul orthonormBasis.toBasis orthonormBasis.toBasis orthonormBasis.toBasis g h
|
||||
map_one' := ContinuousLinearMap.coe_inj.mp $ Matrix.toLin_one orthonormBasis.toBasis
|
||||
|
||||
lemma matrixToLin_star (g : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
matrixToLin (star g) = star (matrixToLin g) :=
|
||||
ContinuousLinearMap.coe_inj.mp $ Matrix.toLin_conjTranspose orthonormBasis orthonormBasis g
|
||||
|
||||
lemma matrixToLin_unitary (g : unitaryGroup (Fin 2) ℂ) :
|
||||
matrixToLin g ∈ unitary (HiggsVec →L[ℂ] HiggsVec) := by
|
||||
rw [@unitary.mem_iff, ← matrixToLin_star, ← matrixToLin.map_mul, ← matrixToLin.map_mul,
|
||||
mem_unitaryGroup_iff.mp g.prop, mem_unitaryGroup_iff'.mp g.prop, matrixToLin.map_one]
|
||||
simp
|
||||
|
||||
/-- The natural homomorphism from unitary `2×2` complex matrices to unitary transformations
|
||||
of `higgsVec`. -/
|
||||
noncomputable def unitaryToLin : unitaryGroup (Fin 2) ℂ →* unitary (HiggsVec →L[ℂ] HiggsVec) where
|
||||
toFun g := ⟨matrixToLin g, matrixToLin_unitary g⟩
|
||||
map_mul' g h := by simp
|
||||
map_one' := by simp
|
||||
|
||||
/-- The inclusion of unitary transformations on `higgsVec` into all linear transformations. -/
|
||||
@[simps!]
|
||||
def unitToLinear : unitary (HiggsVec →L[ℂ] HiggsVec) →* HiggsVec →ₗ[ℂ] HiggsVec :=
|
||||
DistribMulAction.toModuleEnd ℂ HiggsVec
|
||||
|
||||
/-- The representation of the gauge group acting on `higgsVec`. -/
|
||||
@[simps!]
|
||||
def rep : Representation ℂ GaugeGroup HiggsVec :=
|
||||
unitToLinear.comp (unitaryToLin.comp higgsRepUnitary)
|
||||
|
||||
lemma higgsRepUnitary_mul (g : GaugeGroup) (φ : HiggsVec) :
|
||||
(higgsRepUnitary g).1 *ᵥ φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) := by
|
||||
simp [higgsRepUnitary_apply_coe, smul_mulVec_assoc]
|
||||
|
||||
lemma rep_apply (g : GaugeGroup) (φ : HiggsVec) : rep g φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) :=
|
||||
higgsRepUnitary_mul g φ
|
||||
|
||||
/-!
|
||||
|
||||
# Gauge freedom
|
||||
|
||||
-/
|
||||
|
||||
/-- Given a Higgs vector, a rotation matrix which puts the first component of the
|
||||
vector to zero, and the second component to a real -/
|
||||
def rotateMatrix (φ : HiggsVec) : Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
![![φ 1 /‖φ‖ , - φ 0 /‖φ‖], ![conj (φ 0) / ‖φ‖ , conj (φ 1) / ‖φ‖] ]
|
||||
|
||||
lemma rotateMatrix_star (φ : HiggsVec) :
|
||||
star φ.rotateMatrix =
|
||||
![![conj (φ 1) /‖φ‖ , φ 0 /‖φ‖], ![- conj (φ 0) / ‖φ‖ , φ 1 / ‖φ‖] ] := by
|
||||
simp_rw [star, rotateMatrix, conjTranspose]
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;> simp [conj_ofReal]
|
||||
|
||||
lemma rotateMatrix_det {φ : HiggsVec} (hφ : φ ≠ 0) : (rotateMatrix φ).det = 1 := by
|
||||
have h1 : (‖φ‖ : ℂ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
|
||||
field_simp [rotateMatrix, det_fin_two]
|
||||
rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm]
|
||||
|
||||
lemma rotateMatrix_unitary {φ : HiggsVec} (hφ : φ ≠ 0) :
|
||||
(rotateMatrix φ) ∈ unitaryGroup (Fin 2) ℂ := by
|
||||
rw [mem_unitaryGroup_iff', rotateMatrix_star, rotateMatrix]
|
||||
erw [mul_fin_two, one_fin_two]
|
||||
have : (‖φ‖ : ℂ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;> field_simp
|
||||
<;> rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
|
||||
· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm]
|
||||
· ring_nf
|
||||
· ring_nf
|
||||
· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
|
||||
|
||||
lemma rotateMatrix_specialUnitary {φ : HiggsVec} (hφ : φ ≠ 0) :
|
||||
(rotateMatrix φ) ∈ specialUnitaryGroup (Fin 2) ℂ :=
|
||||
mem_specialUnitaryGroup_iff.mpr ⟨rotateMatrix_unitary hφ, rotateMatrix_det hφ⟩
|
||||
|
||||
/-- Given a Higgs vector, an element of the gauge group which puts the first component of the
|
||||
vector to zero, and the second component to a real -/
|
||||
def rotateGuageGroup {φ : HiggsVec} (hφ : φ ≠ 0) : GaugeGroup :=
|
||||
⟨1, ⟨(rotateMatrix φ), rotateMatrix_specialUnitary hφ⟩, 1⟩
|
||||
|
||||
lemma rotateGuageGroup_apply {φ : HiggsVec} (hφ : φ ≠ 0) :
|
||||
rep (rotateGuageGroup hφ) φ = ![0, ofReal ‖φ‖] := by
|
||||
rw [rep_apply]
|
||||
simp only [rotateGuageGroup, rotateMatrix, one_pow, one_smul,
|
||||
Nat.succ_eq_add_one, Nat.reduceAdd, ofReal_eq_coe]
|
||||
ext i
|
||||
fin_cases i
|
||||
· simp only [mulVec, Fin.zero_eta, Fin.isValue, cons_val', empty_val', cons_val_fin_one,
|
||||
cons_val_zero, cons_dotProduct, vecHead, vecTail, Nat.succ_eq_add_one, Nat.reduceAdd,
|
||||
Function.comp_apply, Fin.succ_zero_eq_one, dotProduct_empty, add_zero]
|
||||
ring_nf
|
||||
· simp only [Fin.mk_one, Fin.isValue, cons_val_one, head_cons, mulVec, Fin.isValue,
|
||||
cons_val', empty_val', cons_val_fin_one, vecHead, cons_dotProduct, vecTail, Nat.succ_eq_add_one,
|
||||
Nat.reduceAdd, Function.comp_apply, Fin.succ_zero_eq_one, dotProduct_empty, add_zero]
|
||||
have : (‖φ‖ : ℂ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
|
||||
field_simp
|
||||
rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
|
||||
|
||||
theorem rotate_fst_zero_snd_real (φ : HiggsVec) :
|
||||
∃ (g : GaugeGroup), rep g φ = ![0, ofReal ‖φ‖] := by
|
||||
by_cases h : φ = 0
|
||||
· use ⟨1, 1, 1⟩
|
||||
simp [h]
|
||||
ext i
|
||||
fin_cases i <;> rfl
|
||||
· use rotateGuageGroup h
|
||||
exact rotateGuageGroup_apply h
|
||||
|
||||
end HiggsVec
|
||||
|
||||
/-! TODO: Define the global gauge action on HiggsField. -/
|
||||
/-! TODO: Prove `⟪φ1, φ2⟫_H` invariant under the global gauge action. (norm_map_of_mem_unitary) -/
|
||||
/-! TODO: Prove invariance of potential under global gauge action. -/
|
||||
|
||||
end StandardModel
|
||||
end
|
154
HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean
Normal file
154
HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean
Normal file
|
@ -0,0 +1,154 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.StandardModel.HiggsBoson.Basic
|
||||
import LeanCopilot
|
||||
/-!
|
||||
# The pointwise inner product
|
||||
|
||||
We define the pointwise inner product of two Higgs fields.
|
||||
|
||||
The notation for the inner product is `⟪φ, ψ⟫_H`.
|
||||
|
||||
We also define the pointwise norm squared of a Higgs field `∥φ∥_H ^ 2`.
|
||||
|
||||
-/
|
||||
|
||||
noncomputable section
|
||||
|
||||
namespace StandardModel
|
||||
|
||||
namespace HiggsField
|
||||
|
||||
open Manifold
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
open SpaceTime
|
||||
|
||||
/-!
|
||||
|
||||
## The pointwise inner product
|
||||
|
||||
-/
|
||||
|
||||
/-- Given two `HiggsField`, the map `SpaceTime → ℂ` obtained by taking their pointwise
|
||||
inner product. -/
|
||||
def innerProd (φ1 φ2 : HiggsField) : SpaceTime → ℂ := fun x => ⟪φ1 x, φ2 x⟫_ℂ
|
||||
|
||||
/-- Notation for the inner product of two Higgs fields. -/
|
||||
scoped[StandardModel.HiggsField] notation "⟪" φ1 "," φ2 "⟫_H" => innerProd φ1 φ2
|
||||
|
||||
/-!
|
||||
|
||||
## Properties of the inner product
|
||||
|
||||
-/
|
||||
|
||||
@[simp]
|
||||
lemma innerProd_left_zero (φ : HiggsField) : ⟪0, φ⟫_H = 0 := by
|
||||
funext x
|
||||
simp [innerProd]
|
||||
|
||||
@[simp]
|
||||
lemma innerProd_right_zero (φ : HiggsField) : ⟪φ, 0⟫_H = 0 := by
|
||||
funext x
|
||||
simp [innerProd]
|
||||
example (x : ℝ): RCLike.ofReal x = ofReal' x := by
|
||||
rfl
|
||||
lemma innerProd_expand (φ1 φ2 : HiggsField) :
|
||||
⟪φ1, φ2⟫_H = fun x => equivRealProdCLM.symm (((φ1 x 0).re * (φ2 x 0).re
|
||||
+ (φ1 x 1).re * (φ2 x 1).re+ (φ1 x 0).im * (φ2 x 0).im + (φ1 x 1).im * (φ2 x 1).im),
|
||||
((φ1 x 0).re * (φ2 x 0).im + (φ1 x 1).re * (φ2 x 1).im
|
||||
- (φ1 x 0).im * (φ2 x 0).re - (φ1 x 1).im * (φ2 x 1).re)) := by
|
||||
funext x
|
||||
simp only [innerProd, PiLp.inner_apply, RCLike.inner_apply, Fin.sum_univ_two,
|
||||
equivRealProdCLM_symm_apply, ofReal_add, ofReal_mul, ofReal_sub]
|
||||
rw [RCLike.conj_eq_re_sub_im, RCLike.conj_eq_re_sub_im]
|
||||
nth_rewrite 1 [← RCLike.re_add_im (φ2 x 0)]
|
||||
nth_rewrite 1 [← RCLike.re_add_im (φ2 x 1)]
|
||||
ring_nf
|
||||
repeat rw [show RCLike.ofReal _ = ofReal' _ by rfl]
|
||||
simp only [Algebra.id.map_eq_id, RCLike.re_to_complex, RingHom.id_apply, RCLike.I_to_complex,
|
||||
RCLike.im_to_complex, I_sq, mul_neg, mul_one, neg_mul, sub_neg_eq_add, one_mul]
|
||||
ring
|
||||
|
||||
lemma smooth_innerProd (φ1 φ2 : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℂ) ⟪φ1, φ2⟫_H := by
|
||||
rw [innerProd_expand]
|
||||
exact (ContinuousLinearMap.smooth (equivRealProdCLM.symm : ℝ × ℝ →L[ℝ] ℂ)).comp $
|
||||
(((((φ1.apply_re_smooth 0).smul (φ2.apply_re_smooth 0)).add
|
||||
((φ1.apply_re_smooth 1).smul (φ2.apply_re_smooth 1))).add
|
||||
((φ1.apply_im_smooth 0).smul (φ2.apply_im_smooth 0))).add
|
||||
((φ1.apply_im_smooth 1).smul (φ2.apply_im_smooth 1))).prod_mk_space $
|
||||
((((φ1.apply_re_smooth 0).smul (φ2.apply_im_smooth 0)).add
|
||||
((φ1.apply_re_smooth 1).smul (φ2.apply_im_smooth 1))).sub
|
||||
((φ1.apply_im_smooth 0).smul (φ2.apply_re_smooth 0))).sub
|
||||
((φ1.apply_im_smooth 1).smul (φ2.apply_re_smooth 1))
|
||||
|
||||
/-!
|
||||
|
||||
## The pointwise norm squared
|
||||
|
||||
-/
|
||||
|
||||
/-- Given a `HiggsField`, the map `SpaceTime → ℝ` obtained by taking the square norm of the
|
||||
pointwise Higgs vector. -/
|
||||
@[simp]
|
||||
def normSq (φ : HiggsField) : SpaceTime → ℝ := fun x => ( ‖φ x‖ ^ 2)
|
||||
|
||||
/-- Notation for the norm squared of a Higgs field. -/
|
||||
scoped[StandardModel.HiggsField] notation "‖" φ1 "‖_H ^ 2" => normSq φ1
|
||||
|
||||
/-!
|
||||
|
||||
## Relation between inner prod and norm squared
|
||||
|
||||
-/
|
||||
|
||||
lemma innerProd_self_eq_normSq (φ : HiggsField) (x : SpaceTime) :
|
||||
⟪φ, φ⟫_H x = ‖φ‖_H ^ 2 x := by
|
||||
erw [normSq, @PiLp.norm_sq_eq_of_L2, Fin.sum_univ_two]
|
||||
simp only [innerProd, PiLp.inner_apply, RCLike.inner_apply, conj_mul', norm_eq_abs,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_pow]
|
||||
|
||||
lemma normSq_eq_innerProd_self (φ : HiggsField) (x : SpaceTime) :
|
||||
‖φ x‖ ^ 2 = (⟪φ, φ⟫_H x).re := by
|
||||
rw [innerProd_self_eq_normSq]
|
||||
rfl
|
||||
|
||||
/-!
|
||||
|
||||
# Properties of the norm squared
|
||||
|
||||
-/
|
||||
|
||||
lemma toHiggsVec_norm (φ : HiggsField) (x : SpaceTime) :
|
||||
‖φ x‖ = ‖φ.toHiggsVec x‖ := rfl
|
||||
|
||||
lemma normSq_expand (φ : HiggsField) :
|
||||
φ.normSq = fun x => (conj (φ x 0) * (φ x 0) + conj (φ x 1) * (φ x 1) ).re := by
|
||||
funext x
|
||||
simp [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add, @norm_sq_eq_inner ℂ]
|
||||
|
||||
lemma normSq_nonneg (φ : HiggsField) (x : SpaceTime) : 0 ≤ φ.normSq x := by
|
||||
simp [normSq, ge_iff_le, norm_nonneg, pow_nonneg]
|
||||
|
||||
lemma normSq_zero (φ : HiggsField) (x : SpaceTime) : φ.normSq x = 0 ↔ φ x = 0 := by
|
||||
simp [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero]
|
||||
|
||||
lemma normSq_smooth (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) φ.normSq := by
|
||||
rw [normSq_expand]
|
||||
refine Smooth.add ?_ ?_
|
||||
simp only [mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add]
|
||||
exact ((φ.apply_re_smooth 0).smul (φ.apply_re_smooth 0)).add $
|
||||
(φ.apply_im_smooth 0).smul (φ.apply_im_smooth 0)
|
||||
simp only [mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add]
|
||||
exact ((φ.apply_re_smooth 1).smul (φ.apply_re_smooth 1)).add $
|
||||
(φ.apply_im_smooth 1).smul (φ.apply_im_smooth 1)
|
||||
|
||||
end HiggsField
|
||||
|
||||
end StandardModel
|
||||
end
|
226
HepLean/StandardModel/HiggsBoson/Potential.lean
Normal file
226
HepLean/StandardModel/HiggsBoson/Potential.lean
Normal file
|
@ -0,0 +1,226 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.StandardModel.HiggsBoson.PointwiseInnerProd
|
||||
/-!
|
||||
# The potential of the Higgs field
|
||||
|
||||
We define the potential of the Higgs field.
|
||||
|
||||
We show that the potential is a smooth function on spacetime.
|
||||
|
||||
-/
|
||||
|
||||
noncomputable section
|
||||
|
||||
namespace StandardModel
|
||||
|
||||
namespace HiggsField
|
||||
|
||||
open Manifold
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
open SpaceTime
|
||||
|
||||
/-!
|
||||
|
||||
## The Higgs potential
|
||||
|
||||
-/
|
||||
|
||||
/-- The Higgs potential of the form `- μ² * |φ|² + 𝓵 * |φ|⁴`. -/
|
||||
@[simp]
|
||||
def potential (μ2 𝓵 : ℝ ) (φ : HiggsField) (x : SpaceTime) : ℝ :=
|
||||
- μ2 * ‖φ‖_H ^ 2 x + 𝓵 * ‖φ‖_H ^ 2 x * ‖φ‖_H ^ 2 x
|
||||
|
||||
/-!
|
||||
|
||||
## Smoothness of the potential
|
||||
|
||||
-/
|
||||
|
||||
lemma potential_smooth (μSq lambda : ℝ) (φ : HiggsField) :
|
||||
Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (fun x => φ.potential μSq lambda x) := by
|
||||
simp only [potential, normSq, neg_mul]
|
||||
exact (smooth_const.smul φ.normSq_smooth).neg.add
|
||||
((smooth_const.smul φ.normSq_smooth).smul φ.normSq_smooth)
|
||||
|
||||
namespace potential
|
||||
section lowerBound
|
||||
/-!
|
||||
|
||||
## Lower bound on potential
|
||||
|
||||
-/
|
||||
|
||||
variable {𝓵 : ℝ}
|
||||
variable (μ2 : ℝ)
|
||||
variable (h𝓵 : 0 < 𝓵)
|
||||
|
||||
/-- The second term of the potential is non-negative. -/
|
||||
lemma snd_term_nonneg (φ : HiggsField) (x : SpaceTime) :
|
||||
0 ≤ 𝓵 * ‖φ‖_H ^ 2 x * ‖φ‖_H ^ 2 x := by
|
||||
rw [mul_nonneg_iff]
|
||||
apply Or.inl
|
||||
simp_all only [normSq, gt_iff_lt, mul_nonneg_iff_of_pos_left, ge_iff_le, norm_nonneg, pow_nonneg,
|
||||
and_self]
|
||||
|
||||
lemma as_quad (μ2 𝓵 : ℝ) (φ : HiggsField) (x : SpaceTime) :
|
||||
𝓵 * ‖φ‖_H ^ 2 x * ‖φ‖_H ^ 2 x + (- μ2 ) * ‖φ‖_H ^ 2 x + (- potential μ2 𝓵 φ x) = 0 := by
|
||||
simp only [normSq, neg_mul, potential, neg_add_rev, neg_neg]
|
||||
ring
|
||||
|
||||
/-- The discriminant of the quadratic formed by the potential is non-negative. -/
|
||||
lemma discrim_nonneg (φ : HiggsField) (x : SpaceTime) :
|
||||
0 ≤ discrim 𝓵 (- μ2) (- potential μ2 𝓵 φ x) := by
|
||||
have h1 := as_quad μ2 𝓵 φ x
|
||||
rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1
|
||||
· simp only [h1, ne_eq, div_eq_zero_iff, OfNat.ofNat_ne_zero, or_false]
|
||||
exact sq_nonneg (2 * 𝓵 * ‖φ‖_H ^ 2 x+ - μ2)
|
||||
· exact ne_of_gt h𝓵
|
||||
|
||||
lemma eq_zero_at (φ : HiggsField) (x : SpaceTime)
|
||||
(hV : potential μ2 𝓵 φ x = 0) : φ x = 0 ∨ ‖φ‖_H ^ 2 x = μ2 / 𝓵 := by
|
||||
have h1 := as_quad μ2 𝓵 φ x
|
||||
rw [hV] at h1
|
||||
have h2 : ‖φ‖_H ^ 2 x * (𝓵 * ‖φ‖_H ^ 2 x + - μ2) = 0 := by
|
||||
linear_combination h1
|
||||
simp at h2
|
||||
cases' h2 with h2 h2
|
||||
simp_all
|
||||
apply Or.inr
|
||||
field_simp at h2 ⊢
|
||||
ring_nf
|
||||
linear_combination h2
|
||||
|
||||
lemma eq_zero_at_of_μSq_nonpos {μ2 : ℝ} (hμ2 : μ2 ≤ 0)
|
||||
(φ : HiggsField) (x : SpaceTime) (hV : potential μ2 𝓵 φ x = 0) : φ x = 0 := by
|
||||
cases' (eq_zero_at μ2 h𝓵 φ x hV) with h1 h1
|
||||
exact h1
|
||||
by_cases hμSqZ : μ2 = 0
|
||||
simpa [hμSqZ] using h1
|
||||
refine ((?_ : ¬ 0 ≤ μ2 / 𝓵) (?_)).elim
|
||||
· simp_all [div_nonneg_iff]
|
||||
intro h
|
||||
exact lt_imp_lt_of_le_imp_le (fun _ => h) (lt_of_le_of_ne hμ2 hμSqZ)
|
||||
· rw [← h1]
|
||||
exact normSq_nonneg φ x
|
||||
|
||||
lemma bounded_below (φ : HiggsField) (x : SpaceTime) :
|
||||
- μ2 ^ 2 / (4 * 𝓵) ≤ potential μ2 𝓵 φ x := by
|
||||
have h1 := discrim_nonneg μ2 h𝓵 φ x
|
||||
simp only [discrim, even_two, Even.neg_pow, normSq, neg_mul, neg_add_rev, neg_neg] at h1
|
||||
ring_nf at h1
|
||||
rw [← neg_le_iff_add_nonneg'] at h1
|
||||
rw [show 𝓵 * potential μ2 𝓵 φ x * 4 = (4 * 𝓵) * potential μ2 𝓵 φ x by ring] at h1
|
||||
have h2 := (div_le_iff' (by simp [h𝓵] : 0 < 4 * 𝓵)).mpr h1
|
||||
ring_nf at h2 ⊢
|
||||
exact h2
|
||||
|
||||
lemma bounded_below_of_μSq_nonpos {μ2 : ℝ}
|
||||
(hμSq : μ2 ≤ 0) (φ : HiggsField) (x : SpaceTime) : 0 ≤ potential μ2 𝓵 φ x := by
|
||||
refine add_nonneg ?_ (snd_term_nonneg h𝓵 φ x)
|
||||
field_simp [mul_nonpos_iff]
|
||||
simp_all [ge_iff_le, norm_nonneg, pow_nonneg, and_self, or_true]
|
||||
|
||||
end lowerBound
|
||||
|
||||
section MinimumOfPotential
|
||||
variable {𝓵 : ℝ}
|
||||
variable (μ2 : ℝ)
|
||||
variable (h𝓵 : 0 < 𝓵)
|
||||
|
||||
/-!
|
||||
|
||||
## Minima of potential
|
||||
|
||||
-/
|
||||
|
||||
lemma discrim_eq_zero_of_eq_bound (φ : HiggsField) (x : SpaceTime)
|
||||
(hV : potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵)) :
|
||||
discrim 𝓵 (- μ2) (- potential μ2 𝓵 φ x) = 0 := by
|
||||
rw [discrim, hV]
|
||||
field_simp
|
||||
|
||||
lemma normSq_of_eq_bound (φ : HiggsField) (x : SpaceTime)
|
||||
(hV : potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵)) :
|
||||
‖φ‖_H ^ 2 x = μ2 / (2 * 𝓵) := by
|
||||
have h1 := as_quad μ2 𝓵 φ x
|
||||
rw [quadratic_eq_zero_iff_of_discrim_eq_zero _
|
||||
(discrim_eq_zero_of_eq_bound μ2 h𝓵 φ x hV)] at h1
|
||||
simp_rw [h1, neg_neg]
|
||||
exact ne_of_gt h𝓵
|
||||
|
||||
lemma eq_bound_iff (φ : HiggsField) (x : SpaceTime) :
|
||||
potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵) ↔ ‖φ‖_H ^ 2 x = μ2 / (2 * 𝓵) :=
|
||||
Iff.intro (normSq_of_eq_bound μ2 h𝓵 φ x)
|
||||
(fun h ↦ by
|
||||
rw [potential, h]
|
||||
field_simp
|
||||
ring_nf)
|
||||
|
||||
lemma eq_bound_iff_of_μSq_nonpos {μ2 : ℝ}
|
||||
(hμ2 : μ2 ≤ 0) (φ : HiggsField) (x : SpaceTime) :
|
||||
potential μ2 𝓵 φ x = 0 ↔ φ x = 0 :=
|
||||
Iff.intro (fun h ↦ eq_zero_at_of_μSq_nonpos h𝓵 hμ2 φ x h)
|
||||
(fun h ↦ by simp [potential, h])
|
||||
|
||||
lemma eq_bound_IsMinOn (φ : HiggsField) (x : SpaceTime)
|
||||
(hv : potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵)) :
|
||||
IsMinOn (fun (φ, x) => potential μ2 𝓵 φ x) Set.univ (φ, x) := by
|
||||
rw [isMinOn_univ_iff]
|
||||
simp only [normSq, neg_mul, le_neg_add_iff_add_le, ge_iff_le, hv]
|
||||
exact fun (φ', x') ↦ bounded_below μ2 h𝓵 φ' x'
|
||||
|
||||
lemma eq_bound_IsMinOn_of_μSq_nonpos {μ2 : ℝ}
|
||||
(hμ2 : μ2 ≤ 0) (φ : HiggsField) (x : SpaceTime) (hv : potential μ2 𝓵 φ x = 0) :
|
||||
IsMinOn (fun (φ, x) => potential μ2 𝓵 φ x) Set.univ (φ, x) := by
|
||||
rw [isMinOn_univ_iff]
|
||||
simp only [normSq, neg_mul, le_neg_add_iff_add_le, ge_iff_le, hv]
|
||||
exact fun (φ', x') ↦ bounded_below_of_μSq_nonpos h𝓵 hμ2 φ' x'
|
||||
|
||||
lemma bound_reached_of_μSq_nonneg {μ2 : ℝ} (hμ2 : 0 ≤ μ2) :
|
||||
∃ (φ : HiggsField) (x : SpaceTime),
|
||||
potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵) := by
|
||||
use HiggsVec.toField ![√(μ2/(2 * 𝓵)), 0], 0
|
||||
refine (eq_bound_iff μ2 h𝓵 (HiggsVec.toField ![√(μ2/(2 * 𝓵)), 0]) 0).mpr ?_
|
||||
simp only [normSq, HiggsVec.toField, ContMDiffSection.coeFn_mk, PiLp.norm_sq_eq_of_L2,
|
||||
Nat.succ_eq_add_one, Nat.reduceAdd, norm_eq_abs, Fin.sum_univ_two, Fin.isValue, cons_val_zero,
|
||||
abs_ofReal, _root_.sq_abs, cons_val_one, head_cons, map_zero, ne_eq, OfNat.ofNat_ne_zero,
|
||||
not_false_eq_true, zero_pow, add_zero]
|
||||
field_simp [mul_pow]
|
||||
|
||||
lemma IsMinOn_iff_of_μSq_nonneg {μ2 : ℝ} (hμ2 : 0 ≤ μ2) :
|
||||
IsMinOn (fun (φ, x) => potential μ2 𝓵 φ x) Set.univ (φ, x) ↔
|
||||
‖φ‖_H ^ 2 x = μ2 /(2 * 𝓵) := by
|
||||
apply Iff.intro <;> rw [← eq_bound_iff μ2 h𝓵 φ]
|
||||
· intro h
|
||||
obtain ⟨φm, xm, hφ⟩ := bound_reached_of_μSq_nonneg h𝓵 hμ2
|
||||
have hm := isMinOn_univ_iff.mp h (φm, xm)
|
||||
simp only at hm
|
||||
rw [hφ] at hm
|
||||
exact (Real.partialOrder.le_antisymm _ _ (bounded_below μ2 h𝓵 φ x) hm).symm
|
||||
· exact eq_bound_IsMinOn μ2 h𝓵 φ x
|
||||
|
||||
lemma IsMinOn_iff_of_μSq_nonpos {μ2 : ℝ} (hμ2 : μ2 ≤ 0) :
|
||||
IsMinOn (fun (φ, x) => potential μ2 𝓵 φ x) Set.univ (φ, x) ↔ φ x = 0 := by
|
||||
apply Iff.intro <;> rw [← eq_bound_iff_of_μSq_nonpos h𝓵 hμ2 φ]
|
||||
· intro h
|
||||
have h0 := isMinOn_univ_iff.mp h 0
|
||||
have h1 := bounded_below_of_μSq_nonpos h𝓵 hμ2 φ x
|
||||
simp only at h0
|
||||
rw [(eq_bound_iff_of_μSq_nonpos h𝓵 hμ2 0 0 ).mpr (by rfl)] at h0
|
||||
exact (Real.partialOrder.le_antisymm _ _ h1 h0).symm
|
||||
· exact eq_bound_IsMinOn_of_μSq_nonpos h𝓵 hμ2 φ x
|
||||
|
||||
end MinimumOfPotential
|
||||
|
||||
end potential
|
||||
|
||||
end HiggsField
|
||||
|
||||
end StandardModel
|
||||
end
|
|
@ -1,387 +0,0 @@
|
|||
/-
|
||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||||
Released under Apache 2.0 license.
|
||||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.StandardModel.Basic
|
||||
import HepLean.StandardModel.Representations
|
||||
import Mathlib.Data.Complex.Exponential
|
||||
import Mathlib.Tactic.Polyrith
|
||||
import Mathlib.Geometry.Manifold.Instances.Real
|
||||
import Mathlib.RepresentationTheory.Basic
|
||||
import Mathlib.Analysis.InnerProductSpace.Basic
|
||||
import Mathlib.Analysis.InnerProductSpace.Adjoint
|
||||
import Mathlib.Geometry.Manifold.ContMDiff.Product
|
||||
import Mathlib.Algebra.QuadraticDiscriminant
|
||||
import Mathlib.Geometry.Manifold.ContMDiff.NormedSpace
|
||||
/-!
|
||||
# The Higgs vector space
|
||||
|
||||
This file defines the target vector space of the Higgs boson, the potential on it,
|
||||
and the representation of the SM gauge group acting on it.
|
||||
|
||||
This file is a import of `SM.HiggsBoson.Basic`.
|
||||
|
||||
## References
|
||||
|
||||
- We use conventions given in: [Review of Particle Physics, PDG][ParticleDataGroup:2018ovx]
|
||||
|
||||
-/
|
||||
/-! TODO: Move potential to a seperate file, and combine with HiggsBoson.Basic. -/
|
||||
|
||||
universe v u
|
||||
namespace StandardModel
|
||||
noncomputable section
|
||||
|
||||
open Manifold
|
||||
open Matrix
|
||||
open Complex
|
||||
open ComplexConjugate
|
||||
|
||||
/-!
|
||||
|
||||
## The definition of the Higgs vector space.
|
||||
|
||||
-/
|
||||
|
||||
/-- The complex vector space in which the Higgs field takes values. -/
|
||||
abbrev HiggsVec := EuclideanSpace ℂ (Fin 2)
|
||||
|
||||
/-- The continuous linear map from the vector space `higgsVec` to `(Fin 2 → ℂ)` achieved by
|
||||
casting vectors. -/
|
||||
def higgsVecToFin2ℂ : HiggsVec →L[ℝ] (Fin 2 → ℂ) where
|
||||
toFun x := x
|
||||
map_add' x y := by simp
|
||||
map_smul' a x := by simp
|
||||
|
||||
lemma smooth_higgsVecToFin2ℂ : Smooth 𝓘(ℝ, HiggsVec) 𝓘(ℝ, Fin 2 → ℂ) higgsVecToFin2ℂ :=
|
||||
ContinuousLinearMap.smooth higgsVecToFin2ℂ
|
||||
|
||||
namespace HiggsVec
|
||||
/-- An orthonormal basis of higgsVec. -/
|
||||
noncomputable def orthonormBasis : OrthonormalBasis (Fin 2) ℂ HiggsVec :=
|
||||
EuclideanSpace.basisFun (Fin 2) ℂ
|
||||
|
||||
/-!
|
||||
|
||||
## The representation of the gauge group on the Higgs vector space
|
||||
|
||||
-/
|
||||
|
||||
/-- The Higgs representation as a homomorphism from the gauge group to unitary `2×2`-matrices. -/
|
||||
@[simps!]
|
||||
noncomputable def higgsRepUnitary : GaugeGroup →* unitaryGroup (Fin 2) ℂ where
|
||||
toFun g := repU1 g.2.2 * fundamentalSU2 g.2.1
|
||||
map_mul' := by
|
||||
intro ⟨_, a2, a3⟩ ⟨_, b2, b3⟩
|
||||
change repU1 (a3 * b3) * fundamentalSU2 (a2 * b2) = _
|
||||
rw [repU1.map_mul, fundamentalSU2.map_mul, mul_assoc, mul_assoc,
|
||||
← mul_assoc (repU1 b3) _ _, repU1_fundamentalSU2_commute]
|
||||
repeat rw [mul_assoc]
|
||||
map_one' := by simp
|
||||
|
||||
/-- Takes in a `2×2`-matrix and returns a linear map of `higgsVec`. -/
|
||||
noncomputable def matrixToLin : Matrix (Fin 2) (Fin 2) ℂ →* (HiggsVec →L[ℂ] HiggsVec) where
|
||||
toFun g := LinearMap.toContinuousLinearMap
|
||||
$ Matrix.toLin orthonormBasis.toBasis orthonormBasis.toBasis g
|
||||
map_mul' g h := ContinuousLinearMap.coe_inj.mp $
|
||||
Matrix.toLin_mul orthonormBasis.toBasis orthonormBasis.toBasis orthonormBasis.toBasis g h
|
||||
map_one' := ContinuousLinearMap.coe_inj.mp $ Matrix.toLin_one orthonormBasis.toBasis
|
||||
|
||||
lemma matrixToLin_star (g : Matrix (Fin 2) (Fin 2) ℂ) :
|
||||
matrixToLin (star g) = star (matrixToLin g) :=
|
||||
ContinuousLinearMap.coe_inj.mp $ Matrix.toLin_conjTranspose orthonormBasis orthonormBasis g
|
||||
|
||||
lemma matrixToLin_unitary (g : unitaryGroup (Fin 2) ℂ) :
|
||||
matrixToLin g ∈ unitary (HiggsVec →L[ℂ] HiggsVec) := by
|
||||
rw [@unitary.mem_iff, ← matrixToLin_star, ← matrixToLin.map_mul, ← matrixToLin.map_mul,
|
||||
mem_unitaryGroup_iff.mp g.prop, mem_unitaryGroup_iff'.mp g.prop, matrixToLin.map_one]
|
||||
simp
|
||||
|
||||
/-- The natural homomorphism from unitary `2×2` complex matrices to unitary transformations
|
||||
of `higgsVec`. -/
|
||||
noncomputable def unitaryToLin : unitaryGroup (Fin 2) ℂ →* unitary (HiggsVec →L[ℂ] HiggsVec) where
|
||||
toFun g := ⟨matrixToLin g, matrixToLin_unitary g⟩
|
||||
map_mul' g h := by simp
|
||||
map_one' := by simp
|
||||
|
||||
/-- The inclusion of unitary transformations on `higgsVec` into all linear transformations. -/
|
||||
@[simps!]
|
||||
def unitToLinear : unitary (HiggsVec →L[ℂ] HiggsVec) →* HiggsVec →ₗ[ℂ] HiggsVec :=
|
||||
DistribMulAction.toModuleEnd ℂ HiggsVec
|
||||
|
||||
/-- The representation of the gauge group acting on `higgsVec`. -/
|
||||
@[simps!]
|
||||
def rep : Representation ℂ GaugeGroup HiggsVec :=
|
||||
unitToLinear.comp (unitaryToLin.comp higgsRepUnitary)
|
||||
|
||||
lemma higgsRepUnitary_mul (g : GaugeGroup) (φ : HiggsVec) :
|
||||
(higgsRepUnitary g).1 *ᵥ φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) := by
|
||||
simp [higgsRepUnitary_apply_coe, smul_mulVec_assoc]
|
||||
|
||||
lemma rep_apply (g : GaugeGroup) (φ : HiggsVec) : rep g φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) :=
|
||||
higgsRepUnitary_mul g φ
|
||||
|
||||
section potentialDefn
|
||||
|
||||
variable (μSq lambda : ℝ)
|
||||
local notation "λ" => lambda
|
||||
|
||||
/-!
|
||||
|
||||
## The potential for a Higgs vector
|
||||
|
||||
-/
|
||||
|
||||
/-- The higgs potential for `higgsVec`, i.e. for constant higgs fields. -/
|
||||
def potential (φ : HiggsVec) : ℝ := - μSq * ‖φ‖ ^ 2 + λ * ‖φ‖ ^ 4
|
||||
|
||||
lemma potential_as_quad (φ : HiggsVec) :
|
||||
λ * ‖φ‖ ^ 2 * ‖φ‖ ^ 2 + (- μSq ) * ‖φ‖ ^ 2 + (- potential μSq (λ) φ) = 0 := by
|
||||
simp [potential]; ring
|
||||
|
||||
/-!
|
||||
|
||||
## Invariance of the potential under global gauge transformation
|
||||
|
||||
-/
|
||||
|
||||
lemma norm_invariant (g : GaugeGroup) (φ : HiggsVec) : ‖rep g φ‖ = ‖φ‖ :=
|
||||
ContinuousLinearMap.norm_map_of_mem_unitary (unitaryToLin (higgsRepUnitary g)).2 φ
|
||||
|
||||
lemma potential_invariant (φ : HiggsVec) (g : GaugeGroup) :
|
||||
potential μSq (λ) (rep g φ) = potential μSq (λ) φ := by
|
||||
simp only [potential, neg_mul, norm_invariant]
|
||||
|
||||
|
||||
end potentialDefn
|
||||
section potentialProp
|
||||
|
||||
variable {lambda : ℝ}
|
||||
variable (μSq : ℝ)
|
||||
variable (hLam : 0 < lambda)
|
||||
local notation "λ" => lambda
|
||||
|
||||
/-!
|
||||
|
||||
## Lower bound on potential
|
||||
|
||||
-/
|
||||
|
||||
lemma potential_snd_term_nonneg (φ : HiggsVec) :
|
||||
0 ≤ λ * ‖φ‖ ^ 4 := by
|
||||
rw [mul_nonneg_iff]
|
||||
apply Or.inl
|
||||
simp_all only [ge_iff_le, norm_nonneg, pow_nonneg, and_true]
|
||||
exact le_of_lt hLam
|
||||
|
||||
lemma zero_le_potential_discrim (φ : HiggsVec) :
|
||||
0 ≤ discrim (λ) (- μSq ) (- potential μSq (λ) φ) := by
|
||||
have h1 := potential_as_quad μSq (λ) φ
|
||||
rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1
|
||||
· simp only [h1, ne_eq, div_eq_zero_iff, OfNat.ofNat_ne_zero, or_false]
|
||||
exact sq_nonneg (2 * lambda * ‖φ‖ ^ 2 + -μSq)
|
||||
· exact ne_of_gt hLam
|
||||
|
||||
lemma potential_eq_zero_sol (φ : HiggsVec)
|
||||
(hV : potential μSq (λ) φ = 0) : φ = 0 ∨ ‖φ‖ ^ 2 = μSq / λ := by
|
||||
have h1 := potential_as_quad μSq (λ) φ
|
||||
rw [hV] at h1
|
||||
have h2 : ‖φ‖ ^ 2 * (lambda * ‖φ‖ ^ 2 + -μSq ) = 0 := by
|
||||
linear_combination h1
|
||||
simp at h2
|
||||
cases' h2 with h2 h2
|
||||
simp_all
|
||||
apply Or.inr
|
||||
field_simp at h2 ⊢
|
||||
ring_nf
|
||||
linear_combination h2
|
||||
|
||||
lemma potential_eq_zero_sol_of_μSq_nonpos (hμSq : μSq ≤ 0)
|
||||
(φ : HiggsVec) (hV : potential μSq (λ) φ = 0) : φ = 0 := by
|
||||
cases' (potential_eq_zero_sol μSq hLam φ hV) with h1 h1
|
||||
exact h1
|
||||
by_cases hμSqZ : μSq = 0
|
||||
simpa [hμSqZ] using h1
|
||||
refine ((?_ : ¬ 0 ≤ μSq / λ) (?_)).elim
|
||||
· simp_all [div_nonneg_iff]
|
||||
intro h
|
||||
exact lt_imp_lt_of_le_imp_le (fun _ => h) (lt_of_le_of_ne hμSq hμSqZ)
|
||||
· rw [← h1]
|
||||
exact sq_nonneg ‖φ‖
|
||||
|
||||
lemma potential_bounded_below (φ : HiggsVec) :
|
||||
- μSq ^ 2 / (4 * (λ)) ≤ potential μSq (λ) φ := by
|
||||
have h1 := zero_le_potential_discrim μSq hLam φ
|
||||
simp [discrim] at h1
|
||||
ring_nf at h1
|
||||
rw [← neg_le_iff_add_nonneg'] at h1
|
||||
have h3 : (λ) * potential μSq (λ) φ * 4 = (4 * λ) * potential μSq (λ) φ := by
|
||||
ring
|
||||
rw [h3] at h1
|
||||
have h2 := (div_le_iff' (by simp [hLam] : 0 < 4 * λ )).mpr h1
|
||||
ring_nf at h2 ⊢
|
||||
exact h2
|
||||
|
||||
lemma potential_bounded_below_of_μSq_nonpos {μSq : ℝ}
|
||||
(hμSq : μSq ≤ 0) (φ : HiggsVec) : 0 ≤ potential μSq (λ) φ := by
|
||||
refine add_nonneg ?_ (potential_snd_term_nonneg hLam φ)
|
||||
field_simp [mul_nonpos_iff]
|
||||
simp_all [ge_iff_le, norm_nonneg, pow_nonneg, and_self, or_true]
|
||||
|
||||
/-!
|
||||
|
||||
## Minimum of the potential
|
||||
|
||||
-/
|
||||
|
||||
lemma potential_eq_bound_discrim_zero (φ : HiggsVec)
|
||||
(hV : potential μSq (λ) φ = - μSq ^ 2 / (4 * λ)) :
|
||||
discrim (λ) (- μSq) (- potential μSq (λ) φ) = 0 := by
|
||||
field_simp [discrim, hV]
|
||||
|
||||
lemma potential_eq_bound_higgsVec_sq (φ : HiggsVec)
|
||||
(hV : potential μSq (λ) φ = - μSq ^ 2 / (4 * (λ))) :
|
||||
‖φ‖ ^ 2 = μSq / (2 * λ) := by
|
||||
have h1 := potential_as_quad μSq (λ) φ
|
||||
rw [quadratic_eq_zero_iff_of_discrim_eq_zero _
|
||||
(potential_eq_bound_discrim_zero μSq hLam φ hV)] at h1
|
||||
simp_rw [h1, neg_neg]
|
||||
exact ne_of_gt hLam
|
||||
|
||||
lemma potential_eq_bound_iff (φ : HiggsVec) :
|
||||
potential μSq (λ) φ = - μSq ^ 2 / (4 * (λ)) ↔ ‖φ‖ ^ 2 = μSq / (2 * (λ)) :=
|
||||
Iff.intro (potential_eq_bound_higgsVec_sq μSq hLam φ)
|
||||
(fun h ↦ by
|
||||
have hv : ‖φ‖ ^ 4 = ‖φ‖ ^ 2 * ‖φ‖ ^ 2 := by ring_nf
|
||||
field_simp [potential, hv, h]
|
||||
ring_nf)
|
||||
|
||||
lemma potential_eq_bound_iff_of_μSq_nonpos {μSq : ℝ}
|
||||
(hμSq : μSq ≤ 0) (φ : HiggsVec) : potential μSq (λ) φ = 0 ↔ φ = 0 :=
|
||||
Iff.intro (fun h ↦ potential_eq_zero_sol_of_μSq_nonpos μSq hLam hμSq φ h)
|
||||
(fun h ↦ by simp [potential, h])
|
||||
|
||||
lemma potential_eq_bound_IsMinOn (φ : HiggsVec)
|
||||
(hv : potential μSq lambda φ = - μSq ^ 2 / (4 * lambda)) :
|
||||
IsMinOn (potential μSq lambda) Set.univ φ := by
|
||||
rw [isMinOn_univ_iff, hv]
|
||||
exact fun x ↦ potential_bounded_below μSq hLam x
|
||||
|
||||
lemma potential_eq_bound_IsMinOn_of_μSq_nonpos {μSq : ℝ}
|
||||
(hμSq : μSq ≤ 0) (φ : HiggsVec) (hv : potential μSq lambda φ = 0) :
|
||||
IsMinOn (potential μSq lambda) Set.univ φ := by
|
||||
rw [isMinOn_univ_iff, hv]
|
||||
exact fun x ↦ potential_bounded_below_of_μSq_nonpos hLam hμSq x
|
||||
|
||||
lemma potential_bound_reached_of_μSq_nonneg {μSq : ℝ} (hμSq : 0 ≤ μSq) :
|
||||
∃ (φ : HiggsVec), potential μSq lambda φ = - μSq ^ 2 / (4 * lambda) := by
|
||||
use ![√(μSq/(2 * lambda)), 0]
|
||||
refine (potential_eq_bound_iff μSq hLam _).mpr ?_
|
||||
simp [PiLp.norm_sq_eq_of_L2]
|
||||
field_simp [mul_pow]
|
||||
|
||||
lemma IsMinOn_potential_iff_of_μSq_nonneg {μSq : ℝ} (hμSq : 0 ≤ μSq) :
|
||||
IsMinOn (potential μSq lambda) Set.univ φ ↔ ‖φ‖ ^ 2 = μSq /(2 * lambda) := by
|
||||
apply Iff.intro <;> rw [← potential_eq_bound_iff μSq hLam φ]
|
||||
· intro h
|
||||
obtain ⟨φm, hφ⟩ := potential_bound_reached_of_μSq_nonneg hLam hμSq
|
||||
have hm := isMinOn_univ_iff.mp h φm
|
||||
rw [hφ] at hm
|
||||
exact (Real.partialOrder.le_antisymm _ _ (potential_bounded_below μSq hLam φ) hm).symm
|
||||
· exact potential_eq_bound_IsMinOn μSq hLam φ
|
||||
|
||||
lemma IsMinOn_potential_iff_of_μSq_nonpos {μSq : ℝ} (hμSq : μSq ≤ 0) :
|
||||
IsMinOn (potential μSq lambda) Set.univ φ ↔ φ = 0 := by
|
||||
apply Iff.intro <;> rw [← potential_eq_bound_iff_of_μSq_nonpos hLam hμSq φ]
|
||||
· intro h
|
||||
have h0 := isMinOn_univ_iff.mp h 0
|
||||
have h1 := potential_bounded_below_of_μSq_nonpos hLam hμSq φ
|
||||
rw [(potential_eq_bound_iff_of_μSq_nonpos hLam hμSq 0).mpr (by rfl)] at h0
|
||||
exact (Real.partialOrder.le_antisymm _ _ h1 h0).symm
|
||||
· exact potential_eq_bound_IsMinOn_of_μSq_nonpos hLam hμSq φ
|
||||
|
||||
end potentialProp
|
||||
/-!
|
||||
|
||||
## Gauge freedom
|
||||
|
||||
-/
|
||||
|
||||
/-- Given a Higgs vector, a rotation matrix which puts the first component of the
|
||||
vector to zero, and the second component to a real -/
|
||||
def rotateMatrix (φ : HiggsVec) : Matrix (Fin 2) (Fin 2) ℂ :=
|
||||
![![φ 1 /‖φ‖ , - φ 0 /‖φ‖], ![conj (φ 0) / ‖φ‖ , conj (φ 1) / ‖φ‖] ]
|
||||
|
||||
lemma rotateMatrix_star (φ : HiggsVec) :
|
||||
star φ.rotateMatrix =
|
||||
![![conj (φ 1) /‖φ‖ , φ 0 /‖φ‖], ![- conj (φ 0) / ‖φ‖ , φ 1 / ‖φ‖] ] := by
|
||||
simp_rw [star, rotateMatrix, conjTranspose]
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;> simp [conj_ofReal]
|
||||
|
||||
lemma rotateMatrix_det {φ : HiggsVec} (hφ : φ ≠ 0) : (rotateMatrix φ).det = 1 := by
|
||||
have h1 : (‖φ‖ : ℂ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
|
||||
field_simp [rotateMatrix, det_fin_two]
|
||||
rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm]
|
||||
|
||||
lemma rotateMatrix_unitary {φ : HiggsVec} (hφ : φ ≠ 0) :
|
||||
(rotateMatrix φ) ∈ unitaryGroup (Fin 2) ℂ := by
|
||||
rw [mem_unitaryGroup_iff', rotateMatrix_star, rotateMatrix]
|
||||
erw [mul_fin_two, one_fin_two]
|
||||
have : (‖φ‖ : ℂ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;> field_simp
|
||||
<;> rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
|
||||
· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm]
|
||||
· ring_nf
|
||||
· ring_nf
|
||||
· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
|
||||
|
||||
lemma rotateMatrix_specialUnitary {φ : HiggsVec} (hφ : φ ≠ 0) :
|
||||
(rotateMatrix φ) ∈ specialUnitaryGroup (Fin 2) ℂ :=
|
||||
mem_specialUnitaryGroup_iff.mpr ⟨rotateMatrix_unitary hφ, rotateMatrix_det hφ⟩
|
||||
|
||||
/-- Given a Higgs vector, an element of the gauge group which puts the first component of the
|
||||
vector to zero, and the second component to a real -/
|
||||
def rotateGuageGroup {φ : HiggsVec} (hφ : φ ≠ 0) : GaugeGroup :=
|
||||
⟨1, ⟨(rotateMatrix φ), rotateMatrix_specialUnitary hφ⟩, 1⟩
|
||||
|
||||
lemma rotateGuageGroup_apply {φ : HiggsVec} (hφ : φ ≠ 0) :
|
||||
rep (rotateGuageGroup hφ) φ = ![0, ofReal ‖φ‖] := by
|
||||
rw [rep_apply]
|
||||
simp only [rotateGuageGroup, rotateMatrix, one_pow, one_smul,
|
||||
Nat.succ_eq_add_one, Nat.reduceAdd, ofReal_eq_coe]
|
||||
ext i
|
||||
fin_cases i
|
||||
· simp only [mulVec, Fin.zero_eta, Fin.isValue, cons_val', empty_val', cons_val_fin_one,
|
||||
cons_val_zero, cons_dotProduct, vecHead, vecTail, Nat.succ_eq_add_one, Nat.reduceAdd,
|
||||
Function.comp_apply, Fin.succ_zero_eq_one, dotProduct_empty, add_zero]
|
||||
ring_nf
|
||||
· simp only [Fin.mk_one, Fin.isValue, cons_val_one, head_cons, mulVec, Fin.isValue,
|
||||
cons_val', empty_val', cons_val_fin_one, vecHead, cons_dotProduct, vecTail, Nat.succ_eq_add_one,
|
||||
Nat.reduceAdd, Function.comp_apply, Fin.succ_zero_eq_one, dotProduct_empty, add_zero]
|
||||
have : (‖φ‖ : ℂ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
|
||||
field_simp
|
||||
rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
|
||||
|
||||
theorem rotate_fst_zero_snd_real (φ : HiggsVec) :
|
||||
∃ (g : GaugeGroup), rep g φ = ![0, ofReal ‖φ‖] := by
|
||||
by_cases h : φ = 0
|
||||
· use ⟨1, 1, 1⟩
|
||||
simp [h]
|
||||
ext i
|
||||
fin_cases i <;> rfl
|
||||
· use rotateGuageGroup h
|
||||
exact rotateGuageGroup_apply h
|
||||
|
||||
end HiggsVec
|
||||
|
||||
end
|
||||
end StandardModel
|
Loading…
Add table
Add a link
Reference in a new issue