chore: Bump to 4.14.0
This commit is contained in:
parent
c91ca06272
commit
5dfd29ab8d
32 changed files with 376 additions and 334 deletions
|
@ -53,8 +53,8 @@ def toFin2ℂ : HiggsVec →L[ℝ] (Fin 2 → ℂ) where
|
|||
map_smul' a x := rfl
|
||||
|
||||
/-- The map `toFin2ℂ` is smooth. -/
|
||||
lemma smooth_toFin2ℂ : Smooth 𝓘(ℝ, HiggsVec) 𝓘(ℝ, Fin 2 → ℂ) toFin2ℂ :=
|
||||
ContinuousLinearMap.smooth toFin2ℂ
|
||||
lemma smooth_toFin2ℂ : ContMDiff 𝓘(ℝ, HiggsVec) 𝓘(ℝ, Fin 2 → ℂ) ⊤ toFin2ℂ :=
|
||||
ContinuousLinearMap.contMDiff toFin2ℂ
|
||||
|
||||
/-- An orthonormal basis of `HiggsVec`. -/
|
||||
def orthonormBasis : OrthonormalBasis (Fin 2) ℂ HiggsVec :=
|
||||
|
@ -92,7 +92,7 @@ instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
|
|||
Bundle.Trivial.smoothVectorBundle HiggsVec
|
||||
|
||||
/-- A Higgs field is a smooth section of the Higgs bundle. -/
|
||||
abbrev HiggsField : Type := SmoothSection SpaceTime.asSmoothManifold HiggsVec HiggsBundle
|
||||
abbrev HiggsField : Type := ContMDiffSection SpaceTime.asSmoothManifold HiggsVec ⊤ HiggsBundle
|
||||
|
||||
/-- Given a vector in `HiggsVec` the constant Higgs field with value equal to that
|
||||
section. -/
|
||||
|
@ -101,7 +101,7 @@ def HiggsVec.toField (φ : HiggsVec) : HiggsField where
|
|||
contMDiff_toFun := by
|
||||
intro x
|
||||
rw [Bundle.contMDiffAt_section]
|
||||
exact smoothAt_const
|
||||
exact contMDiffAt_const
|
||||
|
||||
/-- For all spacetime points, the constant Higgs field defined by a Higgs vector,
|
||||
returns that Higgs Vector. -/
|
||||
|
@ -120,7 +120,8 @@ open 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
|
||||
lemma toHiggsVec_smooth (φ : HiggsField) :
|
||||
ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, HiggsVec) ⊤ φ.toHiggsVec := by
|
||||
intro x0
|
||||
have h1 := φ.contMDiff x0
|
||||
rw [Bundle.contMDiffAt_section] at h1
|
||||
|
@ -138,20 +139,20 @@ lemma toFin2ℂ_comp_toHiggsVec (φ : HiggsField) :
|
|||
|
||||
-/
|
||||
|
||||
lemma toVec_smooth (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, Fin 2 → ℂ) φ :=
|
||||
lemma toVec_smooth (φ : HiggsField) : ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, Fin 2 → ℂ) ⊤ φ :=
|
||||
smooth_toFin2ℂ.comp φ.toHiggsVec_smooth
|
||||
|
||||
lemma apply_smooth (φ : HiggsField) :
|
||||
∀ i, Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℂ) (fun (x : SpaceTime) => (φ x i)) :=
|
||||
(smooth_pi_space).mp (φ.toVec_smooth)
|
||||
∀ i, ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℂ) ⊤ (fun (x : SpaceTime) => (φ x i)) :=
|
||||
(contMDiff_pi_space).mp (φ.toVec_smooth)
|
||||
|
||||
lemma apply_re_smooth (φ : HiggsField) (i : Fin 2) :
|
||||
Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (reCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
|
||||
(ContinuousLinearMap.smooth reCLM).comp (φ.apply_smooth i)
|
||||
ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) ⊤ (reCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
|
||||
(ContinuousLinearMap.contMDiff reCLM).comp (φ.apply_smooth i)
|
||||
|
||||
lemma apply_im_smooth (φ : HiggsField) (i : Fin 2) :
|
||||
Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (imCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
|
||||
(ContinuousLinearMap.smooth imCLM).comp (φ.apply_smooth i)
|
||||
ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) ⊤ (imCLM ∘ (fun (x : SpaceTime) => (φ x i))) :=
|
||||
(ContinuousLinearMap.contMDiff imCLM).comp (φ.apply_smooth i)
|
||||
|
||||
/-!
|
||||
|
||||
|
|
|
@ -124,14 +124,21 @@ lemma rotateMatrix_unitary {φ : HiggsVec} (hφ : φ ≠ 0) :
|
|||
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_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]
|
||||
· simp only [Fin.isValue, Fin.zero_eta, Fin.mk_one, of_apply, cons_val', cons_val_one, head_cons,
|
||||
empty_val', cons_val_fin_one, cons_val_zero]
|
||||
ring_nf
|
||||
· simp only [Fin.isValue, Fin.mk_one, Fin.zero_eta, of_apply, cons_val', cons_val_zero,
|
||||
empty_val', cons_val_fin_one, cons_val_one, head_fin_const]
|
||||
ring_nf
|
||||
· field_simp
|
||||
rw [← ofReal_mul, ← sq, ← @real_inner_self_eq_norm_sq]
|
||||
simp only [Fin.isValue, mul_comm, mul_conj, PiLp.inner_apply, Complex.inner, ofReal_re,
|
||||
Fin.sum_univ_two, ofReal_add]
|
||||
|
||||
/-- The `rotateMatrix` for a non-zero Higgs vector is special unitary. -/
|
||||
lemma rotateMatrix_specialUnitary {φ : HiggsVec} (hφ : φ ≠ 0) :
|
||||
|
|
|
@ -91,9 +91,9 @@ lemma innerProd_expand (φ1 φ2 : HiggsField) :
|
|||
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
|
||||
lemma smooth_innerProd (φ1 φ2 : HiggsField) : ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℂ) ⊤ ⟪φ1, φ2⟫_H := by
|
||||
rw [innerProd_expand]
|
||||
exact (ContinuousLinearMap.smooth (equivRealProdCLM.symm : ℝ × ℝ →L[ℝ] ℂ)).comp $
|
||||
exact (ContinuousLinearMap.contMDiff (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
|
||||
|
@ -171,9 +171,9 @@ lemma normSq_zero (φ : HiggsField) (x : SpaceTime) : φ.normSq x = 0 ↔ φ x =
|
|||
simp [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero]
|
||||
|
||||
/-- The norm squared of the Higgs field is a smooth function on space-time. -/
|
||||
lemma normSq_smooth (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) φ.normSq := by
|
||||
lemma normSq_smooth (φ : HiggsField) : ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) ⊤ φ.normSq := by
|
||||
rw [normSq_expand]
|
||||
refine Smooth.add ?_ ?_
|
||||
refine ContMDiff.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)
|
||||
|
|
|
@ -50,10 +50,10 @@ def toFun (φ : HiggsField) (x : SpaceTime) : ℝ :=
|
|||
|
||||
/-- The potential is smooth. -/
|
||||
lemma toFun_smooth (φ : HiggsField) :
|
||||
Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (fun x => P.toFun φ x) := by
|
||||
ContMDiff 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) ⊤ (fun x => P.toFun φ x) := by
|
||||
simp only [toFun, normSq, neg_mul]
|
||||
exact (smooth_const.smul φ.normSq_smooth).neg.add
|
||||
((smooth_const.smul φ.normSq_smooth).smul φ.normSq_smooth)
|
||||
exact (contMDiff_const.smul φ.normSq_smooth).neg.add
|
||||
((contMDiff_const.smul φ.normSq_smooth).smul φ.normSq_smooth)
|
||||
|
||||
/-- The Higgs potential formed by negating the mass squared and the quartic coupling. -/
|
||||
def neg : Potential where
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue