refactor: pass at removing double spaces
This commit is contained in:
parent
1fe51b2e04
commit
1133b883f3
19 changed files with 121 additions and 121 deletions
|
@ -72,7 +72,7 @@ We also define the Higgs bundle.
|
|||
/-- The trivial vector bundle 𔲠à ā². -/
|
||||
abbrev HiggsBundle := Bundle.Trivial SpaceTime HiggsVec
|
||||
|
||||
instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
|
||||
instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
|
||||
Bundle.Trivial.smoothVectorBundle HiggsVec š(ā, SpaceTime)
|
||||
|
||||
/-- A Higgs field is a smooth section of the Higgs bundle. -/
|
||||
|
@ -133,11 +133,11 @@ lemma apply_smooth (Ļ : HiggsField) :
|
|||
(smooth_pi_space).mp (Ļ.toVec_smooth)
|
||||
|
||||
lemma apply_re_smooth (Ļ : HiggsField) (i : Fin 2) :
|
||||
Smooth š(ā, SpaceTime) š(ā, ā) (reCLM ā (fun (x : SpaceTime) => (Ļ x i))) :=
|
||||
Smooth š(ā, SpaceTime) š(ā, ā) (reCLM ā (fun (x : SpaceTime) => (Ļ x i))) :=
|
||||
(ContinuousLinearMap.smooth reCLM).comp (Ļ.apply_smooth i)
|
||||
|
||||
lemma apply_im_smooth (Ļ : HiggsField) (i : Fin 2) :
|
||||
Smooth š(ā, SpaceTime) š(ā, ā) (imCLM ā (fun (x : SpaceTime) => (Ļ x i))) :=
|
||||
Smooth š(ā, SpaceTime) š(ā, ā) (imCLM ā (fun (x : SpaceTime) => (Ļ x i))) :=
|
||||
(ContinuousLinearMap.smooth imCLM).comp (Ļ.apply_smooth i)
|
||||
|
||||
/-!
|
||||
|
|
|
@ -34,9 +34,9 @@ open ComplexConjugate
|
|||
@[simps!]
|
||||
noncomputable def higgsRepUnitary : GaugeGroup ā* unitaryGroup (Fin 2) ā where
|
||||
toFun g := repU1 g.2.2 * fundamentalSU2 g.2.1
|
||||
map_mul' := by
|
||||
map_mul' := by
|
||||
intro āØ_, a2, a3ā© āØ_, b2, b3ā©
|
||||
change repU1 (a3 * b3) * fundamentalSU2 (a2 * b2) = _
|
||||
change repU1 (a3 * b3) * fundamentalSU2 (a2 * b2) = _
|
||||
rw [repU1.map_mul, fundamentalSU2.map_mul, mul_assoc, mul_assoc,
|
||||
ā mul_assoc (repU1 b3) _ _, repU1_fundamentalSU2_commute]
|
||||
repeat rw [mul_assoc]
|
||||
|
@ -97,31 +97,31 @@ def rotateMatrix (Ļ : HiggsVec) : Matrix (Fin 2) (Fin 2) ā :=
|
|||
|
||||
lemma rotateMatrix_star (Ļ : HiggsVec) :
|
||||
star Ļ.rotateMatrix =
|
||||
![![conj (Ļ 1) /āĻā , Ļ 0 /āĻā], ![- conj (Ļ 0) / āĻā , Ļ 1 / āĻā] ] := by
|
||||
![![conj (Ļ 1) /āĻā , Ļ 0 /āĻā], ![- conj (Ļ 0) / āĻā , Ļ 1 / āĻā] ] := by
|
||||
simp_rw [star, rotateMatrix, conjTranspose]
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;> simp [conj_ofReal]
|
||||
|
||||
lemma rotateMatrix_det {Ļ : HiggsVec} (hĻ : Ļ ā 0) : (rotateMatrix Ļ).det = 1 := by
|
||||
have h1 : (āĻā : ā) ā 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hĻ)
|
||||
have h1 : (āĻā : ā) ā 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hĻ)
|
||||
field_simp [rotateMatrix, det_fin_two]
|
||||
rw [ā ofReal_mul, ā sq, ā @real_inner_self_eq_norm_sq]
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm]
|
||||
|
||||
lemma rotateMatrix_unitary {Ļ : HiggsVec} (hĻ : Ļ ā 0) :
|
||||
(rotateMatrix Ļ) ā unitaryGroup (Fin 2) ā := by
|
||||
rw [mem_unitaryGroup_iff', rotateMatrix_star, rotateMatrix]
|
||||
erw [mul_fin_two, one_fin_two]
|
||||
have : (āĻā : ā) ā 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hĻ)
|
||||
have : (āĻā : ā) ā 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hĻ)
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;> field_simp
|
||||
<;> rw [ā ofReal_mul, ā sq, ā @real_inner_self_eq_norm_sq]
|
||||
Ā· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Ā· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm, add_comm]
|
||||
Ā· ring_nf
|
||||
Ā· ring_nf
|
||||
Ā· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Ā· simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
|
||||
|
||||
lemma rotateMatrix_specialUnitary {Ļ : HiggsVec} (hĻ : Ļ ā 0) :
|
||||
|
@ -147,10 +147,10 @@ lemma rotateGuageGroup_apply {Ļ : HiggsVec} (hĻ : Ļ ā 0) :
|
|||
Ā· simp only [Fin.mk_one, Fin.isValue, cons_val_one, head_cons, mulVec, Fin.isValue,
|
||||
cons_val', empty_val', cons_val_fin_one, vecHead, cons_dotProduct, vecTail, Nat.succ_eq_add_one,
|
||||
Nat.reduceAdd, Function.comp_apply, Fin.succ_zero_eq_one, dotProduct_empty, add_zero]
|
||||
have : (āĻā : ā) ā 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hĻ)
|
||||
have : (āĻā : ā) ā 0 := ofReal_inj.mp.mt (norm_ne_zero_iff.mpr hĻ)
|
||||
field_simp
|
||||
rw [ā ofReal_mul, ā sq, ā @real_inner_self_eq_norm_sq]
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
simp [PiLp.inner_apply, Complex.inner, neg_mul, sub_neg_eq_add,
|
||||
Fin.sum_univ_two, ofReal_add, ofReal_mul, mul_conj, mul_comm]
|
||||
|
||||
theorem rotate_fst_zero_snd_real (Ļ : HiggsVec) :
|
||||
|
@ -166,7 +166,7 @@ theorem rotate_fst_zero_snd_real (Ļ : HiggsVec) :
|
|||
end HiggsVec
|
||||
|
||||
/-! TODO: Define the global gauge action on HiggsField. -/
|
||||
/-! TODO: Prove `āŖĻ1, Ļ2ā«_H` invariant under the global gauge action. (norm_map_of_mem_unitary) -/
|
||||
/-! TODO: Prove `āŖĻ1, Ļ2ā«_H` invariant under the global gauge action. (norm_map_of_mem_unitary) -/
|
||||
/-! TODO: Prove invariance of potential under global gauge action. -/
|
||||
|
||||
end StandardModel
|
||||
|
|
|
@ -55,13 +55,13 @@ lemma innerProd_left_zero (Ļ : HiggsField) : āŖ0, Ļā«_H = 0 := by
|
|||
lemma innerProd_right_zero (Ļ : HiggsField) : āŖĻ, 0ā«_H = 0 := by
|
||||
funext x
|
||||
simp [innerProd]
|
||||
example (x : ā): RCLike.ofReal x = ofReal' x := by
|
||||
example (x : ā): RCLike.ofReal x = ofReal' x := by
|
||||
rfl
|
||||
lemma innerProd_expand (Ļ1 Ļ2 : HiggsField) :
|
||||
āŖĻ1, Ļ2ā«_H = fun x => equivRealProdCLM.symm (((Ļ1 x 0).re * (Ļ2 x 0).re
|
||||
āŖĻ1, Ļ2ā«_H = fun x => equivRealProdCLM.symm (((Ļ1 x 0).re * (Ļ2 x 0).re
|
||||
+ (Ļ1 x 1).re * (Ļ2 x 1).re+ (Ļ1 x 0).im * (Ļ2 x 0).im + (Ļ1 x 1).im * (Ļ2 x 1).im),
|
||||
((Ļ1 x 0).re * (Ļ2 x 0).im + (Ļ1 x 1).re * (Ļ2 x 1).im
|
||||
- (Ļ1 x 0).im * (Ļ2 x 0).re - (Ļ1 x 1).im * (Ļ2 x 1).re)) := by
|
||||
- (Ļ1 x 0).im * (Ļ2 x 0).re - (Ļ1 x 1).im * (Ļ2 x 1).re)) := by
|
||||
funext x
|
||||
simp only [innerProd, PiLp.inner_apply, RCLike.inner_apply, Fin.sum_univ_two,
|
||||
equivRealProdCLM_symm_apply, ofReal_add, ofReal_mul, ofReal_sub]
|
||||
|
@ -124,9 +124,9 @@ lemma normSq_eq_innerProd_self (Ļ : HiggsField) (x : SpaceTime) :
|
|||
-/
|
||||
|
||||
lemma toHiggsVec_norm (Ļ : HiggsField) (x : SpaceTime) :
|
||||
āĻ xā = āĻ.toHiggsVec xā := rfl
|
||||
āĻ xā = āĻ.toHiggsVec xā := rfl
|
||||
|
||||
lemma normSq_expand (Ļ : HiggsField) :
|
||||
lemma normSq_expand (Ļ : HiggsField) :
|
||||
Ļ.normSq = fun x => (conj (Ļ x 0) * (Ļ x 0) + conj (Ļ x 1) * (Ļ x 1)).re := by
|
||||
funext x
|
||||
simp [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add, @norm_sq_eq_inner ā]
|
||||
|
|
|
@ -34,7 +34,7 @@ open SpaceTime
|
|||
|
||||
/-- The Higgs potential of the form `- μ² * |Ļ|² + šµ * |Ļ|ā“`. -/
|
||||
@[simp]
|
||||
def potential (μ2 šµ : ā ) (Ļ : HiggsField) (x : SpaceTime) : ā :=
|
||||
def potential (μ2 šµ : ā ) (Ļ : HiggsField) (x : SpaceTime) : ā :=
|
||||
- μ2 * āĻā_H ^ 2 x + šµ * āĻā_H ^ 2 x * āĻā_H ^ 2 x
|
||||
|
||||
/-!
|
||||
|
@ -94,7 +94,7 @@ lemma snd_term_nonneg (Ļ : HiggsField) (x : SpaceTime) :
|
|||
and_self]
|
||||
|
||||
lemma as_quad (μ2 šµ : ā) (Ļ : HiggsField) (x : SpaceTime) :
|
||||
šµ * āĻā_H ^ 2 x * āĻā_H ^ 2 x + (- μ2 ) * āĻā_H ^ 2 x + (- potential μ2 šµ Ļ x) = 0 := by
|
||||
šµ * āĻā_H ^ 2 x * āĻā_H ^ 2 x + (- μ2 ) * āĻā_H ^ 2 x + (- potential μ2 šµ Ļ x) = 0 := by
|
||||
simp only [normSq, neg_mul, potential, neg_add_rev, neg_neg]
|
||||
ring
|
||||
|
||||
|
@ -121,7 +121,7 @@ lemma eq_zero_at (Ļ : HiggsField) (x : SpaceTime)
|
|||
ring_nf
|
||||
linear_combination h2
|
||||
|
||||
lemma eq_zero_at_of_μSq_nonpos {μ2 : ā} (hμ2 : μ2 ⤠0)
|
||||
lemma eq_zero_at_of_μSq_nonpos {μ2 : ā} (hμ2 : μ2 ⤠0)
|
||||
(Ļ : HiggsField) (x : SpaceTime) (hV : potential μ2 šµ Ļ x = 0) : Ļ x = 0 := by
|
||||
cases' (eq_zero_at μ2 hšµ Ļ x hV) with h1 h1
|
||||
exact h1
|
||||
|
@ -141,7 +141,7 @@ lemma bounded_below (Ļ : HiggsField) (x : SpaceTime) :
|
|||
ring_nf at h1
|
||||
rw [ā neg_le_iff_add_nonneg'] at h1
|
||||
rw [show šµ * potential μ2 šµ Ļ x * 4 = (4 * šµ) * potential μ2 šµ Ļ x by ring] at h1
|
||||
have h2 := (div_le_iff' (by simp [hšµ] : 0 < 4 * šµ)).mpr h1
|
||||
have h2 := (div_le_iff' (by simp [hšµ] : 0 < 4 * šµ)).mpr h1
|
||||
ring_nf at h2 ā¢
|
||||
exact h2
|
||||
|
||||
|
@ -165,13 +165,13 @@ variable (hšµ : 0 < šµ)
|
|||
-/
|
||||
|
||||
lemma discrim_eq_zero_of_eq_bound (Ļ : HiggsField) (x : SpaceTime)
|
||||
(hV : potential μ2 šµ Ļ x = - μ2 ^ 2 / (4 * šµ)) :
|
||||
(hV : potential μ2 šµ Ļ x = - μ2 ^ 2 / (4 * šµ)) :
|
||||
discrim šµ (- μ2) (- potential μ2 šµ Ļ x) = 0 := by
|
||||
rw [discrim, hV]
|
||||
field_simp
|
||||
|
||||
lemma normSq_of_eq_bound (Ļ : HiggsField) (x : SpaceTime)
|
||||
(hV : potential μ2 šµ Ļ x = - μ2 ^ 2 / (4 * šµ)) :
|
||||
(hV : potential μ2 šµ Ļ x = - μ2 ^ 2 / (4 * šµ)) :
|
||||
āĻā_H ^ 2 x = μ2 / (2 * šµ) := by
|
||||
have h1 := as_quad μ2 šµ Ļ x
|
||||
rw [quadratic_eq_zero_iff_of_discrim_eq_zero _
|
||||
|
@ -180,7 +180,7 @@ lemma normSq_of_eq_bound (Ļ : HiggsField) (x : SpaceTime)
|
|||
exact ne_of_gt hšµ
|
||||
|
||||
lemma eq_bound_iff (Ļ : HiggsField) (x : SpaceTime) :
|
||||
potential μ2 šµ Ļ x = - μ2 ^ 2 / (4 * šµ) ā āĻā_H ^ 2 x = μ2 / (2 * šµ) :=
|
||||
potential μ2 šµ Ļ x = - μ2 ^ 2 / (4 * šµ) ā āĻā_H ^ 2 x = μ2 / (2 * šµ) :=
|
||||
Iff.intro (normSq_of_eq_bound μ2 hšµ Ļ x)
|
||||
(fun h ⦠by
|
||||
rw [potential, h]
|
||||
|
|
Loadingā¦
Add table
Add a link
Reference in a new issue