Skip to content

Commit

Permalink
Merge pull request #46 from pitmonticone/golf-proofs
Browse files Browse the repository at this point in the history
Golf a few proofs
  • Loading branch information
jstoobysmith authored Jun 9, 2024
2 parents de720c6 + 9850a9c commit 1cb2cdf
Show file tree
Hide file tree
Showing 3 changed files with 101 additions and 190 deletions.
69 changes: 24 additions & 45 deletions HepLean/SpaceTime/Metric.lean
Original file line number Diff line number Diff line change
Expand Up @@ -85,40 +85,22 @@ lemma η_transpose : η.transpose = η := by

@[simp]
lemma det_η : η.det = - 1 := by
simp only [η_explicit, det_succ_row_zero, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue,
of_apply, cons_val', empty_val', cons_val_fin_one, cons_val_zero, submatrix_apply,
Fin.succ_zero_eq_one, cons_val_one, head_cons, submatrix_submatrix, Function.comp_apply,
Fin.succ_one_eq_two, cons_val_two, tail_cons, det_unique, Fin.default_eq_zero, cons_val_succ,
head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, one_mul, Fin.zero_succAbove,
Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, neg_mul,
Fin.succ_succAbove_zero, Finset.sum_neg_distrib, Finset.sum_singleton, Fin.succ_succAbove_one,
even_two, Even.neg_pow, one_pow, mul_one, mul_neg, neg_neg, mul_zero, neg_zero, add_zero,
zero_mul, Finset.sum_const_zero]
simp [η_explicit, det_succ_row_zero, Fin.sum_univ_succ]

@[simp]
lemma η_sq : η * η = 1 := by
funext μ ν
rw [mul_apply, Fin.sum_univ_four]
fin_cases μ <;> fin_cases ν <;>
simp [η_explicit, Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one, Matrix.cons_val_one,
Matrix.cons_val_succ', Matrix.cons_val_zero, Matrix.empty_val', Matrix.head_cons,
Matrix.head_fin_const, Matrix.head_cons, Matrix.vecCons_const, Fin.mk_one, Fin.mk_one,
vecHead, vecTail, Function.comp_apply]
simp [η_explicit, vecHead, vecTail]

lemma η_diag_mul_self (μ : Fin 4) : η μ μ * η μ μ = 1 := by
fin_cases μ
<;> simp [η_explicit]
fin_cases μ <;> simp [η_explicit]

lemma η_mulVec (x : spaceTime) : η *ᵥ x = ![x 0, -x 1, -x 2, -x 3] := by
rw [explicit x]
rw [η_explicit]
rw [explicit x, η_explicit]
funext i
rw [mulVec, dotProduct, Fin.sum_univ_four]
fin_cases i <;>
simp [Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one, Matrix.cons_val_one,
Matrix.cons_val_succ', Matrix.cons_val_zero, Matrix.empty_val', Matrix.head_cons,
Matrix.head_fin_const, Matrix.head_cons, Matrix.vecCons_const, Fin.mk_one, Fin.mk_one,
vecHead, vecTail, Function.comp_apply]
simp [vecHead, vecTail]

/-- Given a point in spaceTime `x` the linear map `y → x ⬝ᵥ (η *ᵥ y)`. -/
@[simps!]
Expand All @@ -128,7 +110,7 @@ def linearMapForSpaceTime (x : spaceTime) : spaceTime →ₗ[ℝ] ℝ where
simp only
rw [mulVec_add, dotProduct_add]
map_smul' c y := by
simp only
simp only [RingHom.id_apply, smul_eq_mul]
rw [mulVec_smul, dotProduct_smul]
rfl

Expand Down Expand Up @@ -168,7 +150,7 @@ lemma time_elm_sq_of_ηLin (x : spaceTime) : x 0 ^ 2 = ηLin x x + ‖x.space‖

lemma ηLin_leq_time_sq (x : spaceTime) : ηLin x x ≤ x 0 ^ 2 := by
rw [time_elm_sq_of_ηLin]
apply (le_add_iff_nonneg_right _).mpr $ sq_nonneg ‖x.space‖
exact (le_add_iff_nonneg_right _).mpr $ sq_nonneg ‖x.space‖

lemma ηLin_space_inner_product (x y : spaceTime) :
ηLin x y = x 0 * y 0 - ⟪x.space, y.space⟫_ℝ := by
Expand Down Expand Up @@ -202,18 +184,18 @@ lemma ηLin_stdBasis_apply (μ : Fin 4) (x : spaceTime) : ηLin (stdBasis μ) x
lemma ηLin_η_stdBasis (μ ν : Fin 4) : ηLin (stdBasis μ) (stdBasis ν) = η μ ν := by
rw [ηLin_stdBasis_apply]
by_cases h : μ = ν
rw [stdBasis_apply]
subst h
simp only [↓reduceIte, mul_one]
rw [stdBasis_not_eq, η_off_diagonal h]
simp only [mul_zero]
exact fun a => h (id a.symm)
· rw [stdBasis_apply]
subst h
simp
· rw [stdBasis_not_eq, η_off_diagonal h]
simp only [mul_zero]
exact fun a h (id a.symm)

