feat: More fixes

This commit is contained in:
jstoobysmith 2024-11-02 08:50:17 +00:00
parent 4df8663cbc
commit c9c9047a0c
31 changed files with 134 additions and 124 deletions

View file

@ -88,7 +88,7 @@ We also define the Higgs bundle.
abbrev HiggsBundle := Bundle.Trivial SpaceTime HiggsVec
instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
Bundle.Trivial.smoothVectorBundle HiggsVec 𝓘(, SpaceTime)
Bundle.Trivial.smoothVectorBundle HiggsVec
/-- A Higgs field is a smooth section of the Higgs bundle. -/
abbrev HiggsField : Type := SmoothSection SpaceTime.asSmoothManifold HiggsVec HiggsBundle

View file

@ -135,10 +135,10 @@ def rotateGuageGroup {φ : HiggsVec} (hφ : φ ≠ 0) : GaugeGroupI :=
⟨1, ⟨(rotateMatrix φ), rotateMatrix_specialUnitary hφ⟩, 1⟩
lemma rotateGuageGroup_apply {φ : HiggsVec} (hφ : φ ≠ 0) :
rep (rotateGuageGroup hφ) φ = ![0, Complex.ofReal ‖φ‖] := by
rep (rotateGuageGroup hφ) φ = ![0, Complex.ofRealHom ‖φ‖] := by
rw [rep_apply]
simp only [rotateGuageGroup, rotateMatrix, one_pow, one_smul,
Nat.succ_eq_add_one, Nat.reduceAdd, ofReal_eq_coe]
Nat.succ_eq_add_one, Nat.reduceAdd, ofRealHom_eq_coe]
ext i
fin_cases i
· simp only [mulVec, Fin.zero_eta, Fin.isValue, cons_val', empty_val', cons_val_fin_one,
@ -181,7 +181,7 @@ theorem rotate_fst_real_snd_zero (φ : HiggsVec) :
rw [rep.map_mul]
change rep P (rep g φ) = _
rw [h, rep_apply]
simp only [one_pow, Nat.succ_eq_add_one, Nat.reduceAdd, ofReal_eq_coe, mulVec_cons, zero_smul,
simp only [one_pow, Nat.succ_eq_add_one, Nat.reduceAdd, ofRealHom_eq_coe, mulVec_cons, zero_smul,
coe_smul, mulVec_empty, add_zero, zero_add, one_smul]
funext i
fin_cases i

View file

@ -87,8 +87,7 @@ lemma innerProd_expand (φ1 φ2 : HiggsField) :
nth_rewrite 1 [← RCLike.re_add_im (φ2 x 0)]
nth_rewrite 1 [← RCLike.re_add_im (φ2 x 1)]
ring_nf
repeat rw [show RCLike.ofReal _ = ofReal' _ by rfl]
simp only [Algebra.id.map_eq_id, RCLike.re_to_complex, RingHom.id_apply, RCLike.I_to_complex,
simp only [Fin.isValue, RCLike.re_to_complex, coe_algebraMap, RCLike.I_to_complex,
RCLike.im_to_complex, I_sq, mul_neg, mul_one, neg_mul, sub_neg_eq_add, one_mul]
ring
@ -144,16 +143,15 @@ lemma normSq_eq_innerProd_self (φ : HiggsField) (x : SpaceTime) :
@[simp]
lemma normSq_apply_im_zero (φ : HiggsField) (x : SpaceTime) :
((Complex.ofReal' ‖φ x‖) ^ 2).im = 0 := by
((Complex.ofRealHom ‖φ 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 only [ofRealHom_eq_coe, mul_im, ofReal_re, 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
((Complex.ofRealHom ‖φ x‖) ^ 2).re = φ.normSq x := by
rw [sq]
simp only [mul_re, ofReal_re, ofReal_im, mul_zero, sub_zero, normSq]
simp only [ofRealHom_eq_coe, mul_re, ofReal_re, ofReal_im, mul_zero, sub_zero, normSq]
exact Eq.symm (pow_two ‖φ x‖)
lemma toHiggsVec_norm (φ : HiggsField) (x : SpaceTime) :

View file

@ -126,7 +126,7 @@ def quadDiscrim (φ : HiggsField) (x : SpaceTime) : := discrim P.𝓵 (- P.
lemma quadDiscrim_nonneg (h : P.𝓵 ≠ 0) (φ : HiggsField) (x : SpaceTime) :
0 ≤ P.quadDiscrim φ x := by
have h1 := P.as_quad φ x
rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1
rw [mul_assoc, quadratic_eq_zero_iff_discrim_eq_sq] at h1
· simp only [h1, ne_eq, quadDiscrim, div_eq_zero_iff, OfNat.ofNat_ne_zero, or_false]
exact sq_nonneg (2 * P.𝓵 * ‖φ‖_H ^ 2 x + - P.μ2)
· exact h
@ -148,7 +148,7 @@ lemma quadDiscrim_eq_zero_iff_normSq (h : P.𝓵 ≠ 0) (φ : HiggsField) (x : S
rw [P.quadDiscrim_eq_zero_iff h]
refine Iff.intro (fun hV => ?_) (fun hF => ?_)
· have h1 := P.as_quad φ x
rw [quadratic_eq_zero_iff_of_discrim_eq_zero h
rw [mul_assoc, quadratic_eq_zero_iff_of_discrim_eq_zero h
((P.quadDiscrim_eq_zero_iff h φ x).mpr hV)] at h1
simp_rw [h1, neg_neg]
· rw [toFun, hF]
@ -241,6 +241,7 @@ lemma neg_𝓵_sol_exists_iff (h𝓵 : P.𝓵 < 0) (c : ) : (∃ φ x, P.toFu
have hdd : discrim P.𝓵 (- P.μ2) (-c) = Real.sqrt (discrim P.𝓵 (- P.μ2) (-c))
* Real.sqrt (discrim P.𝓵 (- P.μ2) (-c)) := by
exact (Real.mul_self_sqrt hd).symm
rw [mul_assoc]
refine (quadratic_eq_zero_iff (ne_of_gt h𝓵).symm hdd _).mpr ?_
simp only [neg_neg, or_true, a]