111 lines
4.4 KiB
Text
111 lines
4.4 KiB
Text
/-
|
||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
||
Released under Apache 2.0 license as described in the file LICENSE.
|
||
Authors: Joseph Tooby-Smith
|
||
-/
|
||
import HepLean.StandardModel.HiggsBoson.Potential
|
||
/-!
|
||
|
||
# The Two Higgs Doublet Model
|
||
|
||
The two Higgs doublet model is the standard model plus an additional Higgs doublet.
|
||
|
||
Currently this file contains the definition of the 2HDM potential.
|
||
|
||
-/
|
||
|
||
namespace TwoHDM
|
||
|
||
open StandardModel
|
||
open ComplexConjugate
|
||
open HiggsField
|
||
|
||
noncomputable section
|
||
|
||
/-- The potential of the two Higgs doublet model. -/
|
||
def potential (māā2 māā2 šµā šµā šµā šµā : ā)
|
||
(māā2 šµā
šµā šµā : ā) (Φ1 Φ2 : HiggsField) (x : SpaceTime) : ā :=
|
||
māā2 * āΦ1ā_H ^ 2 x + māā2 * āΦ2ā_H ^ 2 x - (māā2 * āŖĪ¦1, Φ2ā«_H x + conj māā2 * āŖĪ¦2, Φ1ā«_H x).re
|
||
+ 1/2 * šµā * āΦ1ā_H ^ 2 x * āΦ1ā_H ^ 2 x + 1/2 * šµā * āΦ2ā_H ^ 2 x * āΦ2ā_H ^ 2 x
|
||
+ šµā * āΦ1ā_H ^ 2 x * āΦ2ā_H ^ 2 x
|
||
+ šµā * āāŖĪ¦1, Φ2ā«_H xā ^ 2 + (1/2 * šµā
* āŖĪ¦1, Φ2ā«_H x ^ 2 + 1/2 * conj šµā
* āŖĪ¦2, Φ1ā«_H x ^ 2).re
|
||
+ (šµā * āΦ1ā_H ^ 2 x * āŖĪ¦1, Φ2ā«_H x + conj šµā * āΦ1ā_H ^ 2 x * āŖĪ¦2, Φ1ā«_H x).re
|
||
+ (šµā * āΦ2ā_H ^ 2 x * āŖĪ¦1, Φ2ā«_H x + conj šµā * āΦ2ā_H ^ 2 x * āŖĪ¦2, Φ1ā«_H x).re
|
||
|
||
namespace potential
|
||
|
||
variable (Φ1 Φ2 : HiggsField)
|
||
variable (māā2 māā2 šµā šµā šµā šµā : ā)
|
||
variable (māā2 šµā
šµā šµā : ā)
|
||
/-!
|
||
|
||
## Simple properties of the potential
|
||
|
||
-/
|
||
|
||
/-- Swapping `Φ1` with `Φ2`, and a number of the parameters (with possible conjugation) leads
|
||
to an identical potential. -/
|
||
lemma swap_fields :
|
||
potential māā2 māā2 šµā šµā šµā šµā māā2 šµā
šµā šµā Φ1 Φ2
|
||
= potential māā2 māā2 šµā šµā šµā šµā (conj māā2) (conj šµā
) (conj šµā) (conj šµā) Φ2 Φ1 := by
|
||
funext x
|
||
simp only [potential, HiggsField.normSq, Complex.add_re, Complex.mul_re, Complex.conj_re,
|
||
Complex.conj_im, neg_mul, sub_neg_eq_add, one_div, Complex.norm_eq_abs, Complex.inv_re,
|
||
Complex.re_ofNat, Complex.normSq_ofNat, div_self_mul_self', Complex.inv_im, Complex.im_ofNat,
|
||
neg_zero, zero_div, zero_mul, sub_zero, Complex.mul_im, add_zero, mul_neg, Complex.ofReal_pow,
|
||
RingHomCompTriple.comp_apply, RingHom.id_apply]
|
||
ring_nf
|
||
simp only [one_div, add_left_inj, add_right_inj, mul_eq_mul_left_iff]
|
||
apply Or.inl
|
||
rw [HiggsField.innerProd, HiggsField.innerProd, ā InnerProductSpace.conj_symm, Complex.abs_conj]
|
||
|
||
/-- If `Φā` is zero the potential reduces to the Higgs potential on `Φā`. -/
|
||
lemma right_zero : potential māā2 māā2 šµā šµā šµā šµā māā2 šµā
šµā šµā Φ1 0 =
|
||
StandardModel.HiggsField.potential (- māā2) (šµā/2) Φ1 := by
|
||
funext x
|
||
simp only [potential, normSq, ContMDiffSection.coe_zero, Pi.zero_apply, norm_zero, ne_eq,
|
||
OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, mul_zero, add_zero, innerProd_right_zero,
|
||
innerProd_left_zero, Complex.zero_re, sub_zero, one_div, Complex.ofReal_pow,
|
||
Complex.ofReal_zero, HiggsField.potential, neg_neg, add_right_inj, mul_eq_mul_right_iff,
|
||
pow_eq_zero_iff, norm_eq_zero, or_self_right]
|
||
ring_nf
|
||
simp only [true_or]
|
||
|
||
/-- If `Φā` is zero the potential reduces to the Higgs potential on `Φā`. -/
|
||
lemma left_zero : potential māā2 māā2 šµā šµā šµā šµā māā2 šµā
šµā šµā 0 Φ2 =
|
||
StandardModel.HiggsField.potential (- māā2) (šµā/2) Φ2 := by
|
||
rw [swap_fields, right_zero]
|
||
|
||
/-!
|
||
|
||
## Potential bounded from below
|
||
|
||
-/
|
||
|
||
/-! TODO: Prove bounded properties of the 2HDM potential. -/
|
||
|
||
/-- The proposition on the coefficents for a potential to be bounded. -/
|
||
def IsBounded (māā2 māā2 šµā šµā šµā šµā : ā) (māā2 šµā
šµā šµā : ā) : Prop :=
|
||
ā c, ā Φ1 Φ2 x, c ⤠potential māā2 māā2 šµā šµā šµā šµā māā2 šµā
šµā šµā Φ1 Φ2 x
|
||
|
||
/-! TODO: Show that if the potential is bounded then `0 ⤠šµā` and `0 ⤠šµā`. -/
|
||
/-!
|
||
|
||
## Smoothness of the potential
|
||
|
||
-/
|
||
|
||
/-! TODO: Prove smoothness properties of the 2HDM potential. -/
|
||
|
||
/-!
|
||
|
||
## Invariance of the potential under gauge transformations
|
||
|
||
-/
|
||
|
||
/-! TODO: Prove invariance properties of the 2HDM potential. -/
|
||
|
||
end potential
|
||
|
||
end
|
||
end TwoHDM
|