feat: bounded properties of Higgs potential
This commit is contained in:
parent
0dd6f7390e
commit
27721bc476
4 changed files with 249 additions and 57 deletions
|
@ -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
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue