refactor: pass at removing double spaces

This commit is contained in:
jstoobysmith 2024-07-12 10:36:39 -04:00
parent 1fe51b2e04
commit 1133b883f3
19 changed files with 121 additions and 121 deletions

View file

@ -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)
/-!

View file

@ -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

View file

@ -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 ā„‚]

View file

@ -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]