lemma ηLin_mulVec_left (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ℝ) :
ηLin (Λ *ᵥ x) y = ηLin x ((η * Λᵀ * η) *ᵥ y) := by
simp only [ηLin, LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply, mulVec_mulVec]
rw [(vecMul_transpose Λ x).symm, ← dotProduct_mulVec, mulVec_mulVec]
rw [← mul_assoc, ← mul_assoc, η_sq, one_mul]
simp [ηLin, LinearMap.coe_mk, AddHom.coe_mk, linearMapForSpaceTime_apply,
mulVec_mulVec, (vecMul_transpose Λ x).symm, ← dotProduct_mulVec, mulVec_mulVec,
← mul_assoc, ← mul_assoc, η_sq, one_mul]

lemma ηLin_mulVec_right (x y : spaceTime) (Λ : Matrix (Fin 4) (Fin 4) ℝ) :
ηLin x (Λ *ᵥ y) = ηLin ((η * Λᵀ * η) *ᵥ x) y := by
Expand All @@ -231,14 +213,14 @@ lemma ηLin_matrix_stdBasis' (μ ν : Fin 4) (Λ : Matrix (Fin 4) (Fin 4) ℝ) :
lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ℝ) :
Λ = 1 ↔ ∀ (x y : spaceTime), ηLin x y = ηLin x (Λ *ᵥ y) := by
apply Iff.intro
intro h
subst h
simp only [ηLin, one_mulVec, implies_true]
intro h
funext μ ν
have h1 := h (stdBasis μ) (stdBasis ν)
rw [ηLin_matrix_stdBasis, ηLin_η_stdBasis] at h1
fin_cases μ <;> fin_cases ν <;>
· intro h
subst h
simp only [ηLin, one_mulVec, implies_true]
· intro h
funext μ ν
have h1 := h (stdBasis μ) (stdBasis ν)
rw [ηLin_matrix_stdBasis, ηLin_η_stdBasis] at h1
fin_cases μ <;> fin_cases ν <;>
simp_all [η_explicit, Fin.zero_eta, Matrix.cons_val', Matrix.cons_val_fin_one,
Matrix.cons_val_one,
Matrix.cons_val_succ', Matrix.cons_val_zero, Matrix.empty_val', Matrix.head_cons,
Expand All @@ -248,9 +230,6 @@ lemma ηLin_matrix_eq_identity_iff (Λ : Matrix (Fin 4) (Fin 4) ℝ) :
/-- The metric as a quadratic form on `spaceTime`. -/
def quadraticForm : QuadraticForm ℝ spaceTime := ηLin.toQuadraticForm




end spaceTime

end
60 changes: 18 additions & 42 deletions HepLean/StandardModel/HiggsBoson/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -51,7 +51,7 @@ instance : NormedAddCommGroup (Fin 2 → ℂ) := by
/-- Given a vector `ℂ²` the constant higgs field with value equal to that
section. -/
noncomputable def higgsVec.toField (φ : higgsVec) : higgsField where
toFun := fun _ => φ
toFun := fun _ φ
contMDiff_toFun := by
intro x
rw [Bundle.contMDiffAt_section]
Expand All @@ -63,7 +63,6 @@ open Complex Real
/-- Given a `higgsField`, the corresponding map from `spaceTime` to `higgsVec`. -/
def toHiggsVec (φ : higgsField) : spaceTime → higgsVec := φ


lemma toHiggsVec_smooth (φ : higgsField) : Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, higgsVec) φ.toHiggsVec := by
intro x0
have h1 := φ.contMDiff x0
Expand All @@ -76,29 +75,25 @@ lemma toHiggsVec_smooth (φ : higgsField) : Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ
exact h1

lemma toField_toHiggsVec_apply (φ : higgsField) (x : spaceTime) :
(φ.toHiggsVec x).toField x = φ x := by
rfl
(φ.toHiggsVec x).toField x = φ x := rfl

lemma higgsVecToFin2ℂ_toHiggsVec (φ : higgsField) : higgsVecToFin2ℂ ∘ φ.toHiggsVec = φ := by
ext x
rfl
lemma higgsVecToFin2ℂ_toHiggsVec (φ : higgsField) :
higgsVecToFin2ℂ ∘ φ.toHiggsVec = φ := rfl

lemma toVec_smooth (φ : higgsField) : Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, Fin 2 → ℂ) φ := by
rw [← φ.higgsVecToFin2ℂ_toHiggsVec]
exact Smooth.comp smooth_higgsVecToFin2ℂ (φ.toHiggsVec_smooth)
lemma toVec_smooth (φ : higgsField) : Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, Fin 2 → ℂ) φ :=
smooth_higgsVecToFin2ℂ.comp φ.toHiggsVec_smooth

