refactor: Change structure of SM file

This commit is contained in:
jstoobysmith 2024-05-09 15:09:14 -04:00
parent 8fd0b63edb
commit 5af2eb4d8d
5 changed files with 300 additions and 264 deletions

View file

@ -0,0 +1,200 @@
/-
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.HiggsBoson.TargetSpace
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.RepresentationTheory.Basic
import Mathlib.Analysis.InnerProductSpace.Basic
import Mathlib.Analysis.InnerProductSpace.Adjoint
import Mathlib.Geometry.Manifold.ContMDiff.Product
import Mathlib.Analysis.Complex.RealDeriv
import Mathlib.Algebra.QuadraticDiscriminant
/-!
# The Higgs field
This file defines the basic properties for the higgs field in the standard model.
## References
- We use conventions given in: https://pdg.lbl.gov/2019/reviews/rpp2019-rev-higgs-boson.pdf
-/
universe v u
namespace StandardModel
noncomputable section
open Manifold
open Matrix
open Complex
open ComplexConjugate
/-- The trivial vector bundle 𝓡² × ℂ². (TODO: Make associated bundle.) -/
abbrev higgsBundle := Bundle.Trivial spaceTime higgsVec
instance : SmoothVectorBundle higgsVec higgsBundle (𝓡 4) :=
Bundle.Trivial.smoothVectorBundle higgsVec 𝓘(, spaceTime)
/-- A higgs field is a smooth section of the higgs bundle. -/
abbrev higgsField : Type := SmoothSection (𝓡 4) higgsVec higgsBundle
instance : NormedAddCommGroup (Fin 2 → ) := by
exact Pi.normedAddCommGroup
/-- Given a vector `ℂ²` the constant higgs field with value equal to that
section. -/
noncomputable def higgsVec.toField (φ : higgsVec) : higgsField where
toFun := fun _ => φ
contMDiff_toFun := by
intro x
rw [Bundle.contMDiffAt_section]
exact smoothAt_const
namespace higgsField
open Complex Real
/-- Given a `higgsField`, the corresponding map from `spaceTime` to `higgsVec`. -/
def toHiggsVec (φ : higgsField) : spaceTime → higgsVec := φ
lemma toHiggsVec_smooth (φ : higgsField) : Smooth 𝓘(, spaceTime) 𝓘(, higgsVec) φ.toHiggsVec := by
intro x0
have h1 := φ.contMDiff x0
rw [Bundle.contMDiffAt_section] at h1
have h2 :
(fun x => ((trivializationAt higgsVec (Bundle.Trivial spaceTime higgsVec) x0)
{ proj := x, snd := φ x }).2) = φ := by
rfl
simp only [h2] at h1
exact h1
lemma toField_toHiggsVec_apply (φ : higgsField) (x : spaceTime) :
(φ.toHiggsVec x).toField x = φ x := by
rfl
lemma higgsVecToFin2_toHiggsVec (φ : higgsField) : higgsVecToFin2 ∘ φ.toHiggsVec = φ := by
ext x
rfl
lemma toVec_smooth (φ : higgsField) : Smooth 𝓘(, spaceTime) 𝓘(, Fin 2 → ) φ := by
rw [← φ.higgsVecToFin2_toHiggsVec]
exact Smooth.comp smooth_higgsVecToFin2 (φ.toHiggsVec_smooth)
lemma apply_smooth (φ : higgsField) :
∀ i, Smooth 𝓘(, spaceTime) 𝓘(, ) (fun (x : spaceTime) => (φ x i)) := by
rw [← smooth_pi_space]
exact φ.toVec_smooth
lemma apply_re_smooth (φ : higgsField) (i : Fin 2):
Smooth 𝓘(, spaceTime) 𝓘(, ) (reCLM ∘ (fun (x : spaceTime) => (φ x i))) :=
Smooth.comp (ContinuousLinearMap.smooth reCLM) (φ.apply_smooth i)
lemma apply_im_smooth (φ : higgsField) (i : Fin 2):
Smooth 𝓘(, spaceTime) 𝓘(, ) (imCLM ∘ (fun (x : spaceTime) => (φ x i))) :=
Smooth.comp (ContinuousLinearMap.smooth imCLM) (φ.apply_smooth i)
/-- 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)
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 only [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add]
rw [@norm_sq_eq_inner ]
simp
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 normSq_nonneg (φ : higgsField) (x : spaceTime) : 0 ≤ φ.normSq x := by
simp only [normSq, ge_iff_le, norm_nonneg, pow_nonneg]
lemma normSq_zero (φ : higgsField) (x : spaceTime) : φ.normSq x = 0 ↔ φ x = 0 := by
simp only [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero]
/-- The Higgs potential of the form `- μ² * |φ|² + λ * |φ|⁴`. -/
@[simp]
def potential (φ : higgsField) (μSq lambda : ) (x : spaceTime) : :=
- μSq * φ.normSq x + lambda * φ.normSq x * φ.normSq x
lemma potential_smooth (φ : higgsField) (μSq lambda : ) :
Smooth 𝓘(, spaceTime) 𝓘(, ) (fun x => φ.potential μSq lambda x) := by
simp only [potential, normSq, neg_mul]
exact Smooth.add
(Smooth.neg (Smooth.smul smooth_const φ.normSq_smooth))
(Smooth.smul (Smooth.smul smooth_const φ.normSq_smooth) φ.normSq_smooth)
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
/-- 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
intro x _
simp [higgsVec.toField]
lemma isConst_iff_of_higgsVec (Φ : higgsField) : Φ.isConst ↔ ∃ (φ : higgsVec), Φ = φ.toField := by
apply Iff.intro
intro h
use Φ 0
ext x y
rw [← h x 0]
rfl
intro h
intro x y
obtain ⟨φ, hφ⟩ := h
subst hφ
rfl
lemma normSq_of_higgsVec (φ : higgsVec) : φ.toField.normSq = fun x => (norm φ) ^ 2 := by
simp only [normSq, higgsVec.toField]
funext x
simp
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]
funext x
simp only [neg_mul, add_right_inj]
ring_nf
end higgsField
end
end StandardModel

