feat: More boundedness properties of Higgs potentials

This commit is contained in:
jstoobysmith 2024-09-06 06:38:33 -04:00
parent 5a613b4de6
commit 5cadb8d4b2
2 changed files with 93 additions and 2 deletions

View file

@ -35,7 +35,7 @@ open SpaceTime
/-- 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⟫_
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
@ -46,6 +46,16 @@ scoped[StandardModel.HiggsField] notation "⟪" φ1 "," φ2 "⟫_H" => innerProd
-/
@[simp]
lemma innerProd_neg_left (φ1 φ2 : HiggsField) : ⟪- φ1, φ2⟫_H = -⟪φ1, φ2⟫_H := by
funext x
simp [innerProd]
@[simp]
lemma innerProd_neg_right (φ1 φ2 : HiggsField) : ⟪φ1, - φ2⟫_H = -⟪φ1, φ2⟫_H := by
funext x
simp [innerProd]
@[simp]
lemma innerProd_left_zero (φ : HiggsField) : ⟪0, φ⟫_H = 0 := by
funext x
@ -105,6 +115,7 @@ scoped[StandardModel.HiggsField] notation "‖" φ1 "‖_H ^ 2" => normSq φ1
-/
@[simp]
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]
@ -122,6 +133,20 @@ lemma normSq_eq_innerProd_self (φ : HiggsField) (x : SpaceTime) :
-/
@[simp]
lemma normSq_apply_im_zero (φ : HiggsField) (x : SpaceTime) :
((Complex.ofReal' ‖φ x‖) ^ 2).im = 0 := by
rw [sq]
simp only [Complex.ofReal_eq_coe, Complex.mul_im, Complex.ofReal_re, Complex.ofReal_im,
mul_zero, zero_mul, add_zero]
@[simp]
lemma normSq_apply_re_self (φ : HiggsField) (x : SpaceTime) :
((Complex.ofReal' ‖φ x‖) ^ 2).re = φ.normSq x := by
rw [sq]
simp only [mul_re, ofReal_re, ofReal_im, mul_zero, sub_zero, normSq]
exact Eq.symm (pow_two ‖φ x‖)
lemma toHiggsVec_norm (φ : HiggsField) (x : SpaceTime) :
‖φ x‖ = ‖φ.toHiggsVec x‖ := rfl