lemma apply_smooth (φ : higgsField) :
∀ i, Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, ℂ) (fun (x : spaceTime) => (φ x i)) := by
rw [← smooth_pi_space]
exact φ.toVec_smooth
∀ i, Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, ℂ) (fun (x : spaceTime) => (φ x i)) :=
(smooth_pi_space).mp (φ.toVec_smooth)

lemma apply_re_smooth (φ : higgsField) (i : Fin 2):
Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, ℝ) (reCLM ∘ (fun (x : spaceTime) => (φ x i))) :=
Smooth.comp (ContinuousLinearMap.smooth reCLM) (φ.apply_smooth 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.comp (ContinuousLinearMap.smooth imCLM) (φ.apply_smooth i)
(ContinuousLinearMap.smooth imCLM).comp (φ.apply_smooth i)

/-- Given two `higgsField`, the map `spaceTime → ℂ` obtained by taking their inner product. -/
def innerProd1 φ2 : higgsField) : spaceTime → ℂ := fun x => ⟪φ1 x, φ2 x⟫_ℂ
Expand All @@ -115,9 +110,7 @@ lemma toHiggsVec_norm (φ : higgsField) (x : spaceTime) :
lemma normSq_expand (φ : higgsField) :
φ.normSq = fun x => (conj (φ x 0) * (φ x 0) + conj (φ x 1) * (φ x 1) ).re := by
funext x
simp only [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add]
rw [@norm_sq_eq_inner ℂ]
simp
simp [normSq, add_re, mul_re, conj_re, conj_im, neg_mul, sub_neg_eq_add, @norm_sq_eq_inner ℂ]

lemma normSq_smooth (φ : higgsField) : Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, ℝ) φ.normSq := by
rw [normSq_expand]
Expand All @@ -140,10 +133,10 @@ lemma normSq_smooth (φ : higgsField) : Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ,
exact φ.apply_im_smooth 1

lemma normSq_nonneg (φ : higgsField) (x : spaceTime) : 0 ≤ φ.normSq x := by
simp only [normSq, ge_iff_le, norm_nonneg, pow_nonneg]
simp [normSq, ge_iff_le, norm_nonneg, pow_nonneg]

lemma normSq_zero (φ : higgsField) (x : spaceTime) : φ.normSq x = 0 ↔ φ x = 0 := by
simp only [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero]
simp [normSq, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff, norm_eq_zero]

/-- The Higgs potential of the form `- μ² * |φ|² + λ * |φ|⁴`. -/
@[simp]
Expand All @@ -153,10 +146,8 @@ def potential (φ : higgsField) (μSq lambda : ℝ ) (x : spaceTime) : ℝ :=
lemma potential_smooth (φ : higgsField) (μSq lambda : ℝ) :
Smooth 𝓘(ℝ, spaceTime) 𝓘(ℝ, ℝ) (fun x => φ.potential μSq lambda x) := by
simp only [potential, normSq, neg_mul]
exact Smooth.add
(Smooth.neg (Smooth.smul smooth_const φ.normSq_smooth))
(Smooth.smul (Smooth.smul smooth_const φ.normSq_smooth) φ.normSq_smooth)

exact (smooth_const.smul φ.normSq_smooth).neg.add
((smooth_const.smul φ.normSq_smooth).smul φ.normSq_smooth)

lemma potential_apply (φ : higgsField) (μSq lambda : ℝ) (x : spaceTime) :
(φ.potential μSq lambda) x = higgsVec.potential μSq lambda (φ.toHiggsVec x) := by
Expand All @@ -171,35 +162,20 @@ lemma isConst_of_higgsVec (φ : higgsVec) : φ.toField.isConst := by
intro x _
simp [higgsVec.toField]

lemma isConst_iff_of_higgsVec (Φ : higgsField) : Φ.isConst ↔ ∃ (φ : higgsVec), Φ = φ.toField := by
apply Iff.intro
intro h
use Φ 0
ext x y
rw [← h x 0]
rfl
intro h
intro x y
obtain ⟨φ, hφ⟩ := h
subst hφ
rfl
lemma isConst_iff_of_higgsVec (Φ : higgsField) : Φ.isConst ↔ ∃ (φ : higgsVec), Φ = φ.toField :=
Iff.intro (fun h ↦ ⟨Φ 0, by ext x y; rw [← h x 0]; rfl⟩) (fun ⟨φ, hφ⟩ x y ↦ by subst hφ; rfl)

lemma normSq_of_higgsVec (φ : higgsVec) : φ.toField.normSq = fun x => (norm φ) ^ 2 := by
simp only [normSq, higgsVec.toField]
funext x
simp
simp [normSq, higgsVec.toField]

lemma potential_of_higgsVec (φ : higgsVec) (μSq lambda : ℝ) :
φ.toField.potential μSq lambda = fun _ => higgsVec.potential μSq lambda φ := by
simp [higgsVec.potential]
unfold potential
rw [normSq_of_higgsVec]
funext x
simp only [neg_mul, add_right_inj]
ring_nf



end higgsField
end
end StandardModel
Loading

0 comments on commit 1cb2cdf

Please sign in to comment.