View file

@ -0,0 +1,376 @@
/-
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.VectorBundle.Basic
import Mathlib.Geometry.Manifold.VectorBundle.SmoothSection
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
/-!
# 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: https://pdg.lbl.gov/2019/reviews/rpp2019-rev-higgs-boson.pdf
-/
universe v u
namespace StandardModel
noncomputable section
open Manifold
open Matrix
open Complex
open ComplexConjugate
/-- The complex vector space in which the Higgs field takes values. -/
abbrev higgsVec := EuclideanSpace (Fin 2)
section higgsVec
/-- The continous linear map from the vector space `higgsVec` to `(Fin 2 → )` acheived 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
@[simps!]
noncomputable def higgsRepUnitary : guageGroup →* 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]
rw [mul_assoc, mul_assoc, ← mul_assoc (repU1 b3) _ _, repU1_fundamentalSU2_commute]
repeat rw [mul_assoc]
map_one' := by
simp only [Prod.snd_one, _root_.map_one, Prod.fst_one, mul_one]
/-- An orthonomral basis of higgsVec. -/
noncomputable def orthonormBasis : OrthonormalBasis (Fin 2) higgsVec :=
EuclideanSpace.basisFun (Fin 2)
/-- Takes in a `2×2`-matrix and returns a linear map of `higgsVec`. -/
@[simps!]
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]
rw [mem_unitaryGroup_iff.mp g.prop, mem_unitaryGroup_iff'.mp g.prop, matrixToLin.map_one]
simp
@[simps!]
noncomputable def unitaryToLin : unitaryGroup (Fin 2) →* unitary (higgsVec →L[] higgsVec) where
toFun g := ⟨matrixToLin g, matrixToLin_unitary g⟩
map_mul' g h := by
ext
simp
map_one' := by
ext
simp
@[simps!]
def unitToLinear : unitary (higgsVec →L[] higgsVec) →* higgsVec →ₗ[] higgsVec :=
DistribMulAction.toModuleEnd higgsVec
@[simps!]
def rep : Representation guageGroup higgsVec :=
unitToLinear.comp (unitaryToLin.comp higgsRepUnitary)
lemma higgsRepUnitary_mul (g : guageGroup) (φ : higgsVec) :
(higgsRepUnitary g).1 *ᵥ φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) := by
simp only [higgsRepUnitary_apply_coe]
exact smul_mulVec_assoc (g.2.2 ^ 3) (g.2.1.1) φ
lemma rep_apply (g : guageGroup) (φ : higgsVec) : rep g φ = g.2.2 ^ 3 • (g.2.1.1 *ᵥ φ) :=
higgsRepUnitary_mul g φ
lemma norm_invariant (g : guageGroup) (φ : higgsVec) : ‖rep g φ‖ = ‖φ‖ :=
ContinuousLinearMap.norm_map_of_mem_unitary (unitaryToLin (higgsRepUnitary g)).2 φ
/-- The higgs potential for `higgsVec`, i.e. for constant higgs fields. -/
def potential (μSq lambda : ) (φ : higgsVec) : := - μSq * ‖φ‖ ^ 2 +
lambda * ‖φ‖ ^ 4
lemma potential_invariant (μSq lambda : ) (φ : higgsVec) (g : guageGroup) :
potential μSq lambda (rep g φ) = potential μSq lambda φ := by
simp only [potential, neg_mul]
rw [norm_invariant]
lemma potential_snd_term_nonneg {lambda : } (hLam : 0 < lambda) (φ : higgsVec) :
0 ≤ lambda * ‖φ‖ ^ 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 potential_as_quad (μSq lambda : ) (φ : higgsVec) :
lambda * ‖φ‖ ^ 2 * ‖φ‖ ^ 2 + (- μSq ) * ‖φ‖ ^ 2 + (- potential μSq lambda φ ) = 0 := by
simp [potential]
ring
lemma zero_le_potential_discrim (μSq lambda : ) (φ : higgsVec) (hLam : 0 < lambda) :
0 ≤ discrim (lambda) (- μSq ) (- potential μSq lambda φ) := by
have h1 := potential_as_quad μSq lambda φ
rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1
rw [h1]
exact sq_nonneg (2 * (lambda ) * ‖φ‖ ^ 2 + -μSq)
simp only [ne_eq, div_eq_zero_iff, OfNat.ofNat_ne_zero, or_false]
exact ne_of_gt hLam
lemma potential_eq_zero_sol (μSq lambda : ) (hLam : 0 < lambda)(φ : higgsVec)
(hV : potential μSq lambda φ = 0) : φ = 0 ‖φ‖ ^ 2 = μSq / lambda := by
have h1 := potential_as_quad μSq lambda φ
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 {μSq lambda : } (hLam : 0 < lambda) (hμSq : μSq ≤ 0)
(φ : higgsVec) (hV : potential μSq lambda φ = 0) : φ = 0 := by
cases' (potential_eq_zero_sol μSq lambda hLam φ hV) with h1 h1
exact h1
by_cases hμSqZ : μSq = 0
simpa [hμSqZ] using h1
refine ((?_ : ¬ 0 ≤ μSq / lambda) (?_)).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 (μSq lambda : ) (hLam : 0 < lambda) (φ : higgsVec) :
- μSq ^ 2 / (4 * lambda) ≤ potential μSq lambda φ := by
have h1 := zero_le_potential_discrim μSq lambda φ hLam
simp [discrim] at h1
ring_nf at h1
rw [← neg_le_iff_add_nonneg'] at h1
have h3 : lambda * potential μSq lambda φ * 4 = (4 * lambda) * potential μSq lambda φ := by
ring
rw [h3] at h1
have h2 := (div_le_iff' (by simp [hLam] : 0 < 4 * lambda )).mpr h1
ring_nf at h2 ⊢
exact h2
lemma potential_bounded_below_of_μSq_nonpos {μSq lambda : } (hLam : 0 < lambda)
(hμSq : μSq ≤ 0) (φ : higgsVec) : 0 ≤ potential μSq lambda φ := by
simp only [potential, neg_mul, add_zero]
refine add_nonneg ?_ (potential_snd_term_nonneg hLam φ)
field_simp
rw [@mul_nonpos_iff]
simp_all only [ge_iff_le, norm_nonneg, pow_nonneg, and_self, or_true]
lemma potential_eq_bound_discrim_zero (μSq lambda : ) (hLam : 0 < lambda)(φ : higgsVec)
(hV : potential μSq lambda φ = - μSq ^ 2 / (4 * lambda)) :
discrim (lambda) (- μSq) (- potential μSq lambda φ) = 0 := by
simp [discrim, hV]
field_simp
ring
lemma potential_eq_bound_higgsVec_sq (μSq lambda : ) (hLam : 0 < lambda)(φ : higgsVec)
(hV : potential μSq lambda φ = - μSq ^ 2 / (4 * lambda)) :
‖φ‖ ^ 2 = μSq / (2 * lambda) := by
have h1 := potential_as_quad μSq lambda φ
rw [quadratic_eq_zero_iff_of_discrim_eq_zero _
(potential_eq_bound_discrim_zero μSq lambda hLam φ hV)] at h1
rw [h1]
field_simp
ring_nf
simp only [ne_eq, div_eq_zero_iff, OfNat.ofNat_ne_zero, or_false]
exact ne_of_gt hLam
lemma potential_eq_bound_iff (μSq lambda : ) (hLam : 0 < lambda)(φ : higgsVec) :
potential μSq lambda φ = - μSq ^ 2 / (4 * lambda) ↔ ‖φ‖ ^ 2 = μSq / (2 * lambda) := by
apply Iff.intro
· intro h
exact potential_eq_bound_higgsVec_sq μSq lambda hLam φ h
· intro h
have hv : ‖φ‖ ^ 4 = ‖φ‖ ^ 2 * ‖φ‖ ^ 2 := by
ring_nf
field_simp [potential, hv, h]
ring
lemma potential_eq_bound_iff_of_μSq_nonpos {μSq lambda : } (hLam : 0 < lambda)
(hμSq : μSq ≤ 0) (φ : higgsVec) : potential μSq lambda φ = 0 ↔ φ = 0 := by
apply Iff.intro
· intro h
exact potential_eq_zero_sol_of_μSq_nonpos hLam hμSq φ h
· intro h
simp [potential, h]
lemma potential_eq_bound_IsMinOn (μSq lambda : ) (hLam : 0 < lambda) (φ : higgsVec)
(hv : potential μSq lambda φ = - μSq ^ 2 / (4 * lambda)) :
IsMinOn (potential μSq lambda) Set.univ φ := by
rw [isMinOn_univ_iff]
intro x
rw [hv]
exact potential_bounded_below μSq lambda hLam x
lemma potential_eq_bound_IsMinOn_of_μSq_nonpos {μSq lambda : } (hLam : 0 < lambda)
(hμSq : μSq ≤ 0) (φ : higgsVec) (hv : potential μSq lambda φ = 0) :
IsMinOn (potential μSq lambda) Set.univ φ := by
rw [isMinOn_univ_iff]
intro x
rw [hv]
exact potential_bounded_below_of_μSq_nonpos hLam hμSq x
lemma potential_bound_reached_of_μSq_nonneg {μSq lambda : } (hLam : 0 < lambda) (hμSq : 0 ≤ μSq) :
∃ (φ : higgsVec), potential μSq lambda φ = - μSq ^ 2 / (4 * lambda) := by
use ![√(μSq/(2 * lambda)), 0]
refine (potential_eq_bound_iff μSq lambda hLam _).mpr ?_
simp [@PiLp.norm_sq_eq_of_L2, Fin.sum_univ_two]
field_simp [mul_pow]
lemma IsMinOn_potential_iff_of_μSq_nonneg {μSq lambda : } (hLam : 0 < lambda) (hμSq : 0 ≤ μSq) :
IsMinOn (potential μSq lambda) Set.univ φ ↔ ‖φ‖ ^ 2 = μSq /(2 * lambda) := by
apply Iff.intro
· 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
have h1 := potential_bounded_below μSq lambda hLam φ
rw [← potential_eq_bound_iff μSq lambda hLam φ]
exact (Real.partialOrder.le_antisymm _ _ h1 hm).symm
· intro h
rw [← potential_eq_bound_iff μSq lambda hLam φ] at h
exact potential_eq_bound_IsMinOn μSq lambda hLam φ h
lemma IsMinOn_potential_iff_of_μSq_nonpos {μSq lambda : } (hLam : 0 < lambda) (hμSq : μSq ≤ 0) :
IsMinOn (potential μSq lambda) Set.univ φ ↔ φ = 0 := by
apply Iff.intro
· intro h
have h0 := isMinOn_univ_iff.mp h 0
rw [(potential_eq_bound_iff_of_μSq_nonpos hLam hμSq 0).mpr (by rfl)] at h0
have h1 := potential_bounded_below_of_μSq_nonpos hLam hμSq φ
rw [← (potential_eq_bound_iff_of_μSq_nonpos hLam hμSq φ)]
exact (Real.partialOrder.le_antisymm _ _ h1 h0).symm
· intro h
rw [← potential_eq_bound_iff_of_μSq_nonpos hLam hμSq φ] at h
exact potential_eq_bound_IsMinOn_of_μSq_nonpos hLam hμSq φ h
/-- Given a Higgs vector, a rotation matrix which puts the fst component of the
vector to zero, and the snd componenet 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 [star]
rw [rotateMatrix, conjTranspose]
ext i j
fin_cases i <;> fin_cases j <;> simp [conj_ofReal]
lemma rotateMatrix_det {φ : higgsVec} (hφ : φ ≠ 0) : (rotateMatrix φ).det = 1 := by
simp [rotateMatrix, det_fin_two]
have h1 : (‖φ‖ : ) ≠ 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hφ)
field_simp
rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
simp only [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]
rfl
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φ)
congr
field_simp
ext i j
fin_cases i <;> fin_cases j <;> field_simp
· rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
simp only [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]
rfl
· ring_nf
· ring_nf
· rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
simp only [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
rfl
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 fst component of the
vector to zero, and the snd componenet to a real -/
def rotateGuageGroup {φ : higgsVec} (hφ : φ ≠ 0) : guageGroup :=
⟨1, ⟨(rotateMatrix φ), rotateMatrix_specialUnitary hφ⟩, 1⟩
lemma rotateGuageGroup_apply {φ : higgsVec} (hφ : φ ≠ 0) :
rep (rotateGuageGroup hφ) φ = ![0, ofReal ‖φ‖] := by
rw [rep_apply]
simp [rotateGuageGroup, rotateMatrix]
ext i
fin_cases i
simp [mulVec, vecHead, vecTail]
ring_nf
simp only [Fin.mk_one, Fin.isValue, cons_val_one, head_cons]
simp [mulVec, vecHead, vecTail]
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 only [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
rfl
theorem rotate_fst_zero_snd_real (φ : higgsVec) :
∃ (g : guageGroup), 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 higgsVec
end
end StandardModel