feat: bounded properties of Higgs potential

This commit is contained in:
jstoobysmith 2024-09-05 18:21:43 -04:00
parent 0dd6f7390e
commit 27721bc476
4 changed files with 249 additions and 57 deletions

View file

@ -59,6 +59,18 @@ lemma smooth_toFin2 : Smooth 𝓘(, HiggsVec) 𝓘(, Fin 2 → ) toF
def orthonormBasis : OrthonormalBasis (Fin 2) HiggsVec :=
EuclideanSpace.basisFun (Fin 2)
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
end HiggsVec
/-!
@ -90,6 +102,9 @@ def HiggsVec.toField (φ : HiggsVec) : HiggsField where
rw [Bundle.contMDiffAt_section]
exact smoothAt_const
@[simp]
lemma HiggsVec.toField_apply (φ : HiggsVec) (x : SpaceTime) : (φ.toField x) = φ := rfl
namespace HiggsField
open HiggsVec
@ -149,6 +164,8 @@ lemma isConst_of_higgsVec (φ : HiggsVec) : φ.toField.IsConst := fun _ => congr
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)
def ofReal (a : ) : HiggsField := (HiggsVec.ofReal a).toField
end HiggsField
end