From 0634fac03bbeabf412fd6007b25aed6c1b8a63f6 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Fri, 12 Jul 2024 09:58:40 -0400 Subject: [PATCH 1/5] feat: Add double empty Lint --- HepLean/AnomalyCancellation/Basic.lean | 12 ++++++------ HepLean/AnomalyCancellation/GroupActions.lean | 2 +- HepLean/AnomalyCancellation/MSSMNu/B3.lean | 2 +- HepLean/AnomalyCancellation/MSSMNu/Basic.lean | 4 ++-- scripts/double_line_lint.lean | 11 ++++++++++- 5 files changed, 20 insertions(+), 11 deletions(-) diff --git a/HepLean/AnomalyCancellation/Basic.lean b/HepLean/AnomalyCancellation/Basic.lean index 1a298cf..8ea1866 100644 --- a/HepLean/AnomalyCancellation/Basic.lean +++ b/HepLean/AnomalyCancellation/Basic.lean @@ -108,7 +108,7 @@ instance linSolsAddCommMonoid (χ : ACCSystemLinear) : add_zero S := by apply LinSols.ext exact χ.chargesAddCommMonoid.add_zero _ - nsmul n S := ⟨n • S.val, by + nsmul n S := ⟨n • S.val, by intro i rw [nsmul_eq_smul_cast ℚ] erw [(χ.linearACCs i).map_smul, S.linearSol i] @@ -122,7 +122,7 @@ instance linSolsAddCommMonoid (χ : ACCSystemLinear) : /-- An instance providing the operations and properties for `LinSols` to form an module over `ℚ`. -/ @[simps!] -instance linSolsModule (χ : ACCSystemLinear) : Module ℚ χ.LinSols where +instance linSolsModule (χ : ACCSystemLinear) : Module ℚ χ.LinSols where smul a S := ⟨a • S.val, by intro i rw [(χ.linearACCs i).map_smul, S.linearSol i] @@ -177,13 +177,13 @@ structure QuadSols (χ : ACCSystemQuad) extends χ.LinSols where @[ext] lemma QuadSols.ext {χ : ACCSystemQuad} {S T : χ.QuadSols} (h : S.val = T.val) : S = T := by - have h := ACCSystemLinear.LinSols.ext h + have h := ACCSystemLinear.LinSols.ext h cases' S simp_all only /-- An instance giving the properties and structures to define an action of `ℚ` on `QuadSols`. -/ instance quadSolsMulAction (χ : ACCSystemQuad) : MulAction ℚ χ.QuadSols where - smul a S := ⟨a • S.toLinSols , by + smul a S := ⟨a • S.toLinSols , by intro i erw [(χ.quadraticACCs i).map_smul] rw [S.quadSol i] @@ -198,7 +198,7 @@ instance quadSolsMulAction (χ : ACCSystemQuad) : MulAction ℚ χ.QuadSols wher /-- The inclusion of quadratic solutions into linear solutions. -/ def quadSolsInclLinSols (χ : ACCSystemQuad) : χ.QuadSols →[ℚ] χ.LinSols where - toFun := QuadSols.toLinSols + toFun := QuadSols.toLinSols map_smul' _ _ := rfl /-- If there are no quadratic equations (i.e. no U(1)'s in the underlying gauge group. The inclusion @@ -239,7 +239,7 @@ def IsSolution (χ : ACCSystem) (S : χ.Charges) : Prop := /-- An instance giving the properties and structures to define an action of `ℚ` on `Sols`. -/ instance solsMulAction (χ : ACCSystem) : MulAction ℚ χ.Sols where - smul a S := ⟨a • S.toQuadSols , by + smul a S := ⟨a • S.toQuadSols, by erw [(χ.cubicACC).map_smul] rw [S.cubicSol] simp⟩ diff --git a/HepLean/AnomalyCancellation/GroupActions.lean b/HepLean/AnomalyCancellation/GroupActions.lean index 2ed7378..6f63edd 100644 --- a/HepLean/AnomalyCancellation/GroupActions.lean +++ b/HepLean/AnomalyCancellation/GroupActions.lean @@ -100,7 +100,7 @@ instance quadSolAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) : rfl lemma linSolRep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) - (S : χ.QuadSols) : χ.quadSolsInclLinSols (G.quadSolAction.toFun S g) = + (S : χ.QuadSols) : χ.quadSolsInclLinSols (G.quadSolAction.toFun S g) = G.linSolRep g (χ.quadSolsInclLinSols S) := rfl lemma rep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) diff --git a/HepLean/AnomalyCancellation/MSSMNu/B3.lean b/HepLean/AnomalyCancellation/MSSMNu/B3.lean index 7bb1bfc..b2a3b46 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/B3.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/B3.lean @@ -37,7 +37,7 @@ def B₃AsCharge : MSSMACC.Charges := toSpecies.symm | 1, 2 => 1 | 2, 0 => - 1 | 2, 1 => - 1 - | 2, 2 => 1 + | 2, 2 => 1 | 3, 0 => - 3 | 3, 1 => - 3 | 3, 2 => 3 diff --git a/HepLean/AnomalyCancellation/MSSMNu/Basic.lean b/HepLean/AnomalyCancellation/MSSMNu/Basic.lean index 58632ea..2355fa4 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/Basic.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/Basic.lean @@ -50,7 +50,7 @@ def splitSMPlusH : (Fin 18 ⊕ Fin 2 → ℚ) ≃ (Fin 18 → ℚ) × (Fin 2 → /-- An equivalence between `MSSMCharges.charges` and `(Fin 18 → ℚ) × (Fin 2 → ℚ)`. This splits the charges up into the SM and the additional ones for the MSSM. -/ @[simps!] -def toSplitSMPlusH : MSSMCharges.Charges ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) := +def toSplitSMPlusH : MSSMCharges.Charges ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) := toSMPlusH.trans splitSMPlusH /-- An equivalence between `(Fin 18 → ℚ)` and `(Fin 6 → Fin 3 → ℚ)`. -/ @@ -74,7 +74,7 @@ def toSMSpecies (i : Fin 6) : MSSMCharges.Charges →ₗ[ℚ] MSSMSpecies.Charge map_add' _ _ := by rfl map_smul' _ _ := by rfl -lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin 3 → ℚ) × (Fin 2 → ℚ)) : +lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin 3 → ℚ) × (Fin 2 → ℚ)) : (toSMSpecies i) (toSpecies.symm f) = f.1 i := by change (Prod.fst ∘ toSpecies ∘ toSpecies.symm ) _ i= f.1 i simp diff --git a/scripts/double_line_lint.lean b/scripts/double_line_lint.lean index 3b04875..b39e7d8 100644 --- a/scripts/double_line_lint.lean +++ b/scripts/double_line_lint.lean @@ -31,6 +31,15 @@ def doubleEmptyLineLinter : HepLeanTextLinter := fun lines ↦ Id.run do else none) errors.toArray +/-- Checks if there is a souble space in the line, which is not at the start. -/ +def doubleSpaceLinter : HepLeanTextLinter := fun lines ↦ Id.run do + let enumLines := (lines.toList.enumFrom 1) + let errors := enumLines.filterMap (fun (lno, l) ↦ + if String.containsSubstr l.trimLeft " " then + some (s!" Non-initial double space in line. ", lno) + else none) + errors.toArray + structure HepLeanErrorContext where /-- The underlying `message`. -/ error : String @@ -47,7 +56,7 @@ def hepLeanLintFile (path : FilePath) : IO Bool := do let lines ← IO.FS.lines path let allOutput := (Array.map (fun lint ↦ (Array.map (fun (e, n) ↦ HepLeanErrorContext.mk e n path)) (lint lines))) - #[doubleEmptyLineLinter] + #[doubleEmptyLineLinter, doubleSpaceLinter] let errors := allOutput.flatten printErrors errors return errors.size > 0 From 1fe51b2e04806cf399642793302b409479532782 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Fri, 12 Jul 2024 10:23:59 -0400 Subject: [PATCH 2/5] chore: namechange --- .../HiggsBoson/PointwiseInnerProd.lean | 4 +- .../StandardModel/HiggsBoson/Potential.lean | 40 +++++++++---------- lakefile.toml | 2 +- ...line_lint.lean => hepLean_style_lint.lean} | 16 +++++--- 4 files changed, 34 insertions(+), 28 deletions(-) rename scripts/{double_line_lint.lean => hepLean_style_lint.lean} (83%) diff --git a/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean b/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean index 0f53792..c3871fa 100644 --- a/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean +++ b/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean @@ -69,7 +69,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] + 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, RCLike.im_to_complex, I_sq, mul_neg, mul_one, neg_mul, sub_neg_eq_add, one_mul] ring @@ -127,7 +127,7 @@ lemma toHiggsVec_norm (φ : HiggsField) (x : SpaceTime) : ‖φ x‖ = ‖φ.toHiggsVec x‖ := rfl lemma normSq_expand (φ : HiggsField) : - φ.normSq = fun x => (conj (φ x 0) * (φ x 0) + conj (φ x 1) * (φ x 1) ).re := by + φ.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 ℂ] diff --git a/HepLean/StandardModel/HiggsBoson/Potential.lean b/HepLean/StandardModel/HiggsBoson/Potential.lean index 44d4ce7..9b4c392 100644 --- a/HepLean/StandardModel/HiggsBoson/Potential.lean +++ b/HepLean/StandardModel/HiggsBoson/Potential.lean @@ -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 /-! @@ -43,7 +43,7 @@ def potential (μ2 𝓵 : ℝ ) (φ : HiggsField) (x : SpaceTime) : ℝ := -/ -lemma potential_smooth (μSq lambda : ℝ) (φ : HiggsField) : +lemma potential_smooth (μSq lambda : ℝ) (φ : HiggsField) : Smooth 𝓘(ℝ, SpaceTime) 𝓘(ℝ, ℝ) (fun x => φ.potential μSq lambda x) := by simp only [potential, normSq, neg_mul] exact (smooth_const.smul φ.normSq_smooth).neg.add @@ -56,7 +56,7 @@ namespace potential -/ -lemma complete_square (μ2 𝓵 : ℝ) (h : 𝓵 ≠ 0) (φ : HiggsField) (x : SpaceTime) : +lemma complete_square (μ2 𝓵 : ℝ) (h : 𝓵 ≠ 0) (φ : HiggsField) (x : SpaceTime) : potential μ2 𝓵 φ x = 𝓵 * (‖φ‖_H ^ 2 x - μ2 / (2 * 𝓵)) ^ 2 - μ2 ^ 2 / (4 * 𝓵) := by simp only [potential] field_simp @@ -69,7 +69,7 @@ lemma complete_square (μ2 𝓵 : ℝ) (h : 𝓵 ≠ 0) (φ : HiggsField) (x : S -/ /-- The proposition on the coefficents for a potential to be bounded. -/ -def IsBounded (μ2 𝓵 : ℝ) : Prop := +def IsBounded (μ2 𝓵 : ℝ) : Prop := ∃ c, ∀ Φ x, c ≤ potential μ2 𝓵 Φ x /-! TODO: Show when 𝓵 < 0, the potential is not bounded. -/ @@ -87,19 +87,19 @@ variable (h𝓵 : 0 < 𝓵) /-- The second term of the potential is non-negative. -/ lemma snd_term_nonneg (φ : HiggsField) (x : SpaceTime) : - 0 ≤ 𝓵 * ‖φ‖_H ^ 2 x * ‖φ‖_H ^ 2 x := by + 0 ≤ 𝓵 * ‖φ‖_H ^ 2 x * ‖φ‖_H ^ 2 x := by rw [mul_nonneg_iff] apply Or.inl simp_all only [normSq, gt_iff_lt, mul_nonneg_iff_of_pos_left, ge_iff_le, norm_nonneg, pow_nonneg, 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 /-- The discriminant of the quadratic formed by the potential is non-negative. -/ -lemma discrim_nonneg (φ : HiggsField) (x : SpaceTime) : +lemma discrim_nonneg (φ : HiggsField) (x : SpaceTime) : 0 ≤ discrim 𝓵 (- μ2) (- potential μ2 𝓵 φ x) := by have h1 := as_quad μ2 𝓵 φ x rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1 @@ -111,7 +111,7 @@ lemma eq_zero_at (φ : HiggsField) (x : SpaceTime) (hV : potential μ2 𝓵 φ x = 0) : φ x = 0 ∨ ‖φ‖_H ^ 2 x = μ2 / 𝓵 := by have h1 := as_quad μ2 𝓵 φ x rw [hV] at h1 - have h2 : ‖φ‖_H ^ 2 x * (𝓵 * ‖φ‖_H ^ 2 x + - μ2) = 0 := by + have h2 : ‖φ‖_H ^ 2 x * (𝓵 * ‖φ‖_H ^ 2 x + - μ2) = 0 := by linear_combination h1 simp at h2 cases' h2 with h2 h2 @@ -122,12 +122,12 @@ lemma eq_zero_at (φ : HiggsField) (x : SpaceTime) linear_combination h2 lemma eq_zero_at_of_μSq_nonpos {μ2 : ℝ} (hμ2 : μ2 ≤ 0) - (φ : HiggsField) (x : SpaceTime) (hV : potential μ2 𝓵 φ x = 0) : φ x = 0 := by + (φ : HiggsField) (x : SpaceTime) (hV : potential μ2 𝓵 φ x = 0) : φ x = 0 := by cases' (eq_zero_at μ2 h𝓵 φ x hV) with h1 h1 exact h1 by_cases hμSqZ : μ2 = 0 simpa [hμSqZ] using h1 - refine ((?_ : ¬ 0 ≤ μ2 / 𝓵) (?_)).elim + refine ((?_ : ¬ 0 ≤ μ2 / 𝓵) (?_)).elim · simp_all [div_nonneg_iff] intro h exact lt_imp_lt_of_le_imp_le (fun _ => h) (lt_of_le_of_ne hμ2 hμSqZ) @@ -135,7 +135,7 @@ lemma eq_zero_at_of_μSq_nonpos {μ2 : ℝ} (hμ2 : μ2 ≤ 0) exact normSq_nonneg φ x lemma bounded_below (φ : HiggsField) (x : SpaceTime) : - - μ2 ^ 2 / (4 * 𝓵) ≤ potential μ2 𝓵 φ x := by + - μ2 ^ 2 / (4 * 𝓵) ≤ potential μ2 𝓵 φ x := by have h1 := discrim_nonneg μ2 h𝓵 φ x simp only [discrim, even_two, Even.neg_pow, normSq, neg_mul, neg_add_rev, neg_neg] at h1 ring_nf at h1 @@ -145,7 +145,7 @@ lemma bounded_below (φ : HiggsField) (x : SpaceTime) : ring_nf at h2 ⊢ exact h2 -lemma bounded_below_of_μSq_nonpos {μ2 : ℝ} +lemma bounded_below_of_μSq_nonpos {μ2 : ℝ} (hμSq : μ2 ≤ 0) (φ : HiggsField) (x : SpaceTime) : 0 ≤ potential μ2 𝓵 φ x := by refine add_nonneg ?_ (snd_term_nonneg h𝓵 φ x) field_simp [mul_nonpos_iff] @@ -165,7 +165,7 @@ 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 @@ -180,36 +180,36 @@ 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] field_simp ring_nf) -lemma eq_bound_iff_of_μSq_nonpos {μ2 : ℝ} +lemma eq_bound_iff_of_μSq_nonpos {μ2 : ℝ} (hμ2 : μ2 ≤ 0) (φ : HiggsField) (x : SpaceTime) : potential μ2 𝓵 φ x = 0 ↔ φ x = 0 := Iff.intro (fun h ↦ eq_zero_at_of_μSq_nonpos h𝓵 hμ2 φ x h) (fun h ↦ by simp [potential, h]) lemma eq_bound_IsMinOn (φ : HiggsField) (x : SpaceTime) - (hv : potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵)) : + (hv : potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵)) : IsMinOn (fun (φ, x) => potential μ2 𝓵 φ x) Set.univ (φ, x) := by rw [isMinOn_univ_iff] simp only [normSq, neg_mul, le_neg_add_iff_add_le, ge_iff_le, hv] exact fun (φ', x') ↦ bounded_below μ2 h𝓵 φ' x' -lemma eq_bound_IsMinOn_of_μSq_nonpos {μ2 : ℝ} +lemma eq_bound_IsMinOn_of_μSq_nonpos {μ2 : ℝ} (hμ2 : μ2 ≤ 0) (φ : HiggsField) (x : SpaceTime) (hv : potential μ2 𝓵 φ x = 0) : IsMinOn (fun (φ, x) => potential μ2 𝓵 φ x) Set.univ (φ, x) := by rw [isMinOn_univ_iff] simp only [normSq, neg_mul, le_neg_add_iff_add_le, ge_iff_le, hv] exact fun (φ', x') ↦ bounded_below_of_μSq_nonpos h𝓵 hμ2 φ' x' -lemma bound_reached_of_μSq_nonneg {μ2 : ℝ} (hμ2 : 0 ≤ μ2) : +lemma bound_reached_of_μSq_nonneg {μ2 : ℝ} (hμ2 : 0 ≤ μ2) : ∃ (φ : HiggsField) (x : SpaceTime), - potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵) := by + potential μ2 𝓵 φ x = - μ2 ^ 2 / (4 * 𝓵) := by use HiggsVec.toField ![√(μ2/(2 * 𝓵)), 0], 0 refine (eq_bound_iff μ2 h𝓵 (HiggsVec.toField ![√(μ2/(2 * 𝓵)), 0]) 0).mpr ?_ simp only [normSq, HiggsVec.toField, ContMDiffSection.coeFn_mk, PiLp.norm_sq_eq_of_L2, @@ -218,7 +218,7 @@ lemma bound_reached_of_μSq_nonneg {μ2 : ℝ} (hμ2 : 0 ≤ μ2) : not_false_eq_true, zero_pow, add_zero] field_simp [mul_pow] -lemma IsMinOn_iff_of_μSq_nonneg {μ2 : ℝ} (hμ2 : 0 ≤ μ2) : +lemma IsMinOn_iff_of_μSq_nonneg {μ2 : ℝ} (hμ2 : 0 ≤ μ2) : IsMinOn (fun (φ, x) => potential μ2 𝓵 φ x) Set.univ (φ, x) ↔ ‖φ‖_H ^ 2 x = μ2 /(2 * 𝓵) := by apply Iff.intro <;> rw [← eq_bound_iff μ2 h𝓵 φ] diff --git a/lakefile.toml b/lakefile.toml index bf4acda..a035d1f 100644 --- a/lakefile.toml +++ b/lakefile.toml @@ -39,7 +39,7 @@ name = "type_former_lint" srcDir = "scripts" [[lean_exe]] -name = "double_line_lint" +name = "hepLean_style_lint" srcDir = "scripts" [[lean_exe]] diff --git a/scripts/double_line_lint.lean b/scripts/hepLean_style_lint.lean similarity index 83% rename from scripts/double_line_lint.lean rename to scripts/hepLean_style_lint.lean index b39e7d8..96dbfc1 100644 --- a/scripts/double_line_lint.lean +++ b/scripts/hepLean_style_lint.lean @@ -19,7 +19,7 @@ Parts of this file are adapted from `Mathlib.Tactic.Linter.TextBased`, open Lean System Meta /-- Given a list of lines, outputs an error message and a line number. -/ -def HepLeanTextLinter : Type := Array String → Array (String × ℕ) +def HepLeanTextLinter : Type := Array String → Array (String × ℕ × ℕ) /-- Checks if there are two consecutive empty lines. -/ def doubleEmptyLineLinter : HepLeanTextLinter := fun lines ↦ Id.run do @@ -27,7 +27,7 @@ def doubleEmptyLineLinter : HepLeanTextLinter := fun lines ↦ Id.run do let pairLines := List.zip enumLines (List.tail! enumLines) let errors := pairLines.filterMap (fun ((lno1, l1), _, l2) ↦ if l1.length == 0 && l2.length == 0 then - some (s!" Double empty line. ", lno1) + some (s!" Double empty line. ", lno1, 1) else none) errors.toArray @@ -36,7 +36,11 @@ def doubleSpaceLinter : HepLeanTextLinter := fun lines ↦ Id.run do let enumLines := (lines.toList.enumFrom 1) let errors := enumLines.filterMap (fun (lno, l) ↦ if String.containsSubstr l.trimLeft " " then - some (s!" Non-initial double space in line. ", lno) + let k := (Substring.findAllSubstr l " ").toList.getLast? + let col := match k with + | none => 1 + | some k => k.stopPos.byteIdx + some (s!" Non-initial double space in line.", lno, col) else none) errors.toArray @@ -45,17 +49,19 @@ structure HepLeanErrorContext where error : String /-- The line number -/ lineNumber : ℕ + /-- The column number -/ + columnNumber : ℕ /-- The file path -/ path : FilePath def printErrors (errors : Array HepLeanErrorContext) : IO Unit := do for e in errors do - IO.println (s!"error: {e.path}:{e.lineNumber}: {e.error}") + IO.println (s!"error: {e.path}:{e.lineNumber}:{e.columnNumber}: {e.error}") def hepLeanLintFile (path : FilePath) : IO Bool := do let lines ← IO.FS.lines path let allOutput := (Array.map (fun lint ↦ - (Array.map (fun (e, n) ↦ HepLeanErrorContext.mk e n path)) (lint lines))) + (Array.map (fun (e, n, c) ↦ HepLeanErrorContext.mk e n c path)) (lint lines))) #[doubleEmptyLineLinter, doubleSpaceLinter] let errors := allOutput.flatten printErrors errors From 1133b883f35323de1ea42fbd8edce339f2afbe02 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Fri, 12 Jul 2024 10:36:39 -0400 Subject: [PATCH 3/5] refactor: pass at removing double spaces --- .../StandardParameters.lean | 34 +++++++++---------- HepLean/Mathematics/LinearMaps.lean | 6 ++-- HepLean/Mathematics/SO3/Basic.lean | 14 ++++---- HepLean/SpaceTime/LorentzAlgebra/Basic.lean | 12 +++---- HepLean/SpaceTime/LorentzGroup/Basic.lean | 10 +++--- HepLean/SpaceTime/LorentzGroup/Boosts.lean | 2 +- .../SpaceTime/LorentzGroup/Orthochronous.lean | 12 +++---- HepLean/SpaceTime/LorentzGroup/Proper.lean | 14 ++++---- HepLean/SpaceTime/LorentzTensor/Basic.lean | 8 ++--- .../LorentzVector/AsSelfAdjointMatrix.lean | 18 +++++----- HepLean/SpaceTime/LorentzVector/Basic.lean | 2 +- HepLean/SpaceTime/LorentzVector/NormOne.lean | 14 ++++---- HepLean/SpaceTime/MinkowskiMetric.lean | 30 ++++++++-------- HepLean/SpaceTime/SL2C/Basic.lean | 12 +++---- HepLean/StandardModel/HiggsBoson/Basic.lean | 6 ++-- .../StandardModel/HiggsBoson/GaugeAction.lean | 22 ++++++------ .../HiggsBoson/PointwiseInnerProd.lean | 10 +++--- .../StandardModel/HiggsBoson/Potential.lean | 14 ++++---- scripts/hepLean_style_lint.lean | 2 +- 19 files changed, 121 insertions(+), 121 deletions(-) diff --git a/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean b/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean index ef7c1d3..523c70b 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean @@ -40,19 +40,19 @@ def S₂₃ (V : Quotient CKMMatrixSetoid) : ℝ := else VcbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2) /-- Given a CKM matrix `V` the real number corresponding to `θ₁₂` in the -standard parameterization. --/ +standard parameterization. --/ def θ₁₂ (V : Quotient CKMMatrixSetoid) : ℝ := Real.arcsin (S₁₂ V) /-- Given a CKM matrix `V` the real number corresponding to `θ₁₃` in the -standard parameterization. --/ +standard parameterization. --/ def θ₁₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.arcsin (S₁₃ V) /-- Given a CKM matrix `V` the real number corresponding to `θ₂₃` in the -standard parameterization. --/ +standard parameterization. --/ def θ₂₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.arcsin (S₂₃ V) /-- Given a CKM matrix `V` the real number corresponding to `cos θ₁₂` in the -standard parameterization. --/ +standard parameterization. --/ def C₁₂ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₁₂ V) /-- Given a CKM matrix `V` the real number corresponding to `cos θ₁₃` in the @@ -60,7 +60,7 @@ standard parameterization. --/ def C₁₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₁₃ V) /-- Given a CKM matrix `V` the real number corresponding to `sin θ₂₃` in the -standard parameterization. --/ +standard parameterization. --/ def C₂₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₂₃ V) /-- Given a CKM matrix `V` the real number corresponding to the phase `δ₁₃` in the @@ -126,7 +126,7 @@ lemma S₂₃_leq_one (V : Quotient CKMMatrixSetoid) : S₂₃ V ≤ 1 := by lemma S₁₂_eq_sin_θ₁₂ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₁₂ V) = S₁₂ V := Real.sin_arcsin (le_trans (by simp) (S₁₂_nonneg V)) (S₁₂_leq_one V) -lemma S₁₃_eq_sin_θ₁₃ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₁₃ V) = S₁₃ V := +lemma S₁₃_eq_sin_θ₁₃ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₁₃ V) = S₁₃ V := Real.sin_arcsin (le_trans (by simp) (S₁₃_nonneg V)) (S₁₃_leq_one V) lemma S₂₃_eq_sin_θ₂₃ (V : Quotient CKMMatrixSetoid) : Real.sin (θ₂₃ V) = S₂₃ V := @@ -168,7 +168,7 @@ lemma S₂₃_of_Vub_eq_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V = 1) : rw [S₂₃, if_pos ha] lemma S₂₃_of_Vub_neq_one {V : Quotient CKMMatrixSetoid} (ha : VubAbs V ≠ 1) : - S₂₃ V = VcbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2) := by + S₂₃ V = VcbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2) := by rw [S₂₃, if_neg ha] end sines @@ -336,9 +336,9 @@ namespace standParam open Invariant lemma mulExpδ₁₃_on_param_δ₁₃ (V : CKMMatrix) (δ₁₃ : ℝ) : - mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = + mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧) - * cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) * cexp (I * δ₁₃) := by + * cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) * cexp (I * δ₁₃) := by refine mulExpδ₁₃_eq _ _ _ _ ?_ ?_ ?_ ?_ rw [S₁₂_eq_sin_θ₁₂] exact S₁₂_nonneg _ @@ -364,7 +364,7 @@ lemma mulExpδ₁₃_on_param_eq_zero_iff (V : CKMMatrix) (δ₁₃ : ℝ) : aesop lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ℝ) : - Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) = + Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) = sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧) * cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) := by rw [mulExpδ₁₃_on_param_δ₁₃] @@ -373,19 +373,19 @@ lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ℝ) : complexAbs_sin_θ₂₃, complexAbs_cos_θ₂₃] lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : ℝ) - (h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) : + (h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) : cexp (arg ( mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ) * I) = cexp (δ₁₃ * I) := by have h1a := mulExpδ₁₃_on_param_δ₁₃ V δ₁₃ have habs := mulExpδ₁₃_on_param_abs V δ₁₃ - have h2 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = Complex.abs - (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * exp (δ₁₃ * I) := by + have h2 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = Complex.abs + (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * exp (δ₁₃ * I) := by rw [habs, h1a] ring_nf nth_rewrite 1 [← abs_mul_exp_arg_mul_I (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ )] at h2 have habs_neq_zero : - (Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ℂ) ≠ 0 := by + (Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ℂ) ≠ 0 := by simp only [ne_eq, ofReal_eq_zero, map_eq_zero] exact h1 rw [← mul_right_inj' habs_neq_zero] @@ -514,7 +514,7 @@ lemma eq_standParam_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 funext i fin_cases i simp only [uRow, Fin.isValue, Fin.zero_eta, cons_val_zero, standParam, standParamAsMatrix, - ofReal_cos, ofReal_sin, ofReal_neg, mul_neg, neg_mul, neg_neg, cons_val', empty_val', + ofReal_cos, ofReal_sin, ofReal_neg, mul_neg, neg_mul, neg_neg, cons_val', empty_val', cons_val_fin_one, cons_val_one, head_cons, cons_val_two, tail_cons] rw [hV.1, VudAbs_eq_C₁₂_mul_C₁₃ ⟦V⟧] simp [C₁₂, C₁₃] @@ -613,7 +613,7 @@ theorem exists_δ₁₃ (V : CKMMatrix) : rw [hUV] at hna simp [VAbs] at hna simp_all - have hU' := eq_standParam_of_fstRowThdColRealCond haU hU.2 + have hU' := eq_standParam_of_fstRowThdColRealCond haU hU.2 rw [hU'] at hU use (- arg ([U]ub)) rw [← hUV] @@ -639,7 +639,7 @@ theorem eq_standardParameterization_δ₃ (V : CKMMatrix) : obtain ⟨δ₁₃', hδ₃⟩ := exists_δ₁₃ V have hSV := (Quotient.eq.mpr (hδ₃)) by_cases h : Invariant.mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'⟧ ≠ 0 - have h2 := eq_exp_of_phases (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃' + have h2 := eq_exp_of_phases (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃' (δ₁₃ ⟦V⟧) (by rw [← mulExpδ₁₃_on_param_neq_zero_arg V δ₁₃' h, ← hSV, δ₁₃, Invariant.mulExpδ₁₃]) rw [h2] at hδ₃ exact hδ₃ diff --git a/HepLean/Mathematics/LinearMaps.lean b/HepLean/Mathematics/LinearMaps.lean index d6dbf02..aac7ebd 100644 --- a/HepLean/Mathematics/LinearMaps.lean +++ b/HepLean/Mathematics/LinearMaps.lean @@ -77,7 +77,7 @@ def mk₂ (f : V × V → ℚ) (map_smul : ∀ a S T, f (a • S, T) = a * f (S, exact congrArg (HMul.hMul a) (swap T S) } map_smul' := fun a S => LinearMap.ext fun T => map_smul a S T - map_add' := fun S1 S2 => LinearMap.ext fun T => map_add S1 S2 T + map_add' := fun S1 S2 => LinearMap.ext fun T => map_add S1 S2 T swap' := swap lemma map_smul₁ (f : BiLinearSymm V) (a : ℚ) (S T : V) : f (a • S) T = a * f S T := by @@ -99,7 +99,7 @@ lemma map_add₂ (f : BiLinearSymm V) (S : V) (T1 T2 : V) : rw [f.swap, f.map_add₁, f.swap T1 S, f.swap T2 S] /-- Fixing the second input vectors, the resulting linear map. -/ -def toLinear₁ (f : BiLinearSymm V) (T : V) : V →ₗ[ℚ] ℚ where +def toLinear₁ (f : BiLinearSymm V) (T : V) : V →ₗ[ℚ] ℚ where toFun S := f S T map_add' S1 S2 := map_add₁ f S1 S2 T map_smul' a S := by @@ -166,7 +166,7 @@ end HomogeneousCubic /-- The structure of a symmetric trilinear function. -/ structure TriLinearSymm (V : Type) [AddCommMonoid V] [Module ℚ V] extends V →ₗ[ℚ] V →ₗ[ℚ] V →ₗ[ℚ] ℚ where - swap₁' : ∀ S T L, toFun S T L = toFun T S L + swap₁' : ∀ S T L, toFun S T L = toFun T S L swap₂' : ∀ S T L, toFun S T L = toFun S L T namespace TriLinearSymm diff --git a/HepLean/Mathematics/SO3/Basic.lean b/HepLean/Mathematics/SO3/Basic.lean index e31b657..4df983e 100644 --- a/HepLean/Mathematics/SO3/Basic.lean +++ b/HepLean/Mathematics/SO3/Basic.lean @@ -59,7 +59,7 @@ def toGL : SO(3) →* GL (Fin 3) ℝ where map_one' := by simp rfl - map_mul' x y := by + map_mul' x y := by simp only [_root_.mul_inv_rev, coe_inv] ext rfl @@ -143,16 +143,16 @@ instance : TopologicalGroup SO(3) := Inducing.topologicalGroup toGL toGL_embedding.toInducing lemma det_minus_id (A : SO(3)) : det (A.1 - 1) = 0 := by - have h1 : det (A.1 - 1) = - det (A.1 - 1) := + have h1 : det (A.1 - 1) = - det (A.1 - 1) := calc det (A.1 - 1) = det (A.1 - A.1 * A.1ᵀ) := by simp [A.2.2] - _ = det A.1 * det (1 - A.1ᵀ) := by rw [← det_mul, mul_sub, mul_one] + _ = det A.1 * det (1 - A.1ᵀ) := by rw [← det_mul, mul_sub, mul_one] _ = det (1 - A.1ᵀ):= by simp [A.2.1] _ = det (1 - A.1ᵀ)ᵀ := by rw [det_transpose] _ = det (1 - A.1) := by simp _ = det (- (A.1 - 1)) := by simp - _ = (- 1) ^ 3 * det (A.1 - 1) := by simp only [det_neg, Fintype.card_fin, neg_mul, one_mul] - _ = - det (A.1 - 1) := by simp [pow_three] + _ = (- 1) ^ 3 * det (A.1 - 1) := by simp only [det_neg, Fintype.card_fin, neg_mul, one_mul] + _ = - det (A.1 - 1) := by simp [pow_three] simpa using h1 @[simp] @@ -161,7 +161,7 @@ lemma det_id_minus (A : SO(3)) : det (1 - A.1) = 0 := by calc det (1 - A.1) = det (- (A.1 - 1)) := by simp _ = (- 1) ^ 3 * det (A.1 - 1) := by simp only [det_neg, Fintype.card_fin, neg_mul, one_mul] - _ = - det (A.1 - 1) := by simp [pow_three] + _ = - det (A.1 - 1) := by simp [pow_three] rw [h1, det_minus_id] simp only [neg_zero] @@ -216,7 +216,7 @@ lemma exists_basis_preserved (A : SO(3)) : obtain ⟨v, hv⟩ := exists_stationary_vec A have h3 : FiniteDimensional.finrank ℝ (EuclideanSpace ℝ (Fin 3)) = Fintype.card (Fin 3) := by simp_all only [toEnd_apply, finrank_euclideanSpace, Fintype.card_fin] - obtain ⟨b, hb⟩ := Orthonormal.exists_orthonormalBasis_extension_of_card_eq h3 hv.1 + obtain ⟨b, hb⟩ := Orthonormal.exists_orthonormalBasis_extension_of_card_eq h3 hv.1 simp at hb use b rw [hb, hv.2] diff --git a/HepLean/SpaceTime/LorentzAlgebra/Basic.lean b/HepLean/SpaceTime/LorentzAlgebra/Basic.lean index c585a52..fddd9ef 100644 --- a/HepLean/SpaceTime/LorentzAlgebra/Basic.lean +++ b/HepLean/SpaceTime/LorentzAlgebra/Basic.lean @@ -13,7 +13,7 @@ We define - Define `lorentzAlgebra` via `LieAlgebra.Orthogonal.so'` as a subalgebra of `Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ`. - In `mem_iff` prove that a matrix is in the Lorentz algebra if and only if it satisfies the - condition `Aᵀ * η = - η * A`. + condition `Aᵀ * η = - η * A`. -/ @@ -21,7 +21,7 @@ namespace SpaceTime open Matrix open TensorProduct -/-- The Lorentz algebra as a subalgebra of `Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ`. -/ +/-- The Lorentz algebra as a subalgebra of `Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ`. -/ def lorentzAlgebra : LieSubalgebra ℝ (Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) := (LieAlgebra.Orthogonal.so' (Fin 1) (Fin 3) ℝ) @@ -34,7 +34,7 @@ lemma transpose_eta (A : lorentzAlgebra) : A.1ᵀ * η = - η * A.1 := by simpa [LieAlgebra.Orthogonal.so', IsSkewAdjoint, IsAdjointPair] using h1 lemma mem_of_transpose_eta_eq_eta_mul_self {A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ} - (h : Aᵀ * η = - η * A) : A ∈ lorentzAlgebra := by + (h : Aᵀ * η = - η * A) : A ∈ lorentzAlgebra := by erw [mem_skewAdjointMatricesLieSubalgebra] simpa [LieAlgebra.Orthogonal.so', IsSkewAdjoint, IsAdjointPair] using h @@ -42,8 +42,8 @@ lemma mem_iff {A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ} : A ∈ lorentzAlgebra ↔ Aᵀ * η = - η * A := Iff.intro (fun h => transpose_eta ⟨A, h⟩) (fun h => mem_of_transpose_eta_eq_eta_mul_self h) -lemma mem_iff' (A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) : - A ∈ lorentzAlgebra ↔ A = - η * Aᵀ * η := by +lemma mem_iff' (A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) : + A ∈ lorentzAlgebra ↔ A = - η * Aᵀ * η := by rw [mem_iff] refine Iff.intro (fun h => ?_) (fun h => ?_) · trans -η * (Aᵀ * η) @@ -91,7 +91,7 @@ instance lorentzVectorAsLieRingModule : LieRingModule lorentzAlgebra (LorentzVec @[simps!] instance spaceTimeAsLieModule : LieModule ℝ lorentzAlgebra (LorentzVector 3) where - smul_lie r Λ x := by + smul_lie r Λ x := by simp [Bracket.bracket, smul_mulVec_assoc] lie_smul r Λ x := by simp [Bracket.bracket] diff --git a/HepLean/SpaceTime/LorentzGroup/Basic.lean b/HepLean/SpaceTime/LorentzGroup/Basic.lean index 75bd9cc..5df1c17 100644 --- a/HepLean/SpaceTime/LorentzGroup/Basic.lean +++ b/HepLean/SpaceTime/LorentzGroup/Basic.lean @@ -44,7 +44,7 @@ scoped[LorentzGroup] notation (name := lorentzGroup_notation) "𝓛" => LorentzG open minkowskiMetric -variable {Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ} +variable {Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ} /-! @@ -52,7 +52,7 @@ variable {Λ Λ' : Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ} -/ -lemma mem_iff_norm : Λ ∈ LorentzGroup d ↔ +lemma mem_iff_norm : Λ ∈ LorentzGroup d ↔ ∀ (x : LorentzVector d), ⟪Λ *ᵥ x, Λ *ᵥ x⟫ₘ = ⟪x, x⟫ₘ := by refine Iff.intro (fun h x => h x x) (fun h x y => ?_) have hp := h (x + y) @@ -78,7 +78,7 @@ lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by rw [mem_iff_on_right, matrix_eq_id_iff] exact forall_comm -lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by +lemma mem_iff_self_mul_dual : Λ ∈ LorentzGroup d ↔ Λ * dual Λ = 1 := by rw [mem_iff_dual_mul_self] exact mul_eq_one_comm @@ -147,7 +147,7 @@ open minkowskiMetric variable {Λ Λ' : LorentzGroup d} -lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= by +lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= by refine (inv_eq_left_inv ?h).symm exact mem_iff_dual_mul_self.mp Λ.2 @@ -190,7 +190,7 @@ def toProd : LorentzGroup d →* (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ (Matrix (Fin 1 ⊕ Fin d) (Fin 1 ⊕ Fin d) ℝ)ᵐᵒᵖ := MonoidHom.comp (Units.embedProduct _) toGL -lemma toProd_eq_transpose_η : toProd Λ = (Λ.1, MulOpposite.op $ minkowskiMetric.dual Λ.1) := rfl +lemma toProd_eq_transpose_η : toProd Λ = (Λ.1, MulOpposite.op $ minkowskiMetric.dual Λ.1) := rfl lemma toProd_injective : Function.Injective (@toProd d) := by intro A B h diff --git a/HepLean/SpaceTime/LorentzGroup/Boosts.lean b/HepLean/SpaceTime/LorentzGroup/Boosts.lean index a79f0db..15552ca 100644 --- a/HepLean/SpaceTime/LorentzGroup/Boosts.lean +++ b/HepLean/SpaceTime/LorentzGroup/Boosts.lean @@ -100,7 +100,7 @@ lemma toMatrix_mulVec (u v : FuturePointing d) (x : LorentzVector d) : open minkowskiMatrix LorentzVector in @[simp] lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) : - (toMatrix u v) μ ν = η μ μ * (⟪e μ, e ν⟫ₘ + 2 * ⟪e ν, u⟫ₘ * ⟪e μ, v⟫ₘ + (toMatrix u v) μ ν = η μ μ * (⟪e μ, e ν⟫ₘ + 2 * ⟪e ν, u⟫ₘ * ⟪e μ, v⟫ₘ - ⟪e μ, u + v⟫ₘ * ⟪e ν, u + v⟫ₘ / (1 + ⟪u, v.1.1⟫ₘ)) := by rw [matrix_apply_stdBasis (toMatrix u v) μ ν, toMatrix_mulVec] simp only [genBoost, genBoostAux₁, genBoostAux₂, map_add, smul_add, neg_smul, LinearMap.add_apply, diff --git a/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean b/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean index be1963c..56f7615 100644 --- a/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean +++ b/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean @@ -39,7 +39,7 @@ lemma IsOrthochronous_iff_transpose : IsOrthochronous Λ ↔ IsOrthochronous (transpose Λ) := by rfl lemma IsOrthochronous_iff_ge_one : - IsOrthochronous Λ ↔ 1 ≤ timeComp Λ := by + IsOrthochronous Λ ↔ 1 ≤ timeComp Λ := by rw [IsOrthochronous_iff_futurePointing, NormOneLorentzVector.FuturePointing.mem_iff, NormOneLorentzVector.time_pos_iff] simp only [time, toNormOneLorentzVector, timeVec, Fin.isValue] @@ -68,7 +68,7 @@ def timeCompCont : C(LorentzGroup d, ℝ) := ⟨fun Λ => timeComp Λ , /-- An auxillary function used in the definition of `orthchroMapReal`. -/ def stepFunction : ℝ → ℝ := fun t => if t ≤ -1 then -1 else - if 1 ≤ t then 1 else t + if 1 ≤ t then 1 else t lemma stepFunction_continuous : Continuous stepFunction := by apply Continuous.if ?_ continuous_const (Continuous.if ?_ continuous_const continuous_id) @@ -105,20 +105,20 @@ lemma orthchroMapReal_minus_one_or_one (Λ : LorentzGroup d) : apply Or.inr $ orthchroMapReal_on_IsOrthochronous h apply Or.inl $ orthchroMapReal_on_not_IsOrthochronous h -local notation "ℤ₂" => Multiplicative (ZMod 2) +local notation "ℤ₂" => Multiplicative (ZMod 2) /-- A continuous map from `lorentzGroup` to `ℤ₂` whose kernal are the Orthochronous elements. -/ def orthchroMap : C(LorentzGroup d, ℤ₂) := ContinuousMap.comp coeForℤ₂ { toFun := fun Λ => ⟨orthchroMapReal Λ, orthchroMapReal_minus_one_or_one Λ⟩, - continuous_toFun := Continuous.subtype_mk (ContinuousMap.continuous orthchroMapReal) _} + continuous_toFun := Continuous.subtype_mk (ContinuousMap.continuous orthchroMapReal) _} lemma orthchroMap_IsOrthochronous {Λ : LorentzGroup d} (h : IsOrthochronous Λ) : orthchroMap Λ = 1 := by simp [orthchroMap, orthchroMapReal_on_IsOrthochronous h] lemma orthchroMap_not_IsOrthochronous {Λ : LorentzGroup d} (h : ¬ IsOrthochronous Λ) : - orthchroMap Λ = Additive.toMul (1 : ZMod 2) := by + orthchroMap Λ = Additive.toMul (1 : ZMod 2) := by simp [orthchroMap, orthchroMapReal_on_not_IsOrthochronous h] rfl @@ -136,7 +136,7 @@ lemma mul_othchron_of_not_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : rw [IsOrthochronous, timeComp_mul] exact NormOneLorentzVector.FuturePointing.metric_reflect_not_mem_not_mem h h' -lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ) +lemma mul_not_othchron_of_othchron_not_othchron {Λ Λ' : LorentzGroup d} (h : IsOrthochronous Λ) (h' : ¬ IsOrthochronous Λ') : ¬ IsOrthochronous (Λ * Λ') := by rw [not_orthochronous_iff_le_zero, timeComp_mul] rw [IsOrthochronous_iff_transpose] at h diff --git a/HepLean/SpaceTime/LorentzGroup/Proper.lean b/HepLean/SpaceTime/LorentzGroup/Proper.lean index 99aeb0e..3e76268 100644 --- a/HepLean/SpaceTime/LorentzGroup/Proper.lean +++ b/HepLean/SpaceTime/LorentzGroup/Proper.lean @@ -42,7 +42,7 @@ instance : TopologicalGroup ℤ₂ := TopologicalGroup.mk /-- A continuous function from `({-1, 1} : Set ℝ)` to `ℤ₂`. -/ @[simps!] -def coeForℤ₂ : C(({-1, 1} : Set ℝ), ℤ₂) where +def coeForℤ₂ : C(({-1, 1} : Set ℝ), ℤ₂) where toFun x := if x = ⟨1, Set.mem_insert_of_mem (-1) rfl⟩ then (Additive.toMul 0) else (Additive.toMul (1 : ZMod 2)) continuous_toFun := by @@ -50,7 +50,7 @@ def coeForℤ₂ : C(({-1, 1} : Set ℝ), ℤ₂) where exact continuous_of_discreteTopology /-- The continuous map taking a Lorentz matrix to its determinant. -/ -def detContinuous : C(𝓛 d, ℤ₂) := +def detContinuous : C(𝓛 d, ℤ₂) := ContinuousMap.comp coeForℤ₂ { toFun := fun Λ => ⟨Λ.1.det, Or.symm (LorentzGroup.det_eq_one_or_neg_one _)⟩, continuous_toFun := by @@ -64,7 +64,7 @@ lemma detContinuous_eq_iff_det_eq (Λ Λ' : LorentzGroup d) : apply Iff.intro intro h simp [detContinuous] at h - cases' det_eq_one_or_neg_one Λ with h1 h1 + cases' det_eq_one_or_neg_one Λ with h1 h1 <;> cases' det_eq_one_or_neg_one Λ' with h2 h2 <;> simp_all [h1, h2, h] rw [← toMul_zero, @Equiv.apply_eq_iff_eq] at h @@ -92,16 +92,16 @@ def detRep : 𝓛 d →* ℤ₂ where lemma detRep_continuous : Continuous (@detRep d) := detContinuous.2 -lemma det_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) : +lemma det_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) : Λ.1.det = Λ'.1.det := by obtain ⟨s, hs, hΛ'⟩ := h let f : ContinuousMap s ℤ₂ := ContinuousMap.restrict s detContinuous haveI : PreconnectedSpace s := isPreconnected_iff_preconnectedSpace.mp hs.1 simpa [f, detContinuous_eq_iff_det_eq] using (@IsPreconnected.subsingleton ℤ₂ _ _ _ (isPreconnected_range f.2)) - (Set.mem_range_self ⟨Λ, hs.2⟩) (Set.mem_range_self ⟨Λ', hΛ'⟩) + (Set.mem_range_self ⟨Λ, hs.2⟩) (Set.mem_range_self ⟨Λ', hΛ'⟩) -lemma detRep_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) : +lemma detRep_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) : detRep Λ = detRep Λ' := by simp [detRep_apply, detRep_apply, detContinuous] rw [det_on_connected_component h] @@ -125,7 +125,7 @@ lemma IsProper_iff (Λ : LorentzGroup d) : IsProper Λ ↔ detRep Λ = 1 := by lemma id_IsProper : (@IsProper d) 1 := by simp [IsProper] -lemma isProper_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) : +lemma isProper_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) : IsProper Λ ↔ IsProper Λ' := by simp [detRep_apply, detRep_apply, detContinuous] rw [det_on_connected_component h] diff --git a/HepLean/SpaceTime/LorentzTensor/Basic.lean b/HepLean/SpaceTime/LorentzTensor/Basic.lean index 25b9f2c..c60ae0c 100644 --- a/HepLean/SpaceTime/LorentzTensor/Basic.lean +++ b/HepLean/SpaceTime/LorentzTensor/Basic.lean @@ -64,7 +64,7 @@ lemma indexType_eq {T₁ T₂ : RealLorentzTensor d X} (h : T₁.color = T₂.co rw [h] lemma ext {T₁ T₂ : RealLorentzTensor d X} (h : T₁.color = T₂.color) - (h' : T₁.coord = T₂.coord ∘ Equiv.cast (indexType_eq h)) : T₁ = T₂ := by + (h' : T₁.coord = T₂.coord ∘ Equiv.cast (indexType_eq h)) : T₁ = T₂ := by cases T₁ cases T₂ simp_all only [IndexType, mk.injEq] @@ -97,7 +97,7 @@ def ch {X : FintypeCat} (x : X) (T : RealLorentzTensor d X) : Colors := T.color /-- An equivalence between `X → Fin 1 ⊕ Fin d` and `Y → Fin 1 ⊕ Fin d` given an isomorphism between `X` and `Y`. -/ def congrSetIndexType (d : ℕ) (f : X ≃ Y) (i : X → Colors) : - ((x : X) → ColorsIndex d (i x)) ≃ ((y : Y) → ColorsIndex d ((Equiv.piCongrLeft' _ f) i y)) := + ((x : X) → ColorsIndex d (i x)) ≃ ((y : Y) → ColorsIndex d ((Equiv.piCongrLeft' _ f) i y)) := Equiv.piCongrLeft' _ (f) /-- Given an equivalence of indexing sets, a map on Lorentz tensors. -/ @@ -106,8 +106,8 @@ def congrSetMap (f : X ≃ Y) (T : RealLorentzTensor d X) : RealLorentzTensor d color := (Equiv.piCongrLeft' _ f) T.color coord := (Equiv.piCongrLeft' _ (congrSetIndexType d f T.color)) T.coord -lemma congrSetMap_trans (f : X ≃ Y) (g : Y ≃ Z) (T : RealLorentzTensor d X) : - congrSetMap g (congrSetMap f T) = congrSetMap (f.trans g) T := by +lemma congrSetMap_trans (f : X ≃ Y) (g : Y ≃ Z) (T : RealLorentzTensor d X) : + congrSetMap g (congrSetMap f T) = congrSetMap (f.trans g) T := by apply ext (by rfl) have h1 : (congrSetIndexType d (f.trans g) T.color) = (congrSetIndexType d f T.color).trans (congrSetIndexType d g ((Equiv.piCongrLeft' (fun _ => Colors) f) T.color)) := by diff --git a/HepLean/SpaceTime/LorentzVector/AsSelfAdjointMatrix.lean b/HepLean/SpaceTime/LorentzVector/AsSelfAdjointMatrix.lean index eeb7385..a78a117 100644 --- a/HepLean/SpaceTime/LorentzVector/AsSelfAdjointMatrix.lean +++ b/HepLean/SpaceTime/LorentzVector/AsSelfAdjointMatrix.lean @@ -81,7 +81,7 @@ noncomputable def toSelfAdjointMatrix : rw [← h01, RCLike.conj_eq_re_sub_im] rfl exact conj_eq_iff_re.mp (congrArg (fun M => M 1 1) $ selfAdjoint.mem_iff.mp x.2 ) - map_add' x y := by + map_add' x y := by ext i j : 2 simp only [toSelfAdjointMatrix'_coe, add_apply, ofReal_add, of_apply, cons_val', empty_val', cons_val_fin_one, AddSubmonoid.coe_add, AddSubgroup.coe_toAddSubmonoid, Matrix.add_apply] @@ -109,22 +109,22 @@ noncomputable def toSelfAdjointMatrix : simp only [toSelfAdjointMatrix'_coe, Fin.isValue, of_apply, cons_val', empty_val', cons_val_fin_one, RingHom.id_apply, selfAdjoint.val_smul, smul_apply, real_smul] fin_cases i <;> fin_cases j - · rw [show (r • x) (Sum.inl 0) = r * x (Sum.inl 0) from rfl] - rw [show (r • x) (Sum.inr 2) = r * x (Sum.inr 2) from rfl] + · rw [show (r • x) (Sum.inl 0) = r * x (Sum.inl 0) from rfl] + rw [show (r • x) (Sum.inr 2) = r * x (Sum.inr 2) from rfl] simp only [Fin.isValue, ofReal_mul, Fin.zero_eta, cons_val_zero] ring - · rw [show (r • x) (Sum.inr 0) = r * x (Sum.inr 0) from rfl] - rw [show (r • x) (Sum.inr 1) = r * x (Sum.inr 1) from rfl] + · rw [show (r • x) (Sum.inr 0) = r * x (Sum.inr 0) from rfl] + rw [show (r • x) (Sum.inr 1) = r * x (Sum.inr 1) from rfl] simp only [Fin.isValue, ofReal_mul, Fin.mk_one, cons_val_one, head_cons, Fin.zero_eta, cons_val_zero] ring - · rw [show (r • x) (Sum.inr 0) = r * x (Sum.inr 0) from rfl] - rw [show (r • x) (Sum.inr 1) = r * x (Sum.inr 1) from rfl] + · rw [show (r • x) (Sum.inr 0) = r * x (Sum.inr 0) from rfl] + rw [show (r • x) (Sum.inr 1) = r * x (Sum.inr 1) from rfl] simp only [Fin.isValue, ofReal_mul, Fin.zero_eta, cons_val_zero, Fin.mk_one, cons_val_one, head_fin_const] ring - · rw [show (r • x) (Sum.inl 0) = r * x (Sum.inl 0) from rfl] - rw [show (r • x) (Sum.inr 2) = r * x (Sum.inr 2) from rfl] + · rw [show (r • x) (Sum.inl 0) = r * x (Sum.inl 0) from rfl] + rw [show (r • x) (Sum.inr 2) = r * x (Sum.inr 2) from rfl] simp only [Fin.isValue, ofReal_mul, Fin.mk_one, cons_val_one, head_cons, head_fin_const] ring diff --git a/HepLean/SpaceTime/LorentzVector/Basic.lean b/HepLean/SpaceTime/LorentzVector/Basic.lean index 060e348..4202191 100644 --- a/HepLean/SpaceTime/LorentzVector/Basic.lean +++ b/HepLean/SpaceTime/LorentzVector/Basic.lean @@ -43,7 +43,7 @@ variable (v : LorentzVector d) /-- The space components. -/ @[simp] -def space : EuclideanSpace ℝ (Fin d) := v ∘ Sum.inr +def space : EuclideanSpace ℝ (Fin d) := v ∘ Sum.inr /-- The time component. -/ @[simp] diff --git a/HepLean/SpaceTime/LorentzVector/NormOne.lean b/HepLean/SpaceTime/LorentzVector/NormOne.lean index 905e448..a028227 100644 --- a/HepLean/SpaceTime/LorentzVector/NormOne.lean +++ b/HepLean/SpaceTime/LorentzVector/NormOne.lean @@ -36,7 +36,7 @@ def neg : NormOneLorentzVector d := ⟨- v, by simp only [map_neg, LinearMap.neg_apply, neg_neg] exact v.2⟩ -lemma time_sq : v.1.time ^ 2 = 1 + ‖v.1.space‖ ^ 2 := by +lemma time_sq : v.1.time ^ 2 = 1 + ‖v.1.space‖ ^ 2 := by rw [time_sq_eq_metric_add_space, v.2] lemma abs_time_ge_one : 1 ≤ |v.1.time| := by @@ -48,7 +48,7 @@ lemma norm_space_le_abs_time : ‖v.1.space‖ < |v.1.time| := by rw [(abs_norm _).symm, ← @sq_lt_sq, time_sq] exact lt_one_add (‖(v.1).space‖ ^ 2) -lemma norm_space_leq_abs_time : ‖v.1.space‖ ≤ |v.1.time| := +lemma norm_space_leq_abs_time : ‖v.1.space‖ ≤ |v.1.time| := le_of_lt (norm_space_le_abs_time v) lemma time_le_minus_one_or_ge_one : v.1.time ≤ -1 ∨ 1 ≤ v.1.time := @@ -77,7 +77,7 @@ lemma time_pos_iff : 0 < v.1.time ↔ 1 ≤ v.1.time := by · exact (time_nonneg_iff v).mp (le_of_lt h) · linarith -lemma time_abs_sub_space_norm : +lemma time_abs_sub_space_norm : 0 ≤ |v.1.time| * |w.1.time| - ‖v.1.space‖ * ‖w.1.space‖ := by apply sub_nonneg.mpr apply mul_le_mul (norm_space_leq_abs_time v) (norm_space_leq_abs_time w) ?_ ?_ @@ -167,7 +167,7 @@ lemma metric_reflect_mem_mem (h : v ∈ FuturePointing d) (hw : w ∈ FuturePoin exact abs_real_inner_le_norm (v.1).space (w.1).space lemma metric_reflect_not_mem_not_mem (h : v ∉ FuturePointing d) (hw : w ∉ FuturePointing d) : - 0 ≤ ⟪v.1, w.1.spaceReflection⟫ₘ := by + 0 ≤ ⟪v.1, w.1.spaceReflection⟫ₘ := by have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) ((not_mem_iff_neg w).mp hw) apply le_of_le_of_eq h1 ?_ simp [neg] @@ -179,10 +179,10 @@ lemma metric_reflect_mem_not_mem (h : v ∈ FuturePointing d) (hw : w ∉ Future apply le_of_le_of_eq h1 ?_ simp [neg] -lemma metric_reflect_not_mem_mem (h : v ∉ FuturePointing d) (hw : w ∈ FuturePointing d) : +lemma metric_reflect_not_mem_mem (h : v ∉ FuturePointing d) (hw : w ∈ FuturePointing d) : ⟪v.1, w.1.spaceReflection⟫ₘ ≤ 0 := by rw [show (0 : ℝ) = - 0 from zero_eq_neg.mpr rfl, le_neg] - have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) hw + have h1 := metric_reflect_mem_mem ((not_mem_iff_neg v).mp h) hw apply le_of_le_of_eq h1 ?_ simp [neg] @@ -264,7 +264,7 @@ noncomputable def pathFromTime (u : FuturePointing d) : Path timeVecNormOneFutur lemma isPathConnected : IsPathConnected (@Set.univ (FuturePointing d)) := by use timeVecNormOneFuture - apply And.intro trivial ?_ + apply And.intro trivial ?_ intro y a use pathFromTime y exact fun _ => a diff --git a/HepLean/SpaceTime/MinkowskiMetric.lean b/HepLean/SpaceTime/MinkowskiMetric.lean index 5711fbf..d29c21d 100644 --- a/HepLean/SpaceTime/MinkowskiMetric.lean +++ b/HepLean/SpaceTime/MinkowskiMetric.lean @@ -35,7 +35,7 @@ variable {d : ℕ} scoped[minkowskiMatrix] notation "η" => minkowskiMatrix @[simp] -lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by +lemma sq : @minkowskiMatrix d * minkowskiMatrix = 1 := by simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal] ext1 i j rcases i with i | i <;> rcases j with j | j @@ -58,7 +58,7 @@ lemma eq_transpose : minkowskiMatrixᵀ = @minkowskiMatrix d := by lemma det_eq_neg_one_pow_d : (@minkowskiMatrix d).det = (- 1) ^ d := by simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal] -lemma as_block : @minkowskiMatrix d = ( +lemma as_block : @minkowskiMatrix d = ( Matrix.fromBlocks (1 : Matrix (Fin 1) (Fin 1) ℝ) 0 0 (-1 : Matrix (Fin d) (Fin d) ℝ)) := by rw [minkowskiMatrix] simp [LieAlgebra.Orthogonal.indefiniteDiagonal] @@ -77,7 +77,7 @@ end minkowskiMatrix -/ -/-- Given a Lorentz vector `v` we define the the linear map `w ↦ v * η * w` -/ +/-- Given a Lorentz vector `v` we define the the linear map `w ↦ v * η * w`. -/ @[simps!] def minkowskiLinearForm {d : ℕ} (v : LorentzVector d) : LorentzVector d →ₗ[ℝ] ℝ where toFun w := v ⬝ᵥ (minkowskiMatrix *ᵥ w) @@ -90,7 +90,7 @@ def minkowskiLinearForm {d : ℕ} (v : LorentzVector d) : LorentzVector d →ₗ rfl /-- The Minkowski metric as a bilinear map. -/ -def minkowskiMetric {d : ℕ} : LorentzVector d →ₗ[ℝ] LorentzVector d →ₗ[ℝ] ℝ where +def minkowskiMetric {d : ℕ} : LorentzVector d →ₗ[ℝ] LorentzVector d →ₗ[ℝ] ℝ where toFun v := minkowskiLinearForm v map_add' y z := by ext1 x @@ -134,7 +134,7 @@ lemma as_sum_self : ⟪v, v⟫ₘ = v.time ^ 2 - ‖v.space‖ ^ 2 := by rw [← real_inner_self_eq_norm_sq, PiLp.inner_apply, as_sum] noncomm_ring -lemma eq_time_minus_inner_prod : ⟪v, w⟫ₘ = v.time * w.time - ⟪v.space, w.space⟫_ℝ := by +lemma eq_time_minus_inner_prod : ⟪v, w⟫ₘ = v.time * w.time - ⟪v.space, w.space⟫_ℝ := by rw [as_sum, @PiLp.inner_apply] noncomm_ring @@ -153,7 +153,7 @@ lemma time_sq_eq_metric_add_space : v.time ^ 2 = ⟪v, v⟫ₘ + ‖v.space‖ ^ /-! -# Minkowski metric and space reflections +# Minkowski metric and space reflections -/ @@ -195,7 +195,7 @@ lemma nondegenerate : (∀ w, ⟪w, v⟫ₘ = 0) ↔ v = 0 := by -/ -lemma leq_time_sq : ⟪v, v⟫ₘ ≤ v.time ^ 2 := by +lemma leq_time_sq : ⟪v, v⟫ₘ ≤ v.time ^ 2 := by rw [time_sq_eq_metric_add_space] exact (le_add_iff_nonneg_right _).mpr $ sq_nonneg ‖v.space‖ @@ -227,7 +227,7 @@ lemma dual_id : @dual d 1 = 1 := by @[simp] lemma dual_mul : dual (Λ * Λ') = dual Λ' * dual Λ := by simp only [dual, transpose_mul] - trans η * Λ'ᵀ * (η * η) * Λᵀ * η + trans η * Λ'ᵀ * (η * η) * Λᵀ * η noncomm_ring [minkowskiMatrix.sq] noncomm_ring @@ -256,7 +256,7 @@ lemma det_dual : (dual Λ).det = Λ.det := by simp @[simp] -lemma dual_mulVec_right : ⟪x, (dual Λ) *ᵥ y⟫ₘ = ⟪Λ *ᵥ x, y⟫ₘ := by +lemma dual_mulVec_right : ⟪x, (dual Λ) *ᵥ y⟫ₘ = ⟪Λ *ᵥ x, y⟫ₘ := by simp only [minkowskiMetric, LinearMap.coe_mk, AddHom.coe_mk, dual, minkowskiLinearForm_apply, mulVec_mulVec, ← mul_assoc, minkowskiMatrix.sq, one_mul, (vecMul_transpose Λ x).symm, ← dotProduct_mulVec] @@ -265,7 +265,7 @@ lemma dual_mulVec_right : ⟪x, (dual Λ) *ᵥ y⟫ₘ = ⟪Λ *ᵥ x, y⟫ₘ lemma dual_mulVec_left : ⟪(dual Λ) *ᵥ x, y⟫ₘ = ⟪x, Λ *ᵥ y⟫ₘ := by rw [symm, dual_mulVec_right, symm] -lemma matrix_apply_eq_iff_sub : ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ ↔ ⟪v, (Λ - Λ') *ᵥ w⟫ₘ = 0 := by +lemma matrix_apply_eq_iff_sub : ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ ↔ ⟪v, (Λ - Λ') *ᵥ w⟫ₘ = 0 := by rw [← sub_eq_zero, ← LinearMap.map_sub, sub_mulVec] lemma matrix_eq_iff_eq_forall' : (∀ v, Λ *ᵥ v= Λ' *ᵥ v) ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by @@ -278,7 +278,7 @@ lemma matrix_eq_iff_eq_forall' : (∀ v, Λ *ᵥ v= Λ' *ᵥ v) ↔ ∀ w v, ⟪ have h1 := h v rw [nondegenerate] at h1 simp [sub_mulVec] at h1 - exact h1 + exact h1 lemma matrix_eq_iff_eq_forall : Λ = Λ' ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, Λ' *ᵥ w⟫ₘ := by rw [← matrix_eq_iff_eq_forall'] @@ -302,7 +302,7 @@ lemma matrix_eq_id_iff : Λ = 1 ↔ ∀ w v, ⟪v, Λ *ᵥ w⟫ₘ = ⟪v, w⟫ -/ @[simp] -lemma basis_left (μ : Fin 1 ⊕ Fin d) : ⟪e μ, v⟫ₘ = η μ μ * v μ := by +lemma basis_left (μ : Fin 1 ⊕ Fin d) : ⟪e μ, v⟫ₘ = η μ μ * v μ := by rw [as_sum] rcases μ with μ | μ · fin_cases μ @@ -314,7 +314,7 @@ lemma on_timeVec : ⟪timeVec, @timeVec d⟫ₘ = 1 := by LieAlgebra.Orthogonal.indefiniteDiagonal, diagonal_apply_eq, Sum.elim_inl, stdBasis_apply, ↓reduceIte, mul_one] -lemma on_basis_mulVec (μ ν : Fin 1 ⊕ Fin d) : ⟪e μ, Λ *ᵥ e ν⟫ₘ = η μ μ * Λ μ ν := by +lemma on_basis_mulVec (μ ν : Fin 1 ⊕ Fin d) : ⟪e μ, Λ *ᵥ e ν⟫ₘ = η μ μ * Λ μ ν := by simp [basis_left, mulVec, dotProduct, stdBasis_apply] lemma on_basis (μ ν : Fin 1 ⊕ Fin d) : ⟪e μ, e ν⟫ₘ = η μ ν := by @@ -324,8 +324,8 @@ lemma on_basis (μ ν : Fin 1 ⊕ Fin d) : ⟪e μ, e ν⟫ₘ = η μ ν := by · simp [h, LieAlgebra.Orthogonal.indefiniteDiagonal, minkowskiMatrix] exact fun a => False.elim (h (id (Eq.symm a))) -lemma matrix_apply_stdBasis (ν μ : Fin 1 ⊕ Fin d): - Λ ν μ = η ν ν * ⟪e ν, Λ *ᵥ e μ⟫ₘ := by +lemma matrix_apply_stdBasis (ν μ : Fin 1 ⊕ Fin d): + Λ ν μ = η ν ν * ⟪e ν, Λ *ᵥ e μ⟫ₘ := by rw [on_basis_mulVec, ← mul_assoc] have h1 : η ν ν * η ν ν = 1 := by simp [minkowskiMatrix, LieAlgebra.Orthogonal.indefiniteDiagonal] diff --git a/HepLean/SpaceTime/SL2C/Basic.lean b/HepLean/SpaceTime/SL2C/Basic.lean index 352fe81..ee4399f 100644 --- a/HepLean/SpaceTime/SL2C/Basic.lean +++ b/HepLean/SpaceTime/SL2C/Basic.lean @@ -33,8 +33,8 @@ we can define a representation a representation of `SL(2, ℂ)` on spacetime. -/ -/-- Given an element `M ∈ SL(2, ℂ)` the linear map from `selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)` to - itself defined by `A ↦ M * A * Mᴴ`. -/ +/-- Given an element `M ∈ SL(2, ℂ)` the linear map from `selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)` to + itself defined by `A ↦ M * A * Mᴴ`. -/ @[simps!] def toLinearMapSelfAdjointMatrix (M : SL(2, ℂ)) : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ) →ₗ[ℝ] selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ) where @@ -50,7 +50,7 @@ def toLinearMapSelfAdjointMatrix (M : SL(2, ℂ)) : noncomm_ring [selfAdjoint.val_smul, Algebra.mul_smul_comm, Algebra.smul_mul_assoc, RingHom.id_apply] -/-- The representation of `SL(2, ℂ)` on `selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)` given by +/-- The representation of `SL(2, ℂ)` on `selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)` given by `M A ↦ M * A * Mᴴ`. -/ @[simps!] def repSelfAdjointMatrix : Representation ℝ SL(2, ℂ) $ selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ) where @@ -63,7 +63,7 @@ def repSelfAdjointMatrix : Representation ℝ SL(2, ℂ) $ selfAdjoint (Matrix ( noncomm_ring [toLinearMapSelfAdjointMatrix, SpecialLinearGroup.coe_mul, mul_assoc, conjTranspose_mul, LinearMap.coe_mk, AddHom.coe_mk, LinearMap.mul_apply] -/-- The representation of `SL(2, ℂ)` on `spaceTime` obtained from `toSelfAdjointMatrix` and +/-- The representation of `SL(2, ℂ)` on `spaceTime` obtained from `toSelfAdjointMatrix` and `repSelfAdjointMatrix`. -/ def repLorentzVector : Representation ℝ SL(2, ℂ) (LorentzVector 3) where toFun M := toSelfAdjointMatrix.symm.comp ((repSelfAdjointMatrix M).comp @@ -85,7 +85,7 @@ In the next section we will restrict this homomorphism to the restricted Lorentz -/ -lemma iff_det_selfAdjoint (Λ : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ): Λ ∈ LorentzGroup 3 ↔ +lemma iff_det_selfAdjoint (Λ : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) : Λ ∈ LorentzGroup 3 ↔ ∀ (x : selfAdjoint (Matrix (Fin 2) (Fin 2) ℂ)), det ((toSelfAdjointMatrix ∘ toLin LorentzVector.stdBasis LorentzVector.stdBasis Λ ∘ toSelfAdjointMatrix.symm) x).1 @@ -107,7 +107,7 @@ def toLorentzGroupElem (M : SL(2, ℂ)) : LorentzGroup 3 := /-- The group homomorphism from ` SL(2, ℂ)` to the Lorentz group `𝓛`. -/ @[simps!] -def toLorentzGroup : SL(2, ℂ) →* LorentzGroup 3 where +def toLorentzGroup : SL(2, ℂ) →* LorentzGroup 3 where toFun := toLorentzGroupElem map_one' := by simp only [toLorentzGroupElem, _root_.map_one, LinearMap.toMatrix_one] diff --git a/HepLean/StandardModel/HiggsBoson/Basic.lean b/HepLean/StandardModel/HiggsBoson/Basic.lean index 6ec0721..c6877f1 100644 --- a/HepLean/StandardModel/HiggsBoson/Basic.lean +++ b/HepLean/StandardModel/HiggsBoson/Basic.lean @@ -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) /-! diff --git a/HepLean/StandardModel/HiggsBoson/GaugeAction.lean b/HepLean/StandardModel/HiggsBoson/GaugeAction.lean index 2abcd4c..3c2a359 100644 --- a/HepLean/StandardModel/HiggsBoson/GaugeAction.lean +++ b/HepLean/StandardModel/HiggsBoson/GaugeAction.lean @@ -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 diff --git a/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean b/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean index c3871fa..e94d3d9 100644 --- a/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean +++ b/HepLean/StandardModel/HiggsBoson/PointwiseInnerProd.lean @@ -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 ℂ] diff --git a/HepLean/StandardModel/HiggsBoson/Potential.lean b/HepLean/StandardModel/HiggsBoson/Potential.lean index 9b4c392..4f9b4e7 100644 --- a/HepLean/StandardModel/HiggsBoson/Potential.lean +++ b/HepLean/StandardModel/HiggsBoson/Potential.lean @@ -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] diff --git a/scripts/hepLean_style_lint.lean b/scripts/hepLean_style_lint.lean index 96dbfc1..2cff93a 100644 --- a/scripts/hepLean_style_lint.lean +++ b/scripts/hepLean_style_lint.lean @@ -39,7 +39,7 @@ def doubleSpaceLinter : HepLeanTextLinter := fun lines ↦ Id.run do let k := (Substring.findAllSubstr l " ").toList.getLast? let col := match k with | none => 1 - | some k => k.stopPos.byteIdx + | some k => String.offsetOfPos l k.stopPos some (s!" Non-initial double space in line.", lno, col) else none) errors.toArray From 13f62a50eb392669895146e7d8e4a64c60e16704 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Fri, 12 Jul 2024 11:23:02 -0400 Subject: [PATCH 4/5] reactor: Removal of double spaces --- HepLean/AnomalyCancellation/Basic.lean | 8 +- HepLean/AnomalyCancellation/GroupActions.lean | 4 +- HepLean/AnomalyCancellation/MSSMNu/B3.lean | 2 +- HepLean/AnomalyCancellation/MSSMNu/Basic.lean | 32 ++++---- .../AnomalyCancellation/MSSMNu/LineY3B3.lean | 4 +- .../MSSMNu/OrthogY3B3/Basic.lean | 8 +- .../MSSMNu/OrthogY3B3/PlaneWithY3B3.lean | 16 ++-- .../MSSMNu/OrthogY3B3/ToSols.lean | 12 +-- .../MSSMNu/Permutations.lean | 24 +++--- HepLean/AnomalyCancellation/PureU1/Basic.lean | 10 +-- .../PureU1/BasisLinear.lean | 6 +- .../AnomalyCancellation/PureU1/ConstAbs.lean | 26 +++--- .../PureU1/Even/BasisLinear.lean | 42 +++++----- .../PureU1/Even/LineInCubic.lean | 32 ++++---- .../PureU1/Even/Parameterization.lean | 18 ++-- .../PureU1/LineInPlaneCond.lean | 16 ++-- .../PureU1/LowDim/One.lean | 2 +- .../PureU1/LowDim/Three.lean | 2 +- .../PureU1/LowDim/Two.lean | 2 +- .../PureU1/Odd/BasisLinear.lean | 34 ++++---- .../PureU1/Odd/LineInCubic.lean | 22 ++--- .../PureU1/Odd/Parameterization.lean | 28 +++---- .../PureU1/Permutations.lean | 44 +++++----- HepLean/AnomalyCancellation/PureU1/Sort.lean | 16 ++-- .../PureU1/VectorLike.lean | 10 +-- HepLean/AnomalyCancellation/SM/Basic.lean | 20 ++--- .../AnomalyCancellation/SM/FamilyMaps.lean | 8 +- .../AnomalyCancellation/SM/NoGrav/Basic.lean | 8 +- .../SM/NoGrav/One/Lemmas.lean | 10 +-- .../SM/NoGrav/One/LinearParameterization.lean | 18 ++-- .../AnomalyCancellation/SM/Permutations.lean | 16 ++-- HepLean/AnomalyCancellation/SMNu/Basic.lean | 38 ++++----- .../AnomalyCancellation/SMNu/FamilyMaps.lean | 30 +++---- .../SMNu/NoGrav/Basic.lean | 8 +- .../SMNu/Ordinary/Basic.lean | 12 +-- .../SMNu/Ordinary/DimSevenPlane.lean | 16 ++-- .../SMNu/Permutations.lean | 16 ++-- .../SMNu/PlusU1/Basic.lean | 16 ++-- .../SMNu/PlusU1/HyperCharge.lean | 16 ++-- .../SMNu/PlusU1/PlaneNonSols.lean | 10 +-- .../SMNu/PlusU1/QuadSol.lean | 8 +- .../SMNu/PlusU1/QuadSolToSol.lean | 20 ++--- .../BeyondTheStandardModel/TwoHDM/Basic.lean | 8 +- HepLean/FeynmanDiagrams/Basic.lean | 82 +++++++++---------- .../Instances/ComplexScalar.lean | 10 +-- HepLean/FeynmanDiagrams/Instances/Phi4.lean | 6 +- HepLean/FeynmanDiagrams/Momentum.lean | 20 ++--- HepLean/FlavorPhysics/CKMMatrix/Basic.lean | 16 ++-- .../FlavorPhysics/CKMMatrix/Invariants.lean | 8 +- .../FlavorPhysics/CKMMatrix/PhaseFreedom.lean | 40 ++++----- .../FlavorPhysics/CKMMatrix/Relations.lean | 30 +++---- HepLean/FlavorPhysics/CKMMatrix/Rows.lean | 54 ++++++------ .../StandardParameterization/Basic.lean | 14 ++-- .../StandardParameters.lean | 38 ++++----- HepLean/Mathematics/LinearMaps.lean | 10 +-- HepLean/Mathematics/SO3/Basic.lean | 8 +- HepLean/SpaceTime/LorentzAlgebra/Basic.lean | 8 +- HepLean/SpaceTime/LorentzGroup/Basic.lean | 8 +- HepLean/SpaceTime/LorentzGroup/Boosts.lean | 2 +- .../SpaceTime/LorentzGroup/Orthochronous.lean | 6 +- HepLean/SpaceTime/LorentzGroup/Proper.lean | 6 +- HepLean/SpaceTime/LorentzTensor/Basic.lean | 16 ++-- scripts/README.md | 8 +- scripts/hepLean_style_lint.lean | 8 +- 64 files changed, 550 insertions(+), 546 deletions(-) diff --git a/HepLean/AnomalyCancellation/Basic.lean b/HepLean/AnomalyCancellation/Basic.lean index 8ea1866..3976895 100644 --- a/HepLean/AnomalyCancellation/Basic.lean +++ b/HepLean/AnomalyCancellation/Basic.lean @@ -197,7 +197,7 @@ instance quadSolsMulAction (χ : ACCSystemQuad) : MulAction ℚ χ.QuadSols wher exact one_smul _ _ /-- The inclusion of quadratic solutions into linear solutions. -/ -def quadSolsInclLinSols (χ : ACCSystemQuad) : χ.QuadSols →[ℚ] χ.LinSols where +def quadSolsInclLinSols (χ : ACCSystemQuad) : χ.QuadSols →[ℚ] χ.LinSols where toFun := QuadSols.toLinSols map_smul' _ _ := rfl @@ -229,7 +229,7 @@ structure Sols (χ : ACCSystem) extends χ.QuadSols where /-- Two solutions are equal if the underlying charges are equal. -/ lemma Sols.ext {χ : ACCSystem} {S T : χ.Sols} (h : S.val = T.val) : S = T := by - have h := ACCSystemQuad.QuadSols.ext h + have h := ACCSystemQuad.QuadSols.ext h cases' S simp_all only @@ -251,8 +251,8 @@ instance solsMulAction (χ : ACCSystem) : MulAction ℚ χ.Sols where exact one_smul _ _ /-- The inclusion of `Sols` into `QuadSols`. -/ -def solsInclQuadSols (χ : ACCSystem) : χ.Sols →[ℚ] χ.QuadSols where - toFun := Sols.toQuadSols +def solsInclQuadSols (χ : ACCSystem) : χ.Sols →[ℚ] χ.QuadSols where + toFun := Sols.toQuadSols map_smul' _ _ := rfl /-- The inclusion of `Sols` into `LinSols`. -/ diff --git a/HepLean/AnomalyCancellation/GroupActions.lean b/HepLean/AnomalyCancellation/GroupActions.lean index 6f63edd..be298c2 100644 --- a/HepLean/AnomalyCancellation/GroupActions.lean +++ b/HepLean/AnomalyCancellation/GroupActions.lean @@ -104,7 +104,7 @@ lemma linSolRep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction G.linSolRep g (χ.quadSolsInclLinSols S) := rfl lemma rep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) - (S : χ.QuadSols) : χ.quadSolsIncl (G.quadSolAction.toFun S g) = + (S : χ.QuadSols) : χ.quadSolsIncl (G.quadSolAction.toFun S g) = G.rep g (χ.quadSolsIncl S) := rfl /-- The group action acting on solutions to the anomaly cancellation conditions. -/ @@ -126,7 +126,7 @@ instance solAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) : MulAction G. rfl lemma quadSolAction_solAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) - (S : χ.Sols) : χ.solsInclQuadSols (G.solAction.toFun S g) = + (S : χ.Sols) : χ.solsInclQuadSols (G.solAction.toFun S g) = G.quadSolAction.toFun (χ.solsInclQuadSols S) g := rfl lemma linSolRep_solAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) diff --git a/HepLean/AnomalyCancellation/MSSMNu/B3.lean b/HepLean/AnomalyCancellation/MSSMNu/B3.lean index b2a3b46..68deb5a 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/B3.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/B3.lean @@ -34,7 +34,7 @@ def B₃AsCharge : MSSMACC.Charges := toSpecies.symm | 0, 2 => - 1 | 1, 0 => - 1 | 1, 1 => - 1 - | 1, 2 => 1 + | 1, 2 => 1 | 2, 0 => - 1 | 2, 1 => - 1 | 2, 2 => 1 diff --git a/HepLean/AnomalyCancellation/MSSMNu/Basic.lean b/HepLean/AnomalyCancellation/MSSMNu/Basic.lean index 2355fa4..d347022 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/Basic.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/Basic.lean @@ -43,7 +43,7 @@ def toSMPlusH : MSSMCharges.Charges ≃ (Fin 18 ⊕ Fin 2 → ℚ) := @[simps!] def splitSMPlusH : (Fin 18 ⊕ Fin 2 → ℚ) ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) where toFun f := (f ∘ Sum.inl , f ∘ Sum.inr) - invFun f := Sum.elim f.1 f.2 + invFun f := Sum.elim f.1 f.2 left_inv f := Sum.elim_comp_inl_inr f right_inv _ := rfl @@ -107,7 +107,7 @@ def Hu : MSSMCharges.Charges →ₗ[ℚ] ℚ where map_smul' _ _ := by rfl lemma charges_eq_toSpecies_eq (S T : MSSMCharges.Charges) : - S = T ↔ (∀ i, toSMSpecies i S = toSMSpecies i T) ∧ Hd S = Hd T ∧ Hu S = Hu T := by + S = T ↔ (∀ i, toSMSpecies i S = toSMSpecies i T) ∧ Hd S = Hd T ∧ Hu S = Hu T := by apply Iff.intro intro h rw [h] @@ -159,7 +159,7 @@ def accGrav : MSSMCharges.Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accGrav`. -/ lemma accGrav_ext {S T : MSSMCharges.Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) + (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) (hd : Hd S = Hd T) (hu : Hu S = Hu T) : accGrav S = accGrav T := by simp only [accGrav, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, @@ -192,7 +192,7 @@ def accSU2 : MSSMCharges.Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accSU2`. -/ lemma accSU2_ext {S T : MSSMCharges.Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) + (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) (hd : Hd S = Hd T) (hu : Hu S = Hu T) : accSU2 S = accSU2 T := by simp only [accSU2, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, @@ -224,7 +224,7 @@ def accSU3 : MSSMCharges.Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accSU3`. -/ lemma accSU3_ext {S T : MSSMCharges.Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) : + (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) : accSU3 S = accSU3 T := by simp only [accSU3, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, LinearMap.coe_mk, AddHom.coe_mk] @@ -256,7 +256,7 @@ def accYY : MSSMCharges.Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accGrav`. -/ lemma accYY_ext {S T : MSSMCharges.Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) + (hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) (hd : Hd S = Hd T) (hu : Hu S = Hu T) : accYY S = accYY T := by simp only [accYY, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, @@ -269,9 +269,9 @@ lemma accYY_ext {S T : MSSMCharges.Charges} /-- The symmetric bilinear function used to define the quadratic ACC. -/ @[simps!] -def quadBiLin : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ ( - fun (S, T) => ∑ i, (Q S i * Q T i + (- 2) * (U S i * U T i) + - D S i * D T i + (- 1) * (L S i * L T i) + E S i * E T i) + +def quadBiLin : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ ( + fun (S, T) => ∑ i, (Q S i * Q T i + (- 2) * (U S i * U T i) + + D S i * D T i + (- 1) * (L S i * L T i) + E S i * E T i) + (- Hd S * Hd T + Hu S * Hu T)) (by intro a S T @@ -346,7 +346,7 @@ def cubeTriLinToFun + (2 * Hd S.1 * Hd S.2.1 * Hd S.2.2 + 2 * Hu S.1 * Hu S.2.1 * Hu S.2.2) -lemma cubeTriLinToFun_map_smul₁ (a : ℚ) (S T R : MSSMCharges.Charges) : +lemma cubeTriLinToFun_map_smul₁ (a : ℚ) (S T R : MSSMCharges.Charges) : cubeTriLinToFun (a • S, T, R) = a * cubeTriLinToFun (S, T, R) := by simp only [cubeTriLinToFun] rw [mul_add] @@ -366,7 +366,7 @@ lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.Charges) : rw [add_assoc, ← add_assoc (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _ _] rw [add_comm (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _] rw [add_assoc] - rw [← add_assoc _ _ (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L + + rw [← add_assoc _ _ (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L + (2 * Hd T * Hd R * Hd L + 2 * Hu T * Hu R * Hu L))] congr 1 rw [← Finset.sum_add_distrib] @@ -415,7 +415,7 @@ def accCube : HomogeneousCubic MSSMCharges.Charges := cubeTriLin.toCubic lemma accCube_ext {S T : MSSMCharges.Charges} (h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSMSpecies j S) i = ∑ i, ((fun a => a^3) ∘ toSMSpecies j T) i) - (hd : Hd S = Hd T) (hu : Hu S = Hu T) : + (hd : Hd S = Hd T) (hu : Hu S = Hu T) : accCube S = accCube T := by simp only [HomogeneousCubic, accCube, cubeTriLin, TriLinearSymm.toCubic_apply, TriLinearSymm.mk₃_toFun_apply_apply, cubeTriLinToFun] @@ -452,7 +452,7 @@ def MSSMACC : ACCSystem where namespace MSSMACC open MSSMCharges -lemma quadSol (S : MSSMACC.QuadSols) : accQuad S.val = 0 := by +lemma quadSol (S : MSSMACC.QuadSols) : accQuad S.val = 0 := by have hS := S.quadSol simp only [MSSMACC_numberQuadratic, HomogeneousQuadratic, MSSMACC_quadraticACCs] at hS exact hS 0 @@ -516,9 +516,9 @@ lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols) /-- The dot product on the vector space of charges. -/ @[simps!] -def dot : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ - (fun S => ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i + - D S.1 i * D S.2 i + L S.1 i * L S.2 i + E S.1 i * E S.2 i +def dot : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ + (fun S => ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i + + D S.1 i * D S.2 i + L S.1 i * L S.2 i + E S.1 i * E S.2 i + N S.1 i * N S.2 i) + Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2) (by intro a S T diff --git a/HepLean/AnomalyCancellation/MSSMNu/LineY3B3.lean b/HepLean/AnomalyCancellation/MSSMNu/LineY3B3.lean index 2d2a26e..d4b0a0b 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/LineY3B3.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/LineY3B3.lean @@ -11,7 +11,7 @@ import Mathlib.Tactic.Polyrith # The line through B₃ and Y₃ We give properties of lines through `B₃` and `Y₃`. We show that every point on this line -is a solution to the quadratic `lineY₃B₃Charges_quad` and a double point of the cubic +is a solution to the quadratic `lineY₃B₃Charges_quad` and a double point of the cubic `lineY₃B₃_doublePoint`. # References @@ -52,7 +52,7 @@ lemma lineY₃B₃Charges_cubic (a b : ℚ) : accCube (lineY₃B₃Charges a b). rw [cubeTriLin.toCubic.map_smul] rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₃] rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₃] - erw [Y₃.cubicSol, B₃.cubicSol] + erw [Y₃.cubicSol, B₃.cubicSol] rw [show cubeTriLin Y₃.val Y₃.val B₃.val = 0 by rfl] rw [show cubeTriLin B₃.val B₃.val Y₃.val = 0 by rfl] simp diff --git a/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/Basic.lean b/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/Basic.lean index 721e6e6..5ec2a7d 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/Basic.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/Basic.lean @@ -158,7 +158,7 @@ lemma cube_proj_proj_B₃ (T : MSSMACC.LinSols) : lemma cube_proj_proj_self (T : MSSMACC.Sols) : cubeTriLin (proj T.1.1).val (proj T.1.1).val T.val = 2 * dot Y₃.val B₃.val * - ((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin T.val T.val Y₃.val + + ((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin T.val T.val Y₃.val + ( dot Y₃.val T.val- 2 * dot B₃.val T.val) * cubeTriLin T.val T.val B₃.val) := by rw [proj_val] rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂] @@ -168,13 +168,13 @@ lemma cube_proj_proj_self (T : MSSMACC.Sols) : repeat rw [cubeTriLin.map_add₂] repeat rw [cubeTriLin.map_smul₂] erw [T.cubicSol] - rw [cubeTriLin.swap₁ Y₃.val T.val T.val, cubeTriLin.swap₂ T.val Y₃.val T.val] - rw [cubeTriLin.swap₁ B₃.val T.val T.val, cubeTriLin.swap₂ T.val B₃.val T.val] + rw [cubeTriLin.swap₁ Y₃.val T.val T.val, cubeTriLin.swap₂ T.val Y₃.val T.val] + rw [cubeTriLin.swap₁ B₃.val T.val T.val, cubeTriLin.swap₂ T.val B₃.val T.val] ring lemma cube_proj (T : MSSMACC.Sols) : cubeTriLin (proj T.1.1).val (proj T.1.1).val (proj T.1.1).val = - 3 * dot Y₃.val B₃.val ^ 2 * + 3 * dot Y₃.val B₃.val ^ 2 * ((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin T.val T.val Y₃.val + (dot Y₃.val T.val - 2 * dot B₃.val T.val) * cubeTriLin T.val T.val B₃.val) := by nth_rewrite 3 [proj_val] diff --git a/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/PlaneWithY3B3.lean b/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/PlaneWithY3B3.lean index 9832ed2..2c1073a 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/PlaneWithY3B3.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/PlaneWithY3B3.lean @@ -42,7 +42,7 @@ lemma planeY₃B₃_smul (R : MSSMACC.AnomalyFreePerp) (a b c d : ℚ) : rw [smul_add, smul_add] rw [smul_smul, smul_smul, smul_smul] -lemma planeY₃B₃_eq (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (h : a = a' ∧ b = b' ∧ c = c') : +lemma planeY₃B₃_eq (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (h : a = a' ∧ b = b' ∧ c = c') : (planeY₃B₃ R a b c) = (planeY₃B₃ R a' b' c') := by rw [h.1, h.2.1, h.2.2] @@ -82,7 +82,7 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (hR' : R have h2 := congrArg (fun S => S i) h1i change _ = 0 at h2 simp [HSMul.hSMul] at h2 - have hc : c + -c' = 0 := by + have hc : c + -c' = 0 := by cases h2 <;> rename_i h2 exact h2 exact (hi h2).elim @@ -105,8 +105,8 @@ lemma planeY₃B₃_quad (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) : lemma planeY₃B₃_cubic (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) : accCube (planeY₃B₃ R a b c).val = c ^ 2 * - (3 * a * cubeTriLin R.val R.val Y₃.val - + 3 * b * cubeTriLin R.val R.val B₃.val + c * cubeTriLin R.val R.val R.val) := by + (3 * a * cubeTriLin R.val R.val Y₃.val + + 3 * b * cubeTriLin R.val R.val B₃.val + c * cubeTriLin R.val R.val R.val) := by rw [planeY₃B₃_val] erw [TriLinearSymm.toCubic_add] erw [lineY₃B₃Charges_cubic] @@ -178,7 +178,7 @@ def lineCube (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ℚ) : MSSMACC.LinSols := planeY₃B₃ R (a₂ * cubeTriLin R.val R.val R.val - 3 * a₃ * cubeTriLin R.val R.val B₃.val) - (3 * a₃ * cubeTriLin R.val R.val Y₃.val - a₁ * cubeTriLin R.val R.val R.val) + (3 * a₃ * cubeTriLin R.val R.val Y₃.val - a₁ * cubeTriLin R.val R.val R.val) (3 * (a₁ * cubeTriLin R.val R.val B₃.val - a₂ * cubeTriLin R.val R.val Y₃.val)) lemma lineCube_smul (R : MSSMACC.AnomalyFreePerp) (a b c d : ℚ) : @@ -205,7 +205,7 @@ lemma lineCube_quad (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ℚ) : section proj -lemma α₃_proj (T : MSSMACC.Sols) : α₃ (proj T.1.1) = +lemma α₃_proj (T : MSSMACC.Sols) : α₃ (proj T.1.1) = 6 * dot Y₃.val B₃.val ^ 3 * ( cubeTriLin T.val T.val Y₃.val * quadBiLin B₃.val T.val - cubeTriLin T.val T.val B₃.val * quadBiLin Y₃.val T.val) := by @@ -214,13 +214,13 @@ lemma α₃_proj (T : MSSMACC.Sols) : α₃ (proj T.1.1) = ring lemma α₂_proj (T : MSSMACC.Sols) : α₂ (proj T.1.1) = - - α₃ (proj T.1.1) * (dot Y₃.val T.val - 2 * dot B₃.val T.val) := by + - α₃ (proj T.1.1) * (dot Y₃.val T.val - 2 * dot B₃.val T.val) := by rw [α₃_proj, α₂] rw [cube_proj_proj_Y₃, quad_Y₃_proj, quad_proj, cube_proj] ring lemma α₁_proj (T : MSSMACC.Sols) : α₁ (proj T.1.1) = - - α₃ (proj T.1.1) * (dot B₃.val T.val - dot Y₃.val T.val) := by + - α₃ (proj T.1.1) * (dot B₃.val T.val - dot Y₃.val T.val) := by rw [α₃_proj, α₁] rw [cube_proj_proj_B₃, quad_B₃_proj, quad_proj, cube_proj] ring diff --git a/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/ToSols.lean b/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/ToSols.lean index 659f847..78917a7 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/ToSols.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/OrthogY3B3/ToSols.lean @@ -146,7 +146,7 @@ def InCubeSolProp (R : MSSMACC.Sols) : Prop := /-- A rational which has two properties. It is zero for a solution `T` if and only if that solution satisfies `inCubeSolProp`. It appears in the definition of `inLineEqProj`. -/ def cubicCoeff (T : MSSMACC.Sols) : ℚ := - 3 * (dot Y₃.val B₃.val) ^ 3 * (cubeTriLin T.val T.val Y₃.val ^ 2 + + 3 * (dot Y₃.val B₃.val) ^ 3 * (cubeTriLin T.val T.val Y₃.val ^ 2 + cubeTriLin T.val T.val B₃.val ^ 2 ) lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) : @@ -243,7 +243,7 @@ def toSolNSProj (T : MSSMACC.Sols) : MSSMACC.AnomalyFreePerp × ℚ × ℚ × lemma toSolNS_proj (T : NotInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val := by apply ACCSystem.Sols.ext rw [toSolNS, toSolNSProj] - change (lineEqCoeff T.val)⁻¹ • (toSolNSQuad _).1.1 = _ + change (lineEqCoeff T.val)⁻¹ • (toSolNSQuad _).1.1 = _ rw [toSolNSQuad_eq_planeY₃B₃_on_α] rw [planeY₃B₃_val] rw [Y₃_plus_B₃_plus_proj] @@ -254,7 +254,7 @@ lemma toSolNS_proj (T : NotInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val := rw [lineEqCoeff] ring rw [h1] - have h1 := (lineEqPropSol_iff_lineEqCoeff_zero T.val).mpr.mt T.prop + have h1 := (lineEqPropSol_iff_lineEqCoeff_zero T.val).mpr.mt T.prop rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h1] simp @@ -302,7 +302,7 @@ lemma inLineEqToSol_proj (T : InLineEqSol) : inLineEqToSol (inLineEqProj T) = T. simp /-- Given a element of `inQuad × ℚ × ℚ × ℚ`, a solution to the ACCs. -/ -def inQuadToSol : InQuad × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a₁, a₂, a₃) => +def inQuadToSol : InQuad × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a₁, a₂, a₃) => AnomalyFreeMk' (lineCube R.val.val a₁ a₂ a₃) (by erw [planeY₃B₃_quad] @@ -391,8 +391,8 @@ lemma inQuadCubeToSol_proj (T : InQuadCubeSol) : /-- Given an element of `MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ` a solution. We will show that this map is a surjection. -/ -def toSol : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a, b, c) => - if h₃ : LineEqProp R ∧ InQuadProp R ∧ InCubeProp R then +def toSol : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a, b, c) => + if h₃ : LineEqProp R ∧ InQuadProp R ∧ InCubeProp R then inQuadCubeToSol (⟨⟨⟨R, h₃.1⟩, h₃.2.1⟩, h₃.2.2⟩, a, b, c) else if h₂ : LineEqProp R ∧ InQuadProp R then diff --git a/HepLean/AnomalyCancellation/MSSMNu/Permutations.lean b/HepLean/AnomalyCancellation/MSSMNu/Permutations.lean index 1935340..b5da680 100644 --- a/HepLean/AnomalyCancellation/MSSMNu/Permutations.lean +++ b/HepLean/AnomalyCancellation/MSSMNu/Permutations.lean @@ -17,7 +17,7 @@ and its action on the MSSM. universe v u open Nat -open Finset +open Finset namespace MSSM @@ -30,7 +30,7 @@ open BigOperators def PermGroup := Fin 6 → Equiv.Perm (Fin 3) @[simp] -instance : Group PermGroup := Pi.group +instance : Group PermGroup := Pi.group /-- The image of an element of `permGroup` under the representation on charges. -/ @[simps!] @@ -65,7 +65,7 @@ def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where apply LinearMap.ext intro S rw [charges_eq_toSpecies_eq] - refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl + refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl intro i simp only [ Pi.mul_apply, Pi.inv_apply, Equiv.Perm.coe_mul, LinearMap.mul_apply] rw [chargeMap_toSpecies, chargeMap_toSpecies] @@ -76,7 +76,7 @@ def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where apply LinearMap.ext intro S rw [charges_eq_toSpecies_eq] - refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl + refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl intro i erw [toSMSpecies_toSpecies_inv] rfl @@ -91,46 +91,46 @@ lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup) (S : MSSMCharges.Charges rw [repCharges_toSMSpecies] exact Equiv.sum_comp (f⁻¹ j) ((fun a => a ^ m) ∘ (toSMSpecies j) S) -lemma Hd_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma Hd_invariant (f : PermGroup) (S : MSSMCharges.Charges) : Hd (repCharges f S) = Hd S := rfl -lemma Hu_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma Hu_invariant (f : PermGroup) (S : MSSMCharges.Charges) : Hu (repCharges f S) = Hu S := rfl -lemma accGrav_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma accGrav_invariant (f : PermGroup) (S : MSSMCharges.Charges) : accGrav (repCharges f S) = accGrav S := accGrav_ext (by simpa using toSpecies_sum_invariant 1 f S) (Hd_invariant f S) (Hu_invariant f S) -lemma accSU2_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma accSU2_invariant (f : PermGroup) (S : MSSMCharges.Charges) : accSU2 (repCharges f S) = accSU2 S := accSU2_ext (by simpa using toSpecies_sum_invariant 1 f S) (Hd_invariant f S) (Hu_invariant f S) -lemma accSU3_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma accSU3_invariant (f : PermGroup) (S : MSSMCharges.Charges) : accSU3 (repCharges f S) = accSU3 S := accSU3_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accYY_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma accYY_invariant (f : PermGroup) (S : MSSMCharges.Charges) : accYY (repCharges f S) = accYY S := accYY_ext (by simpa using toSpecies_sum_invariant 1 f S) (Hd_invariant f S) (Hu_invariant f S) -lemma accQuad_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma accQuad_invariant (f : PermGroup) (S : MSSMCharges.Charges) : accQuad (repCharges f S) = accQuad S := accQuad_ext (toSpecies_sum_invariant 2 f S) (Hd_invariant f S) (Hu_invariant f S) -lemma accCube_invariant (f : PermGroup) (S : MSSMCharges.Charges) : +lemma accCube_invariant (f : PermGroup) (S : MSSMCharges.Charges) : accCube (repCharges f S) = accCube S := accCube_ext (fun j => toSpecies_sum_invariant 3 f S j) diff --git a/HepLean/AnomalyCancellation/PureU1/Basic.lean b/HepLean/AnomalyCancellation/PureU1/Basic.lean index b5ab703..096b2e8 100644 --- a/HepLean/AnomalyCancellation/PureU1/Basic.lean +++ b/HepLean/AnomalyCancellation/PureU1/Basic.lean @@ -16,7 +16,7 @@ We define the anomaly cancellation conditions for a pure U(1) gauge theory with universe v u open Nat -open Finset +open Finset namespace PureU1 open BigOperators @@ -37,7 +37,7 @@ def accGrav (n : ℕ) : ((PureU1Charges n).Charges →ₗ[ℚ] ℚ) where /-- The symmetric trilinear form used to define the cubic anomaly. -/ @[simps!] def accCubeTriLinSymm {n : ℕ} : TriLinearSymm (PureU1Charges n).Charges := TriLinearSymm.mk₃ - (fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i) + (fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i) (by intro a S L T simp [HSMul.hSMul] @@ -70,13 +70,13 @@ def accCubeTriLinSymm {n : ℕ} : TriLinearSymm (PureU1Charges n).Charges := Tri /-- The cubic anomaly equation. -/ @[simp] -def accCube (n : ℕ) : HomogeneousCubic ((PureU1Charges n).Charges) := +def accCube (n : ℕ) : HomogeneousCubic ((PureU1Charges n).Charges) := (accCubeTriLinSymm).toCubic lemma accCube_explicit (n : ℕ) (S : (PureU1Charges n).Charges) : accCube n S = ∑ i : Fin n, S i ^ 3:= by rw [accCube, TriLinearSymm.toCubic] - change accCubeTriLinSymm S S S = _ + change accCubeTriLinSymm S S S = _ rw [accCubeTriLinSymm] simp only [PureU1Charges_numberCharges, TriLinearSymm.mk₃_toFun_apply_apply] apply Finset.sum_congr @@ -99,7 +99,7 @@ def PureU1 (n : ℕ) : ACCSystem where /-- An equivalence of vector spaces of charges when the number of fermions is equal. -/ def pureU1EqCharges {n m : ℕ} (h : n = m): - (PureU1 n).Charges ≃ₗ[ℚ] (PureU1 m).Charges where + (PureU1 n).Charges ≃ₗ[ℚ] (PureU1 m).Charges where toFun f := f ∘ Fin.cast h.symm invFun f := f ∘ Fin.cast h map_add' _ _ := rfl diff --git a/HepLean/AnomalyCancellation/PureU1/BasisLinear.lean b/HepLean/AnomalyCancellation/PureU1/BasisLinear.lean index 746e2d6..fcf53fd 100644 --- a/HepLean/AnomalyCancellation/PureU1/BasisLinear.lean +++ b/HepLean/AnomalyCancellation/PureU1/BasisLinear.lean @@ -84,7 +84,7 @@ noncomputable def coordinateMap : ((PureU1 n.succ).LinSols) ≃ₗ[ℚ] Fin n →₀ ℚ where toFun S := Finsupp.equivFunOnFinite.invFun (S.1 ∘ Fin.castSucc) map_add' S T := Finsupp.ext (congrFun rfl) - map_smul' a S := Finsupp.ext (congrFun rfl) + map_smul' a S := Finsupp.ext (congrFun rfl) invFun f := ∑ i : Fin n, f i • asLinSols i left_inv S := by simp only [PureU1_numberCharges, Equiv.invFun_as_coe, Finsupp.equivFunOnFinite_symm_apply_toFun, @@ -92,7 +92,7 @@ def coordinateMap : ((PureU1 n.succ).LinSols) ≃ₗ[ℚ] Fin n →₀ ℚ where apply pureU1_anomalyFree_ext intro j rw [sum_of_vectors] - simp only [HSMul.hSMul, SMul.smul, PureU1_numberCharges, + simp only [HSMul.hSMul, SMul.smul, PureU1_numberCharges, asLinSols_val, Equiv.toFun_as_coe, Fin.coe_eq_castSucc, mul_ite, mul_one, mul_neg, mul_zero, Equiv.invFun_as_coe] rw [Finset.sum_eq_single j] @@ -127,7 +127,7 @@ instance : Module.Finite ℚ ((PureU1 n.succ).LinSols) := lemma finrank_AnomalyFreeLinear : FiniteDimensional.finrank ℚ (((PureU1 n.succ).LinSols)) = n := by - have h := Module.mk_finrank_eq_card_basis (@asBasis n) + have h := Module.mk_finrank_eq_card_basis (@asBasis n) simp only [Nat.succ_eq_add_one, finrank_eq_rank, Cardinal.mk_fintype, Fintype.card_fin] at h exact FiniteDimensional.finrank_eq_of_rank_eq h diff --git a/HepLean/AnomalyCancellation/PureU1/ConstAbs.lean b/HepLean/AnomalyCancellation/PureU1/ConstAbs.lean index 3189e4c..1d3ea4a 100644 --- a/HepLean/AnomalyCancellation/PureU1/ConstAbs.lean +++ b/HepLean/AnomalyCancellation/PureU1/ConstAbs.lean @@ -14,7 +14,7 @@ We look at charge assignments in which all charges have the same absolute value. universe v u open Nat -open Finset +open Finset open BigOperators namespace PureU1 @@ -53,7 +53,7 @@ section charges variable {S : (PureU1 n.succ).Charges} {A : (PureU1 n.succ).LinSols} variable (hS : ConstAbsSorted S) (hA : ConstAbsSorted A.val) -lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := by +lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := by have hSS := hS.2 i k hik have ht := hS.1 i k rw [sq_eq_sq_iff_eq_or_eq_neg] at ht @@ -61,7 +61,7 @@ lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := exact h linarith -lemma val_le_zero {i : Fin n.succ} (hi : S i ≤ 0) : S i = S (0 : Fin n.succ) := by +lemma val_le_zero {i : Fin n.succ} (hi : S i ≤ 0) : S i = S (0 : Fin n.succ) := by symm apply lt_eq hS hi simp @@ -108,7 +108,7 @@ lemma boundary_succ {k : Fin n} (hk : Boundary S k) : S k.succ = - S (0 : Fin n. rw [opposite_signs_eq_neg hS (le_of_lt hk.left) (le_of_lt hk.right)] at hn linear_combination -(1 * hn) -lemma boundary_split (k : Fin n) : k.succ.val + (n.succ - k.succ.val) = n.succ := by +lemma boundary_split (k : Fin n) : k.succ.val + (n.succ - k.succ.val) = n.succ := by omega lemma boundary_accGrav' (k : Fin n) : accGrav n.succ S = @@ -128,7 +128,7 @@ lemma boundary_accGrav'' (k : Fin n) (hk : Boundary S k) : S (Fin.cast (boundary_split k) (Fin.castAdd (n.succ - k.succ.val) i)) = S k.castSucc := by apply lt_eq hS (le_of_lt hk.left) (by rw [Fin.le_def]; simp; omega) have hsnd (i : Fin (n.succ - k.succ.val)) : - S (Fin.cast (boundary_split k) (Fin.natAdd (k.succ.val) i)) = S k.succ := by + S (Fin.cast (boundary_split k) (Fin.natAdd (k.succ.val) i)) = S k.succ := by apply gt_eq hS (le_of_lt hk.right) (by rw [Fin.le_def]; simp) simp only [hfst, hsnd] simp only [Fin.val_succ, sum_const, card_fin, nsmul_eq_mul, cast_add, cast_one, @@ -160,7 +160,7 @@ lemma not_hasBoundry_zero (hnot : ¬ (HasBoundary S)) (i : Fin n.succ) : simp at hi exact zero_gt hS hi i -lemma not_hasBoundary_grav (hnot : ¬ (HasBoundary S)) : +lemma not_hasBoundary_grav (hnot : ¬ (HasBoundary S)) : accGrav n.succ S = n.succ * S (0 : Fin n.succ) := by simp [accGrav, ← not_hasBoundry_zero hS hnot] @@ -196,7 +196,7 @@ lemma AFL_odd_zero {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted A.val) theorem AFL_odd (A : (PureU1 (2 * n + 1)).LinSols) (h : ConstAbsSorted A.val) : A = 0 := by apply ACCSystemLinear.LinSols.ext - exact is_zero h (AFL_odd_zero h) + exact is_zero h (AFL_odd_zero h) lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val) (hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) {k : Fin (2 * n + 1)} (hk : Boundary A.val k) : @@ -209,8 +209,8 @@ lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted linear_combination h0 / 2 lemma AFL_even_below' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val) - (hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) (i : Fin n.succ) : - A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) = A.val (0 : Fin (2*n.succ)) := by + (hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) (i : Fin n.succ) : + A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) = A.val (0 : Fin (2*n.succ)) := by obtain ⟨k, hk⟩ := AFL_hasBoundary h hA rw [← boundary_castSucc h hk] apply lt_eq h (le_of_lt hk.left) @@ -221,7 +221,7 @@ lemma AFL_even_below' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A. lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.val) (i : Fin n.succ) : - A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) + A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) = A.val (0 : Fin (2*n.succ)) := by by_cases hA : A.val (0 : Fin (2*n.succ)) = 0 rw [is_zero h hA] @@ -230,8 +230,8 @@ lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.v exact AFL_even_below' h hA i lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val) - (hA : A.val (0 : Fin (2*n.succ)) ≠ 0) (i : Fin n.succ) : - A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) = + (hA : A.val (0 : Fin (2*n.succ)) ≠ 0) (i : Fin n.succ) : + A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) = - A.val (0 : Fin (2*n.succ)) := by obtain ⟨k, hk⟩ := AFL_hasBoundary h hA rw [← boundary_succ h hk] @@ -243,7 +243,7 @@ lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A. lemma AFL_even_above (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.val) (i : Fin n.succ) : - A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) = + A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) = - A.val (0 : Fin (2*n.succ)) := by by_cases hA : A.val (0 : Fin (2*n.succ)) = 0 rw [is_zero h hA] diff --git a/HepLean/AnomalyCancellation/PureU1/Even/BasisLinear.lean b/HepLean/AnomalyCancellation/PureU1/Even/BasisLinear.lean index cdbf343..83dd9df 100644 --- a/HepLean/AnomalyCancellation/PureU1/Even/BasisLinear.lean +++ b/HepLean/AnomalyCancellation/PureU1/Even/BasisLinear.lean @@ -16,7 +16,7 @@ that splits into two planes on which every point is a solution to the ACCs. universe v u open Nat -open Finset +open Finset open BigOperators namespace PureU1 @@ -51,7 +51,7 @@ def δ!₃ : Fin (2 * n.succ) := (Fin.cast (n_cond₂ n) (Fin.castAdd ((n + n) + def δ!₄ : Fin (2 * n.succ) := (Fin.cast (n_cond₂ n) (Fin.natAdd 1 (Fin.natAdd (n + n) 0))) lemma ext_δ (S T : Fin (2 * n.succ) → ℚ) (h1 : ∀ i, S (δ₁ i) = T (δ₁ i)) - (h2 : ∀ i, S (δ₂ i) = T (δ₂ i)) : S = T := by + (h2 : ∀ i, S (δ₂ i) = T (δ₂ i)) : S = T := by funext i by_cases hi : i.val < n.succ let j : Fin n.succ := ⟨i, hi⟩ @@ -68,7 +68,7 @@ lemma ext_δ (S T : Fin (2 * n.succ) → ℚ) (h1 : ∀ i, S (δ₁ i) = T (δ rw [h3] at h2 exact h2 -lemma sum_δ₁_δ₂ (S : Fin (2 * n.succ) → ℚ) : +lemma sum_δ₁_δ₂ (S : Fin (2 * n.succ) → ℚ) : ∑ i, S i = ∑ i : Fin n.succ, ((S ∘ δ₁) i + (S ∘ δ₂) i) := by have h1 : ∑ i, S i = ∑ i : Fin (n.succ + n.succ), S (Fin.cast (split_equal n.succ) i) := by rw [Finset.sum_equiv (Fin.castOrderIso (split_equal n.succ)).symm.toEquiv] @@ -80,7 +80,7 @@ lemma sum_δ₁_δ₂ (S : Fin (2 * n.succ) → ℚ) : rw [Fin.sum_univ_add, Finset.sum_add_distrib] rfl -lemma sum_δ₁_δ₂' (S : Fin (2 * n.succ) → ℚ) : +lemma sum_δ₁_δ₂' (S : Fin (2 * n.succ) → ℚ) : ∑ i, S i = ∑ i : Fin n.succ, ((S ∘ δ₁) i + (S ∘ δ₂) i) := by have h1 : ∑ i, S i = ∑ i : Fin (n.succ + n.succ), S (Fin.cast (split_equal n.succ) i) := by rw [Finset.sum_equiv (Fin.castOrderIso (split_equal n.succ)).symm.toEquiv] @@ -92,8 +92,8 @@ lemma sum_δ₁_δ₂' (S : Fin (2 * n.succ) → ℚ) : rw [Fin.sum_univ_add, Finset.sum_add_distrib] rfl -lemma sum_δ!₁_δ!₂ (S : Fin (2 * n.succ) → ℚ) : - ∑ i, S i = S δ!₃ + S δ!₄ + ∑ i : Fin n, ((S ∘ δ!₁) i + (S ∘ δ!₂) i) := by +lemma sum_δ!₁_δ!₂ (S : Fin (2 * n.succ) → ℚ) : + ∑ i, S i = S δ!₃ + S δ!₄ + ∑ i : Fin n, ((S ∘ δ!₁) i + (S ∘ δ!₂) i) := by have h1 : ∑ i, S i = ∑ i : Fin (1 + ((n + n) + 1)), S (Fin.cast (n_cond₂ n) i) := by rw [Finset.sum_equiv (Fin.castOrderIso (n_cond₂ n)).symm.toEquiv] intro i @@ -180,12 +180,12 @@ lemma basis_on_δ₁_other {k j : Fin n.succ} (h : k ≠ j) : omega rfl -lemma basis_on_other {k : Fin n.succ} {j : Fin (2 * n.succ)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) : +lemma basis_on_other {k : Fin n.succ} {j : Fin (2 * n.succ)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) : basisAsCharges k j = 0 := by simp [basisAsCharges] simp_all only [ne_eq, ↓reduceIte] -lemma basis!_on_other {k : Fin n} {j : Fin (2 * n.succ)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) : +lemma basis!_on_other {k : Fin n} {j : Fin (2 * n.succ)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) : basis!AsCharges k j = 0 := by simp [basis!AsCharges] simp_all only [ne_eq, ↓reduceIte] @@ -338,11 +338,11 @@ def basisa : (Fin n.succ) ⊕ (Fin n) → (PureU1 (2 * n.succ)).LinSols := fun i /-- Swapping the elements δ!₁ j and δ!₂ j is equivalent to adding a vector basis!AsCharges j. -/ lemma swap!_as_add {S S' : (PureU1 (2 * n.succ)).LinSols} (j : Fin n) (hS : ((FamilyPermutations (2 * n.succ)).linSolRep - (pairSwap (δ!₁ j) (δ!₂ j))) S = S') : + (pairSwap (δ!₁ j) (δ!₂ j))) S = S') : S'.val = S.val + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j := by funext i rw [← hS, FamilyPermutations_anomalyFreeLinear_apply] - by_cases hi : i = δ!₁ j + by_cases hi : i = δ!₁ j subst hi simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst] by_cases hi2 : i = δ!₂ j @@ -350,7 +350,7 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n.succ)).LinSols} (j : Fin n) simp [HSMul.hSMul, basis!_on_δ!₂_self, pairSwap_inv_snd] simp [HSMul.hSMul] rw [basis!_on_other hi hi2] - change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_ + change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_ erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)] simp @@ -374,7 +374,7 @@ lemma P_δ₁ (f : Fin n.succ → ℚ) (j : Fin n.succ) : P f (δ₁ j) = f j := simp only [mul_zero] simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff] -lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by +lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by rw [P!, sum_of_charges] simp [HSMul.hSMul, SMul.smul] rw [Finset.sum_eq_single j] @@ -424,7 +424,7 @@ lemma P!_δ!₃ (f : Fin n → ℚ) : P! f (δ!₃) = 0 := by rw [P!, sum_of_charges] simp [HSMul.hSMul, SMul.smul, basis!_on_δ!₃] -lemma Pa_δ!₃ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₃) = f 0 := by +lemma Pa_δ!₃ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₃) = f 0 := by rw [Pa] simp only [ACCSystemCharges.chargesAddCommMonoid_add] rw [P!_δ!₃, δ!₃_δ₁0, P_δ₁] @@ -434,13 +434,13 @@ lemma P!_δ!₄ (f : Fin n → ℚ) : P! f (δ!₄) = 0 := by rw [P!, sum_of_charges] simp [HSMul.hSMul, SMul.smul, basis!_on_δ!₄] -lemma Pa_δ!₄ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₄) = - f (Fin.last n) := by +lemma Pa_δ!₄ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₄) = - f (Fin.last n) := by rw [Pa] simp only [ACCSystemCharges.chargesAddCommMonoid_add] rw [P!_δ!₄, δ!₄_δ₂Last, P_δ₂] simp -lemma P_δ₁_δ₂ (f : Fin n.succ → ℚ) : P f ∘ δ₂ = - P f ∘ δ₁ := by +lemma P_δ₁_δ₂ (f : Fin n.succ → ℚ) : P f ∘ δ₂ = - P f ∘ δ₁ := by funext j simp only [PureU1_numberCharges, Function.comp_apply, Pi.neg_apply] rw [P_δ₁, P_δ₂] @@ -484,7 +484,7 @@ lemma P_P_P!_accCube (g : Fin n.succ → ℚ) (j : Fin n) : lemma P_P!_P!_accCube (g : Fin n → ℚ) (j : Fin n.succ) : accCubeTriLinSymm (P! g) (P! g) (basisAsCharges j) - = (P! g (δ₁ j))^2 - (P! g (δ₂ j))^2 := by + = (P! g (δ₁ j))^2 - (P! g (δ₂ j))^2 := by simp [accCubeTriLinSymm] rw [sum_δ₁_δ₂] simp only [Function.comp_apply] @@ -604,7 +604,7 @@ theorem basisa_linear_independent : LinearIndependent ℚ (@basisa n) := by simp_all simp_all -lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by +lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by apply Iff.intro intro h funext i @@ -625,7 +625,7 @@ lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ℚ) : Pa' f = Pa' f' ↔ f = /-! TODO: Replace the definition of `join` with a Mathlib definition, most likely `Sum.elim`. -/ /-- A helper function for what follows. -/ -def join (g : Fin n.succ → ℚ) (f : Fin n → ℚ) : (Fin n.succ) ⊕ (Fin n) → ℚ := fun i => +def join (g : Fin n.succ → ℚ) (f : Fin n → ℚ) : (Fin n.succ) ⊕ (Fin n) → ℚ := fun i => match i with | .inl i => g i | .inr i => f i @@ -661,7 +661,7 @@ lemma Pa_eq (g g' : Fin n.succ → ℚ) (f f' : Fin n → ℚ) : rw [← join_ext] exact Pa'_eq _ _ -lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) = +lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) = FiniteDimensional.finrank ℚ (PureU1 (2 * n.succ)).LinSols := by erw [BasisLinear.finrank_AnomalyFreeLinear] simp only [Fintype.card_sum, Fintype.card_fin, mul_eq] @@ -673,7 +673,7 @@ noncomputable def basisaAsBasis : basisOfLinearIndependentOfCardEqFinrank (@basisa_linear_independent n) basisa_card lemma span_basis (S : (PureU1 (2 * n.succ)).LinSols) : - ∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f := by + ∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f := by have h := (mem_span_range_iff_exists_fun ℚ).mp (Basis.mem_span basisaAsBasis S) obtain ⟨f, hf⟩ := h simp [basisaAsBasis] at hf @@ -706,7 +706,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n) S'.val = P g' + P! f' ∧ P! f' = P! f + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∧ g' = g := by let X := P! f + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j - have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by + have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by apply Submodule.add_mem exact (P!_in_span f) exact (smul_basis!AsCharges_in_span S j) diff --git a/HepLean/AnomalyCancellation/PureU1/Even/LineInCubic.lean b/HepLean/AnomalyCancellation/PureU1/Even/LineInCubic.lean index 88e5991..1cee03f 100644 --- a/HepLean/AnomalyCancellation/PureU1/Even/LineInCubic.lean +++ b/HepLean/AnomalyCancellation/PureU1/Even/LineInCubic.lean @@ -34,7 +34,7 @@ open BigOperators variable {n : ℕ} open VectorLikeEvenPlane -/-- A property on `LinSols`, satisfied if every point on the line between the two planes +/-- A property on `LinSols`, satisfied if every point on the line between the two planes in the basis through that point is in the cubic. -/ def LineInCubic (S : (PureU1 (2 * n.succ)).LinSols) : Prop := ∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) , @@ -61,18 +61,18 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S) then `accCubeTriLinSymm.toFun (P g, P g, P! f) = 0`. -/ lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S) : - ∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f), + ∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f), accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by intro g f hS linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1) - (lineInCubic_expand h g f hS 1 2) / 6 -/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/ +/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/ def LineInCubicPerm (S : (PureU1 (2 * n.succ)).LinSols) : Prop := ∀ (M : (FamilyPermutations (2 * n.succ)).group ), LineInCubic ((FamilyPermutations (2 * n.succ)).linSolRep M S) -/-- If `lineInCubicPerm S` then `lineInCubic S`. -/ +/-- If `lineInCubicPerm S` then `lineInCubic S`. -/ lemma lineInCubicPerm_self {S : (PureU1 (2 * n.succ)).LinSols} (hS : LineInCubicPerm S) : LineInCubic S := hS 1 @@ -94,7 +94,7 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols} (S.val (δ!₂ j) - S.val (δ!₁ j)) * accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by intro j g f h - let S' := (FamilyPermutations (2 * n.succ)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j)) S + let S' := (FamilyPermutations (2 * n.succ)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j)) S have hSS' : ((FamilyPermutations (2 * n.succ)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j))) S = S' := rfl obtain ⟨g', f', hall⟩ := span_basis_swap! j hSS' g f h have h1 := line_in_cubic_P_P_P! (lineInCubicPerm_self LIC) g f h @@ -106,21 +106,21 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols} lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ )).LinSols} (f : Fin n.succ.succ → ℚ) (g : Fin n.succ → ℚ) (hS : S.val = Pa f g) : - accCubeTriLinSymm (P f) (P f) (basis!AsCharges (Fin.last n)) = + accCubeTriLinSymm (P f) (P f) (basis!AsCharges (Fin.last n)) = - (S.val (δ!₂ (Fin.last n)) + S.val (δ!₁ (Fin.last n))) * (2 * S.val δ!₄ + - S.val (δ!₂ (Fin.last n)) + S.val (δ!₁ (Fin.last n))) := by - rw [P_P_P!_accCube f (Fin.last n)] + S.val (δ!₂ (Fin.last n)) + S.val (δ!₁ (Fin.last n))) := by + rw [P_P_P!_accCube f (Fin.last n)] have h1 := Pa_δ!₄ f g have h2 := Pa_δ!₁ f g (Fin.last n) have h3 := Pa_δ!₂ f g (Fin.last n) simp at h1 h2 h3 - have hl : f (Fin.succ (Fin.last (n ))) = - Pa f g δ!₄ := by + have hl : f (Fin.succ (Fin.last (n ))) = - Pa f g δ!₄ := by simp_all only [Fin.succ_last, neg_neg] erw [hl] at h2 - have hg : g (Fin.last n) = Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄ := by + have hg : g (Fin.last n) = Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄ := by linear_combination -(1 * h2) - have hll : f (Fin.castSucc (Fin.last (n ))) = - - (Pa f g (δ!₂ (Fin.last n)) + Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄) := by + have hll : f (Fin.castSucc (Fin.last (n ))) = + - (Pa f g (δ!₂ (Fin.last n)) + Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄) := by linear_combination h3 - 1 * hg rw [← hS] at hl hll rw [hl, hll] @@ -145,7 +145,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols} apply Or.inr exact h1 -lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols} +lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols} (LIC : LineInCubicPerm S) : LineInPlaneCond S := by refine @Prop_three (2 * n.succ.succ) LineInPlaneProp S (δ!₂ (Fin.last n)) (δ!₁ (Fin.last n)) δ!₄ ?_ ?_ ?_ ?_ @@ -156,15 +156,15 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols} intro M exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M) -lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ)).Sols} +lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ)).Sols} (LIC : LineInCubicPerm S.1.1) : ConstAbs S.val := linesInPlane_constAbs_AF S (lineInCubicPerm_last_perm LIC) -theorem lineInCubicPerm_vectorLike {S : (PureU1 (2 * n.succ.succ)).Sols} +theorem lineInCubicPerm_vectorLike {S : (PureU1 (2 * n.succ.succ)).Sols} (LIC : LineInCubicPerm S.1.1) : VectorLikeEven S.val := ConstAbs.boundary_value_even S.1.1 (lineInCubicPerm_constAbs LIC) -theorem lineInCubicPerm_in_plane (S : (PureU1 (2 * n.succ.succ)).Sols) +theorem lineInCubicPerm_in_plane (S : (PureU1 (2 * n.succ.succ)).Sols) (LIC : LineInCubicPerm S.1.1) : ∃ (M : (FamilyPermutations (2 * n.succ.succ)).group), (FamilyPermutations (2 * n.succ.succ)).linSolRep M S.1.1 ∈ Submodule.span ℚ (Set.range basis) := diff --git a/HepLean/AnomalyCancellation/PureU1/Even/Parameterization.lean b/HepLean/AnomalyCancellation/PureU1/Even/Parameterization.lean index db11575..f6dd281 100644 --- a/HepLean/AnomalyCancellation/PureU1/Even/Parameterization.lean +++ b/HepLean/AnomalyCancellation/PureU1/Even/Parameterization.lean @@ -14,7 +14,7 @@ import Mathlib.Tactic.Polyrith /-! # Parameterization in even case -Given maps `g : Fin n.succ → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly +Given maps `g : Fin n.succ → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly equations. We show that every solution can be got in this way, up to permutation, unless it, up to permutation, lives in the plane spanned by the first part of the basis vector. @@ -60,7 +60,7 @@ lemma parameterizationCharge_cube (g : Fin n.succ → ℚ) (f : Fin n → ℚ) ( accCubeTriLinSymm.map_smul₃] ring -/-- The construction of a `Sol` from a `Fin n.succ → ℚ`, a `Fin n → ℚ` and a `ℚ`. -/ +/-- The construction of a `Sol` from a `Fin n.succ → ℚ`, a `Fin n → ℚ` and a `ℚ`. -/ def parameterization (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (a : ℚ) : (PureU1 (2 * n.succ)).Sols := ⟨⟨parameterizationAsLinear g f a, by intro i; simp at i; exact Fin.elim0 i⟩, @@ -77,15 +77,15 @@ lemma anomalyFree_param {S : (PureU1 (2 * n.succ)).Sols} erw [P!_accCube] at hC linear_combination hC / 3 -/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. +/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. In this case our parameterization above will be able to recover this point. -/ def GenericCase (S : (PureU1 (2 * n.succ)).Sols) : Prop := ∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) , - accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 + accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols) (hs : ∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f ∧ - accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by + accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by intro g f hS hC obtain ⟨g', f', hS', hC'⟩ := hs rw [hS] at hS' @@ -94,13 +94,13 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols) exact hC' hC /-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) = 0`.-/ -def SpecialCase (S : (PureU1 (2 * n.succ)).Sols) : Prop := +def SpecialCase (S : (PureU1 (2 * n.succ)).Sols) : Prop := ∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) , accCubeTriLinSymm (P g) (P g) (P! f) = 0 lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols) (hs : ∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f ∧ - accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by + accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by intro g f hS obtain ⟨g', f', hS', hC'⟩ := hs rw [hS] at hS' @@ -111,7 +111,7 @@ lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols) lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) : GenericCase S ∨ SpecialCase S := by obtain ⟨g, f, h⟩ := span_basis S.1.1 - have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨ + have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨ accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by exact ne_or_eq _ _ cases h1 <;> rename_i h1 @@ -119,7 +119,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) : exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩) theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : GenericCase S) : - ∃ g f a, S = parameterization g f a := by + ∃ g f a, S = parameterization g f a := by obtain ⟨g, f, hS⟩ := span_basis S.1.1 use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹ rw [parameterization] diff --git a/HepLean/AnomalyCancellation/PureU1/LineInPlaneCond.lean b/HepLean/AnomalyCancellation/PureU1/LineInPlaneCond.lean index 4cfb260..cccdc3f 100644 --- a/HepLean/AnomalyCancellation/PureU1/LineInPlaneCond.lean +++ b/HepLean/AnomalyCancellation/PureU1/LineInPlaneCond.lean @@ -14,7 +14,7 @@ import Mathlib.RepresentationTheory.Basic We say a `LinSol` satisfies the `line in plane` condition if for all distinct `i1`, `i2`, `i3` in `Fin n`, we have -`S i1 = S i2` or `S i1 = - S i2` or `2 S i3 + S i1 + S i2 = 0`. +`S i1 = S i2` or `S i1 = - S i2` or `2 S i3 + S i1 + S i2 = 0`. We look at various consequences of this. The main reference for this material is @@ -60,7 +60,7 @@ lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : LineIn simp only [LineInPlaneProp] at hS simp [not_or] at h have h1 (i : Fin n) : S.val i.castSucc.castSucc = - - (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by + - (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by have h1S := hS (Fin.last n).castSucc ((Fin.last n).succ) i.castSucc.castSucc (by simp; rw [Fin.ext_iff]; simp) (by simp; rw [Fin.ext_iff]; simp; omega) @@ -82,14 +82,14 @@ lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols have h := lineInPlaneCond_eq_last' hS hn rw [sq_eq_sq_iff_eq_or_eq_neg] at hn simp [or_not] at hn - have hx : ((2 : ℚ) - ↑(n + 3)) ≠ 0 := by + have hx : ((2 : ℚ) - ↑(n + 3)) ≠ 0 := by rw [Nat.cast_add] simp only [Nat.cast_ofNat, ne_eq] apply Not.intro intro a linarith - have ht : S.val ((Fin.last n.succ.succ.succ).succ) = - - S.val ((Fin.last n.succ.succ.succ).castSucc) := by + have ht : S.val ((Fin.last n.succ.succ.succ).succ) = + - S.val ((Fin.last n.succ.succ.succ).castSucc) := by rw [← mul_right_inj' hx] simp [Nat.succ_eq_add_one] simp at h @@ -115,7 +115,7 @@ theorem linesInPlane_constAbs {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols rw [hij] lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) : - ConstAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by + ConstAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by simp [ConstAbsProp] by_contra hn have hLin := pureU1_linear S.1.1 @@ -128,7 +128,7 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) : have l013 := hS 0 1 3 (by simp) (by simp) (by simp) have l023 := hS 0 2 3 (by simp) (by simp) (by simp) simp_all [LineInPlaneProp] - have h1 : S.val (2 : Fin 4) = S.val (3 : Fin 4) := by + have h1 : S.val (2 : Fin 4) = S.val (3 : Fin 4) := by linear_combination l012 / 2 + -1 * l013 / 2 by_cases h2 : S.val (0 : Fin 4) = S.val (2 : Fin 4) simp_all @@ -161,7 +161,7 @@ lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols} refine Prop_two ConstAbsProp (by simp : (0 : Fin 4) ≠ 1) ?_ intro M let S' := (FamilyPermutations 4).solAction.toFun S M - have hS' : LineInPlaneCond S'.1.1 := + have hS' : LineInPlaneCond S'.1.1 := (lineInPlaneCond_perm hS M) exact linesInPlane_four S' hS' diff --git a/HepLean/AnomalyCancellation/PureU1/LowDim/One.lean b/HepLean/AnomalyCancellation/PureU1/LowDim/One.lean index fd91578..8493fbe 100644 --- a/HepLean/AnomalyCancellation/PureU1/LowDim/One.lean +++ b/HepLean/AnomalyCancellation/PureU1/LowDim/One.lean @@ -13,7 +13,7 @@ We show that in this case the charge must be zero. universe v u open Nat -open Finset +open Finset namespace PureU1 diff --git a/HepLean/AnomalyCancellation/PureU1/LowDim/Three.lean b/HepLean/AnomalyCancellation/PureU1/LowDim/Three.lean index be0d389..5e1db12 100644 --- a/HepLean/AnomalyCancellation/PureU1/LowDim/Three.lean +++ b/HepLean/AnomalyCancellation/PureU1/LowDim/Three.lean @@ -14,7 +14,7 @@ We define a surjective map from `LinSols` with a charge equal to zero to `Sols`. universe v u open Nat -open Finset +open Finset namespace PureU1 diff --git a/HepLean/AnomalyCancellation/PureU1/LowDim/Two.lean b/HepLean/AnomalyCancellation/PureU1/LowDim/Two.lean index 78cb609..612a0ac 100644 --- a/HepLean/AnomalyCancellation/PureU1/LowDim/Two.lean +++ b/HepLean/AnomalyCancellation/PureU1/LowDim/Two.lean @@ -13,7 +13,7 @@ We define an equivalence between `LinSols` and `Sols`. universe v u open Nat -open Finset +open Finset namespace PureU1 diff --git a/HepLean/AnomalyCancellation/PureU1/Odd/BasisLinear.lean b/HepLean/AnomalyCancellation/PureU1/Odd/BasisLinear.lean index 5af2a8e..9e3bdc9 100644 --- a/HepLean/AnomalyCancellation/PureU1/Odd/BasisLinear.lean +++ b/HepLean/AnomalyCancellation/PureU1/Odd/BasisLinear.lean @@ -16,7 +16,7 @@ that splits into two planes on which every point is a solution to the ACCs. universe v u open Nat -open Finset +open Finset open BigOperators namespace PureU1 @@ -111,7 +111,7 @@ lemma δ₂_δ!₂ (j : Fin n) : δ₂ j = δ!₂ j := by simp [δ₂, δ!₂] omega -lemma sum_δ (S : Fin (2 * n + 1) → ℚ) : +lemma sum_δ (S : Fin (2 * n + 1) → ℚ) : ∑ i, S i = S δ₃ + ∑ i : Fin n, ((S ∘ δ₁) i + (S ∘ δ₂) i) := by have h1 : ∑ i, S i = ∑ i : Fin (n + 1 + n), S (Fin.cast (split_odd n) i) := by rw [Finset.sum_equiv (Fin.castOrderIso (split_odd n)).symm.toEquiv] @@ -127,7 +127,7 @@ lemma sum_δ (S : Fin (2 * n + 1) → ℚ) : rw [Finset.sum_add_distrib] rfl -lemma sum_δ! (S : Fin (2 * n + 1) → ℚ) : +lemma sum_δ! (S : Fin (2 * n + 1) → ℚ) : ∑ i, S i = S δ!₃ + ∑ i : Fin n, ((S ∘ δ!₁) i + (S ∘ δ!₂) i) := by have h1 : ∑ i, S i = ∑ i : Fin ((1+n)+n), S (Fin.cast (split_odd! n) i) := by rw [Finset.sum_equiv (Fin.castOrderIso (split_odd! n)).symm.toEquiv] @@ -210,12 +210,12 @@ lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) : omega rfl -lemma basis_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) : +lemma basis_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) : basisAsCharges k j = 0 := by simp [basisAsCharges] simp_all only [ne_eq, ↓reduceIte] -lemma basis!_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) : +lemma basis!_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) : basis!AsCharges k j = 0 := by simp [basis!AsCharges] simp_all only [ne_eq, ↓reduceIte] @@ -333,11 +333,11 @@ end theBasisVectors /-- Swapping the elements δ!₁ j and δ!₂ j is equivalent to adding a vector basis!AsCharges j. -/ lemma swap!_as_add {S S' : (PureU1 (2 * n + 1)).LinSols} (j : Fin n) (hS : ((FamilyPermutations (2 * n + 1)).linSolRep - (pairSwap (δ!₁ j) (δ!₂ j))) S = S') : + (pairSwap (δ!₁ j) (δ!₂ j))) S = S') : S'.val = S.val + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j := by funext i rw [← hS, FamilyPermutations_anomalyFreeLinear_apply] - by_cases hi : i = δ!₁ j + by_cases hi : i = δ!₁ j subst hi simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst] by_cases hi2 : i = δ!₂ j @@ -345,7 +345,7 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n + 1)).LinSols} (j : Fin n) simp [HSMul.hSMul,basis!_on_δ!₂_self, pairSwap_inv_snd] simp [HSMul.hSMul] rw [basis!_on_other hi hi2] - change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_ + change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_ erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)] simp @@ -369,7 +369,7 @@ lemma P_δ₁ (f : Fin n → ℚ) (j : Fin n) : P f (δ₁ j) = f j := by simp only [mul_zero] simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff] -lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by +lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by rw [P!, sum_of_charges] simp [HSMul.hSMul, SMul.smul] rw [Finset.sum_eq_single j] @@ -418,14 +418,14 @@ lemma Pa_δa₁ (f g : Fin n.succ → ℚ) : Pa f g δa₁ = f 0 := by rw [P_δ₁, P!_δ!₃] simp -lemma Pa_δa₂ (f g : Fin n.succ → ℚ) (j : Fin n) : Pa f g (δa₂ j) = f j.succ + g j.castSucc := by +lemma Pa_δa₂ (f g : Fin n.succ → ℚ) (j : Fin n) : Pa f g (δa₂ j) = f j.succ + g j.castSucc := by rw [Pa] simp only [ACCSystemCharges.chargesAddCommMonoid_add] nth_rewrite 1 [δa₂_δ₁] rw [δa₂_δ!₁] rw [P_δ₁, P!_δ!₁] -lemma Pa_δa₃ (f g : Fin n.succ → ℚ) : Pa f g (δa₃) = g (Fin.last n) := by +lemma Pa_δa₃ (f g : Fin n.succ → ℚ) : Pa f g (δa₃) = g (Fin.last n) := by rw [Pa] simp only [ACCSystemCharges.chargesAddCommMonoid_add] nth_rewrite 1 [δa₃_δ₃] @@ -495,7 +495,7 @@ lemma P!_zero (f : Fin n → ℚ) (h : P! f = 0) : ∀ i, f i = 0 := by rw [h] rfl -lemma Pa_zero (f g : Fin n.succ → ℚ) (h : Pa f g = 0) : +lemma Pa_zero (f g : Fin n.succ → ℚ) (h : Pa f g = 0) : ∀ i, f i = 0 := by have h₃ := Pa_δa₁ f g rw [h] at h₃ @@ -589,7 +589,7 @@ theorem basisa_linear_independent : LinearIndependent ℚ (@basisa n.succ) := by simp_all simp_all -lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by +lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by apply Iff.intro intro h funext i @@ -610,7 +610,7 @@ lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ℚ) : Pa' f = Pa' f' /-! TODO: Replace the definition of `join` with a Mathlib definition, most likely `Sum.elim`. -/ /-- A helper function for what follows. -/ -def join (g f : Fin n → ℚ) : Fin n ⊕ Fin n → ℚ := fun i => +def join (g f : Fin n → ℚ) : Fin n ⊕ Fin n → ℚ := fun i => match i with | .inl i => g i | .inr i => f i @@ -646,7 +646,7 @@ lemma Pa_eq (g g' : Fin n.succ → ℚ) (f f' : Fin n.succ → ℚ) : rw [← join_ext] exact Pa'_eq _ _ -lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n.succ)) = +lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n.succ)) = FiniteDimensional.finrank ℚ (PureU1 (2 * n.succ + 1)).LinSols := by erw [BasisLinear.finrank_AnomalyFreeLinear] simp only [Fintype.card_sum, Fintype.card_fin] @@ -658,7 +658,7 @@ noncomputable def basisaAsBasis : basisOfLinearIndependentOfCardEqFinrank (@basisa_linear_independent n) basisa_card lemma span_basis (S : (PureU1 (2 * n.succ + 1)).LinSols) : - ∃ (g f : Fin n.succ → ℚ) , S.val = P g + P! f := by + ∃ (g f : Fin n.succ → ℚ) , S.val = P g + P! f := by have h := (mem_span_range_iff_exists_fun ℚ).mp (Basis.mem_span basisaAsBasis S) obtain ⟨f, hf⟩ := h simp [basisaAsBasis] at hf @@ -686,7 +686,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ + 1)).LinSols} (j : Fin n.succ) apply SetLike.mem_of_subset apply Submodule.subset_span simp_all only [Set.mem_range, exists_apply_eq_apply] - have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by + have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by apply Submodule.add_mem exact hf exact hP diff --git a/HepLean/AnomalyCancellation/PureU1/Odd/LineInCubic.lean b/HepLean/AnomalyCancellation/PureU1/Odd/LineInCubic.lean index 3cc030f..6d909eb 100644 --- a/HepLean/AnomalyCancellation/PureU1/Odd/LineInCubic.lean +++ b/HepLean/AnomalyCancellation/PureU1/Odd/LineInCubic.lean @@ -34,10 +34,10 @@ open BigOperators variable {n : ℕ} open VectorLikeOddPlane -/-- A property on `LinSols`, satisfied if every point on the line between the two planes +/-- A property on `LinSols`, satisfied if every point on the line between the two planes in the basis through that point is in the cubic. -/ def LineInCubic (S : (PureU1 (2 * n + 1)).LinSols) : Prop := - ∀ (g f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) , + ∀ (g f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) , accCube (2 * n + 1) (a • P g + b • P! f) = 0 lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) : @@ -55,18 +55,18 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) ring lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) : - ∀ (g : Fin n → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f), + ∀ (g : Fin n → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f), accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by intro g f hS linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1 ) - (lineInCubic_expand h g f hS 1 2 ) / 6 -/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/ +/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/ def LineInCubicPerm (S : (PureU1 (2 * n + 1)).LinSols) : Prop := ∀ (M : (FamilyPermutations (2 * n + 1)).group ), LineInCubic ((FamilyPermutations (2 * n + 1)).linSolRep M S) -/-- If `lineInCubicPerm S` then `lineInCubic S`. -/ +/-- If `lineInCubicPerm S` then `lineInCubic S`. -/ lemma lineInCubicPerm_self {S : (PureU1 (2 * n + 1)).LinSols} (hS : LineInCubicPerm S) : LineInCubic S := hS 1 @@ -89,7 +89,7 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ + 1)).LinSols} (S.val (δ!₂ j) - S.val (δ!₁ j)) * accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by intro j g f h - let S' := (FamilyPermutations (2 * n.succ + 1)).linSolRep + let S' := (FamilyPermutations (2 * n.succ + 1)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j)) S have hSS' : ((FamilyPermutations (2 * n.succ + 1)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j))) S = S' := rfl @@ -119,7 +119,7 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} have h2 := Pa_δa₂ f g 0 rw [← hS] at h1 h2 h4 simp at h2 - have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0):= by + have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0):= by linear_combination -(1 * h1) - 1 * h4 - 1 * h2 rw [h5] rw [δa₄_δ!₂] @@ -133,7 +133,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols} (LIC : LineInCubicPerm S) : LineInPlaneProp ((S.val (δ!₂ 0)), ((S.val (δ!₁ 0)), (S.val δ!₃))) := by obtain ⟨g, f, hfg⟩ := span_basis S - have h1 := lineInCubicPerm_swap LIC 0 g f hfg + have h1 := lineInCubicPerm_swap LIC 0 g f hfg rw [P_P_P!_accCube' g f hfg] at h1 simp at h1 cases h1 <;> rename_i h1 @@ -147,7 +147,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols} apply Or.inr linear_combination h1 -lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} +lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} (LIC : LineInCubicPerm S) : LineInPlaneCond S := by refine @Prop_three (2 * n.succ.succ + 1) LineInPlaneProp S (δ!₂ 0) (δ!₁ 0) δ!₃ ?_ ?_ ?_ ?_ @@ -157,11 +157,11 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} intro M exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M) -lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} +lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} (LIC : LineInCubicPerm S) : ConstAbs S.val := linesInPlane_constAbs (lineInCubicPerm_last_perm LIC) -theorem lineInCubicPerm_zero {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} +theorem lineInCubicPerm_zero {S : (PureU1 (2 * n.succ.succ + 1)).LinSols} (LIC : LineInCubicPerm S) : S = 0 := ConstAbs.boundary_value_odd S (lineInCubicPerm_constAbs LIC) diff --git a/HepLean/AnomalyCancellation/PureU1/Odd/Parameterization.lean b/HepLean/AnomalyCancellation/PureU1/Odd/Parameterization.lean index 56ddddd..a4f28d2 100644 --- a/HepLean/AnomalyCancellation/PureU1/Odd/Parameterization.lean +++ b/HepLean/AnomalyCancellation/PureU1/Odd/Parameterization.lean @@ -14,7 +14,7 @@ import Mathlib.RepresentationTheory.Basic /-! # Parameterization in odd case -Given maps `g : Fin n → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly +Given maps `g : Fin n → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly equations. We show that every solution can be got in this way, up to permutation, unless it is zero. The main reference is: @@ -31,7 +31,7 @@ open VectorLikeOddPlane /-- Given a `g f : Fin n → ℚ` and a `a : ℚ` we form a linear solution. We will later show that this can be extended to a complete solution. -/ -def parameterizationAsLinear (g f : Fin n → ℚ) (a : ℚ) : +def parameterizationAsLinear (g f : Fin n → ℚ) (a : ℚ) : (PureU1 (2 * n + 1)).LinSols := a • ((accCubeTriLinSymm (P! f) (P! f) (P g)) • P' g + (- accCubeTriLinSymm (P g) (P g) (P! f)) • P!' f) @@ -44,7 +44,7 @@ lemma parameterizationAsLinear_val (g f : Fin n → ℚ) (a : ℚ) : change a • (_ • (P' g).val + _ • (P!' f).val) = _ rw [P'_val, P!'_val] -lemma parameterizationCharge_cube (g f : Fin n → ℚ) (a : ℚ): +lemma parameterizationCharge_cube (g f : Fin n → ℚ) (a : ℚ): (accCube (2 * n + 1)) (parameterizationAsLinear g f a).val = 0 := by change accCubeTriLinSymm.toCubic _ = 0 rw [parameterizationAsLinear_val] @@ -64,7 +64,7 @@ def parameterization (g f : Fin n → ℚ) (a : ℚ) : parameterizationCharge_cube g f a⟩ lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols} - (g f : Fin n → ℚ) (hS : S.val = P g + P! f) : + (g f : Fin n → ℚ) (hS : S.val = P g + P! f) : accCubeTriLinSymm (P g) (P g) (P! f) = - accCubeTriLinSymm (P! f) (P! f) (P g) := by have hC := S.cubicSol @@ -75,15 +75,15 @@ lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols} erw [P!_accCube] at hC linear_combination hC / 3 -/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. +/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. In this case our parameterization above will be able to recover this point. -/ def GenericCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop := - ∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) , - accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 + ∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) , + accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols) (hs : ∃ (g f : Fin n.succ → ℚ), S.val = P g + P! f ∧ - accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by + accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by intro g f hS hC obtain ⟨g', f', hS', hC'⟩ := hs rw [hS] at hS' @@ -91,15 +91,15 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols) rw [hS'.1, hS'.2] at hC exact hC' hC -/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. +/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`. In this case we will show that S is zero if it is true for all permutations. -/ -def SpecialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop := +def SpecialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop := ∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) , accCubeTriLinSymm (P g) (P g) (P! f) = 0 lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols) (hs : ∃ (g f : Fin n.succ → ℚ), S.val = P g + P! f ∧ - accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by + accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by intro g f hS obtain ⟨g', f', hS', hC'⟩ := hs rw [hS] at hS' @@ -110,7 +110,7 @@ lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols) lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) : GenericCase S ∨ SpecialCase S := by obtain ⟨g, f, h⟩ := span_basis S.1.1 - have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨ + have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨ accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by exact ne_or_eq _ _ cases h1 <;> rename_i h1 @@ -118,7 +118,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) : exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩) theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) : - ∃ g f a, S = parameterization g f a := by + ∃ g f a, S = parameterization g f a := by obtain ⟨g, f, hS⟩ := span_basis S.1.1 use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹ rw [parameterization] @@ -164,7 +164,7 @@ theorem special_case {S : (PureU1 (2 * n.succ.succ + 1)).Sols} (h : ∀ (M : (FamilyPermutations (2 * n.succ.succ + 1)).group), SpecialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun S M)) : S.1.1 = 0 := by - have ht := special_case_lineInCubic_perm h + have ht := special_case_lineInCubic_perm h exact lineInCubicPerm_zero ht end Odd end PureU1 diff --git a/HepLean/AnomalyCancellation/PureU1/Permutations.lean b/HepLean/AnomalyCancellation/PureU1/Permutations.lean index b20fb59..fd44944 100644 --- a/HepLean/AnomalyCancellation/PureU1/Permutations.lean +++ b/HepLean/AnomalyCancellation/PureU1/Permutations.lean @@ -18,13 +18,13 @@ We further define the action on the ACC System. universe v u open Nat -open Finset +open Finset namespace PureU1 /-- The permutation group of the n-fermions. -/ @[simp] -def PermGroup (n : ℕ) := Equiv.Perm (Fin n) +def PermGroup (n : ℕ) := Equiv.Perm (Fin n) instance {n : ℕ} : Group (PermGroup n) := by simp [PermGroup] @@ -35,7 +35,7 @@ section Charges /-- The image of an element of `permGroup` under the representation on charges. -/ @[simps!] def chargeMap {n : ℕ} (f : PermGroup n) : - (PureU1 n).Charges →ₗ[ℚ] (PureU1 n).Charges where + (PureU1 n).Charges →ₗ[ℚ] (PureU1 n).Charges where toFun S := S ∘ f.toFun map_add' S T := by funext i @@ -73,7 +73,7 @@ open BigOperators lemma accCube_invariant {n : ℕ} (f : (PermGroup n)) (S : (PureU1 n).Charges) : accCube n (permCharges f S) = accCube n S := by rw [accCube_explicit, accCube_explicit] - change ∑ i : Fin n, ((((fun a => a^3) ∘ S) (f.symm i))) = _ + change ∑ i : Fin n, ((((fun a => a^3) ∘ S) (f.symm i))) = _ apply (Equiv.Perm.sum_comp _ _ _ ?_) simp @@ -130,22 +130,22 @@ def pairSwap {n : ℕ} (i j : Fin n) : (FamilyPermutations n).group where right_inv s := by aesop -lemma pairSwap_self_inv {n : ℕ} (i j : Fin n) : (pairSwap i j)⁻¹ = (pairSwap i j) := by +lemma pairSwap_self_inv {n : ℕ} (i j : Fin n) : (pairSwap i j)⁻¹ = (pairSwap i j) := by rfl -lemma pairSwap_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun i = j := by +lemma pairSwap_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun i = j := by simp [pairSwap] -lemma pairSwap_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun j = i := by +lemma pairSwap_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun j = i := by simp [pairSwap] -lemma pairSwap_inv_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun i = j := by +lemma pairSwap_inv_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun i = j := by simp [pairSwap] -lemma pairSwap_inv_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun j = i := by +lemma pairSwap_inv_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun j = i := by simp [pairSwap] -lemma pairSwap_other {n : ℕ} (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) : +lemma pairSwap_other {n : ℕ} (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) : (pairSwap i j).toFun k = k := by simp [pairSwap] split @@ -154,7 +154,7 @@ lemma pairSwap_other {n : ℕ} (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) simp_all rfl -lemma pairSwap_inv_other {n : ℕ} {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) : +lemma pairSwap_inv_other {n : ℕ} {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) : (pairSwap i j).invFun k = k := by simp [pairSwap] split @@ -164,12 +164,12 @@ lemma pairSwap_inv_other {n : ℕ} {i j k : Fin n} (hik : i ≠ k) (hjk : j rfl /-- A permutation of fermions which takes one ordered subset into another. -/ -noncomputable def permOfInjection (f g : Fin m ↪ Fin n) : (FamilyPermutations n).group := +noncomputable def permOfInjection (f g : Fin m ↪ Fin n) : (FamilyPermutations n).group := Equiv.extendSubtype (g.toEquivRange.symm.trans f.toEquivRange) section permTwo -variable {i j i' j' : Fin n} (hij : i ≠ j) (hij' : i' ≠ j') +variable {i j i' j' : Fin n} (hij : i ≠ j) (hij' : i' ≠ j') /-- Given two distinct elements, an embedding of `Fin 2` into `Fin n`. -/ def permTwoInj : Fin 2 ↪ Fin n where @@ -185,7 +185,7 @@ lemma permTwoInj_fst : i ∈ Set.range ⇑(permTwoInj hij) := by rfl lemma permTwoInj_fst_apply : - (Function.Embedding.toEquivRange (permTwoInj hij)).symm ⟨i, permTwoInj_fst hij⟩ = 0 := by + (Function.Embedding.toEquivRange (permTwoInj hij)).symm ⟨i, permTwoInj_fst hij⟩ = 0 := by exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permTwoInj hij))).mpr rfl lemma permTwoInj_snd : j ∈ Set.range ⇑(permTwoInj hij) := by @@ -195,12 +195,12 @@ lemma permTwoInj_snd : j ∈ Set.range ⇑(permTwoInj hij) := by lemma permTwoInj_snd_apply : (Function.Embedding.toEquivRange (permTwoInj hij)).symm - ⟨j, permTwoInj_snd hij⟩ = 1 := by + ⟨j, permTwoInj_snd hij⟩ = 1 := by exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permTwoInj hij))).mpr rfl /-- A permutation which swaps `i` with `i'` and `j` with `j'`. -/ noncomputable def permTwo : (FamilyPermutations n).group := - permOfInjection (permTwoInj hij) (permTwoInj hij') + permOfInjection (permTwoInj hij) (permTwoInj hij') lemma permTwo_fst : (permTwo hij hij').toFun i' = i := by rw [permTwo, permOfInjection] @@ -243,7 +243,7 @@ lemma permThreeInj_fst : i ∈ Set.range ⇑(permThreeInj hij hjk hik) := by lemma permThreeInj_fst_apply : (Function.Embedding.toEquivRange (permThreeInj hij hjk hik)).symm - ⟨i, permThreeInj_fst hij hjk hik⟩ = 0 := by + ⟨i, permThreeInj_fst hij hjk hik⟩ = 0 := by exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permThreeInj hij hjk hik))).mpr rfl lemma permThreeInj_snd : j ∈ Set.range ⇑(permThreeInj hij hjk hik) := by @@ -253,7 +253,7 @@ lemma permThreeInj_snd : j ∈ Set.range ⇑(permThreeInj hij hjk hik) := by lemma permThreeInj_snd_apply : (Function.Embedding.toEquivRange (permThreeInj hij hjk hik)).symm - ⟨j, permThreeInj_snd hij hjk hik⟩ = 1 := by + ⟨j, permThreeInj_snd hij hjk hik⟩ = 1 := by exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permThreeInj hij hjk hik))).mpr rfl lemma permThreeInj_thd : k ∈ Set.range ⇑(permThreeInj hij hjk hik) := by @@ -268,7 +268,7 @@ lemma permThreeInj_thd_apply : /-- A permutation which swaps three distinct elements with another three. -/ noncomputable def permThree : (FamilyPermutations n).group := - permOfInjection (permThreeInj hij hjk hik) (permThreeInj hij' hjk' hik') + permOfInjection (permThreeInj hij hjk hik) (permThreeInj hij' hjk' hik') lemma permThree_fst : (permThree hij hjk hik hij' hjk' hik').toFun i' = i := by rw [permThree, permOfInjection] @@ -310,7 +310,7 @@ lemma Prop_two (P : ℚ × ℚ → Prop) {S : (PureU1 n).LinSols} have h1 := h (permTwo hij hab ).symm rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply] at h1 simp at h1 - change P + change P (S.val ((permTwo hij hab ).toFun a), S.val ((permTwo hij hab ).toFun b)) at h1 erw [permTwo_fst,permTwo_snd] at h1 @@ -322,14 +322,14 @@ lemma Prop_three (P : ℚ × ℚ × ℚ → Prop) {S : (PureU1 n).LinSols} P ((((FamilyPermutations n).linSolRep f S).val a),( (((FamilyPermutations n).linSolRep f S).val b), (((FamilyPermutations n).linSolRep f S).val c) - ))) : ∀ (i j k : Fin n) (_ : i ≠ j) (_ : j ≠ k) (_ : i ≠ k), + ))) : ∀ (i j k : Fin n) (_ : i ≠ j) (_ : j ≠ k) (_ : i ≠ k), P (S.val i, (S.val j, S.val k)) := by intro i j k hij hjk hik have h1 := h (permThree hij hjk hik hab hbc hac).symm rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply] at h1 simp at h1 - change P + change P (S.val ((permThree hij hjk hik hab hbc hac).toFun a), S.val ((permThree hij hjk hik hab hbc hac).toFun b), S.val ((permThree hij hjk hik hab hbc hac).toFun c)) at h1 diff --git a/HepLean/AnomalyCancellation/PureU1/Sort.lean b/HepLean/AnomalyCancellation/PureU1/Sort.lean index eccf84b..43e3698 100644 --- a/HepLean/AnomalyCancellation/PureU1/Sort.lean +++ b/HepLean/AnomalyCancellation/PureU1/Sort.lean @@ -15,7 +15,7 @@ We define the sort function for Pure U(1) charges, and prove some basic properti universe v u open Nat -open Finset +open Finset namespace PureU1 @@ -23,15 +23,15 @@ variable {n : ℕ} /-- We say a charge is shorted if for all `i ≤ j`, then `S i ≤ S j`. -/ @[simp] -def Sorted {n : ℕ} (S : (PureU1 n).Charges) : Prop := - ∀ i j (_ : i ≤ j), S i ≤ S j +def Sorted {n : ℕ} (S : (PureU1 n).Charges) : Prop := + ∀ i j (_ : i ≤ j), S i ≤ S j /-- Given a charge assignment `S`, the corresponding sorted charge assignment. -/ @[simp] def sort {n : ℕ} (S : (PureU1 n).Charges) : (PureU1 n).Charges := ((FamilyPermutations n).rep (Tuple.sort S).symm S) -lemma sort_sorted {n : ℕ} (S : (PureU1 n).Charges) : Sorted (sort S) := by +lemma sort_sorted {n : ℕ} (S : (PureU1 n).Charges) : Sorted (sort S) := by simp only [Sorted, PureU1_numberCharges, sort, FamilyPermutations, PermGroup, permCharges, MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply] intro i j hij @@ -45,7 +45,7 @@ lemma sort_apply {n : ℕ} (S : (PureU1 n).Charges) (j : Fin n) : sort S j = S ((Tuple.sort S) j) := by rfl -lemma sort_zero {n : ℕ} (S : (PureU1 n).Charges) (hS : sort S = 0) : S = 0 := by +lemma sort_zero {n : ℕ} (S : (PureU1 n).Charges) (hS : sort S = 0) : S = 0 := by funext i have hj : ∀ j, sort S j = 0 := by rw [hS] @@ -61,13 +61,13 @@ lemma sort_projection {n : ℕ} (S : (PureU1 n).Charges) : sort (sort S) = sort sort_perm S (Tuple.sort S).symm /-- The sort function acting on `LinSols`. -/ -def sortAFL {n : ℕ} (S : (PureU1 n).LinSols) : (PureU1 n).LinSols := +def sortAFL {n : ℕ} (S : (PureU1 n).LinSols) : (PureU1 n).LinSols := ((FamilyPermutations n).linSolRep (Tuple.sort S.val).symm S) -lemma sortAFL_val {n : ℕ} (S : (PureU1 n).LinSols) : (sortAFL S).val = sort S.val := by +lemma sortAFL_val {n : ℕ} (S : (PureU1 n).LinSols) : (sortAFL S).val = sort S.val := by rfl -lemma sortAFL_zero {n : ℕ} (S : (PureU1 n).LinSols) (hS : sortAFL S = 0) : S = 0 := by +lemma sortAFL_zero {n : ℕ} (S : (PureU1 n).LinSols) (hS : sortAFL S = 0) : S = 0 := by apply ACCSystemLinear.LinSols.ext have h1 : sort S.val = 0 := by rw [← sortAFL_val] diff --git a/HepLean/AnomalyCancellation/PureU1/VectorLike.lean b/HepLean/AnomalyCancellation/PureU1/VectorLike.lean index d0d9ce0..54f66c1 100644 --- a/HepLean/AnomalyCancellation/PureU1/VectorLike.lean +++ b/HepLean/AnomalyCancellation/PureU1/VectorLike.lean @@ -15,7 +15,7 @@ For the `n`-even case we define the property of a charge assignment being vector universe v u open Nat -open Finset +open Finset open BigOperators namespace PureU1 @@ -30,10 +30,10 @@ lemma split_equal (n : ℕ) : n + n = 2 * n := (Nat.two_mul n).symm lemma split_odd (n : ℕ) : n + 1 + n = 2 * n + 1 := by omega /-- A charge configuration for n even is vector like if when sorted the `i`th element -is equal to the negative of the `n + i`th element. -/ +is equal to the negative of the `n + i`th element. -/ @[simp] -def VectorLikeEven (S : (PureU1 (2 * n)).Charges) : Prop := - ∀ (i : Fin n), (sort S) (Fin.cast (split_equal n) (Fin.castAdd n i)) - = - (sort S) (Fin.cast (split_equal n) (Fin.natAdd n i)) +def VectorLikeEven (S : (PureU1 (2 * n)).Charges) : Prop := + ∀ (i : Fin n), (sort S) (Fin.cast (split_equal n) (Fin.castAdd n i)) + = - (sort S) (Fin.cast (split_equal n) (Fin.natAdd n i)) end PureU1 diff --git a/HepLean/AnomalyCancellation/SM/Basic.lean b/HepLean/AnomalyCancellation/SM/Basic.lean index f067f5e..fa49b1c 100644 --- a/HepLean/AnomalyCancellation/SM/Basic.lean +++ b/HepLean/AnomalyCancellation/SM/Basic.lean @@ -31,7 +31,7 @@ def SMSpecies (n : ℕ) : ACCSystemCharges := ACCSystemChargesMk n namespace SMCharges -variable {n : ℕ} +variable {n : ℕ} /-- An equivalence between the set `(SMCharges n).charges` and the set `(Fin 5 → Fin n → ℚ)`. -/ @@ -54,7 +54,7 @@ lemma charges_eq_toSpecies_eq (S T : (SMCharges n).Charges) : apply toSpeciesEquiv.injective exact (Set.eqOn_univ (toSpeciesEquiv S) (toSpeciesEquiv T)).mp fun ⦃x⦄ _ => h x -lemma toSMSpecies_toSpecies_inv (i : Fin 5) (f : (Fin 5 → Fin n → ℚ) ) : +lemma toSMSpecies_toSpecies_inv (i : Fin 5) (f : (Fin 5 → Fin n → ℚ) ) : (toSpecies i) (toSpeciesEquiv.symm f) = f i := by change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i= f i simp @@ -80,7 +80,7 @@ namespace SMACCs open SMCharges -variable {n : ℕ} +variable {n : ℕ} /-- The gravitational anomaly equation. -/ @[simp] @@ -89,7 +89,7 @@ def accGrav : (SMCharges n).Charges →ₗ[ℚ] ℚ where map_add' S T := by simp only repeat rw [map_add] - simp [Pi.add_apply, mul_add] + simp [Pi.add_apply, mul_add] repeat erw [Finset.sum_add_distrib] ring map_smul' a S := by @@ -103,7 +103,7 @@ def accGrav : (SMCharges n).Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accGrav`. -/ lemma accGrav_ext {S T : (SMCharges n).Charges} - (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accGrav S = accGrav T := by simp only [accGrav, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, AddHom.coe_mk] @@ -133,7 +133,7 @@ def accSU2 : (SMCharges n).Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accSU2`. -/ lemma accSU2_ext {S T : (SMCharges n).Charges} - (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accSU2 S = accSU2 T := by simp only [accSU2, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, AddHom.coe_mk] @@ -162,7 +162,7 @@ def accSU3 : (SMCharges n).Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accSU3`. -/ lemma accSU3_ext {S T : (SMCharges n).Charges} - (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accSU3 S = accSU3 T := by simp only [accSU3, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, AddHom.coe_mk] @@ -193,7 +193,7 @@ def accYY : (SMCharges n).Charges →ₗ[ℚ] ℚ where /-- Extensionality lemma for `accYY`. -/ lemma accYY_ext {S T : (SMCharges n).Charges} - (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accYY S = accYY T := by simp only [accYY, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk, AddHom.coe_mk] @@ -238,7 +238,7 @@ def quadBiLin : BiLinearSymm (SMCharges n).Charges := BiLinearSymm.mk₂ /-- The quadratic anomaly cancellation condition. -/ @[simp] -def accQuad : HomogeneousQuadratic (SMCharges n).Charges := +def accQuad : HomogeneousQuadratic (SMCharges n).Charges := (@quadBiLin n).toHomogeneousQuad /-- Extensionality lemma for `accQuad`. -/ @@ -263,7 +263,7 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃ + 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i)) + 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i)) + 2 * ((L S.1 i) * (L S.2.1 i) * (L S.2.2 i)) - + ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i)))) + + ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i)))) (by intro a S T R simp only diff --git a/HepLean/AnomalyCancellation/SM/FamilyMaps.lean b/HepLean/AnomalyCancellation/SM/FamilyMaps.lean index ae6ee63..d500301 100644 --- a/HepLean/AnomalyCancellation/SM/FamilyMaps.lean +++ b/HepLean/AnomalyCancellation/SM/FamilyMaps.lean @@ -37,7 +37,7 @@ def chargesMapOfSpeciesMap {n m : ℕ} (f : (SMSpecies n).Charges →ₗ[ℚ] (S rw [map_smul] rfl -/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/ +/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/ @[simps!] def speciesFamilyProj {m n : ℕ} (h : n ≤ m) : (SMSpecies m).Charges →ₗ[ℚ] (SMSpecies n).Charges where @@ -50,12 +50,12 @@ def speciesFamilyProj {m n : ℕ} (h : n ≤ m) : simp [HSMul.hSMul] --rw [show Rat.cast a = a from rfl] -/-- The projection of the `m`-family charges onto the first `n`-family charges. -/ +/-- The projection of the `m`-family charges onto the first `n`-family charges. -/ def familyProjection {m n : ℕ} (h : n ≤ m) : (SMCharges m).Charges →ₗ[ℚ] (SMCharges n).Charges := chargesMapOfSpeciesMap (speciesFamilyProj h) /-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all -other charges zero. -/ +other charges zero. -/ @[simps!] def speciesEmbed (m n : ℕ) : (SMSpecies m).Charges →ₗ[ℚ] (SMSpecies n).Charges where @@ -80,7 +80,7 @@ def speciesEmbed (m n : ℕ) : simp /-- The embedding of the `m`-family charges onto the `n`-family charges, with all -other charges zero. -/ +other charges zero. -/ def familyEmbedding (m n : ℕ) : (SMCharges m).Charges →ₗ[ℚ] (SMCharges n).Charges := chargesMapOfSpeciesMap (speciesEmbed m n) diff --git a/HepLean/AnomalyCancellation/SM/NoGrav/Basic.lean b/HepLean/AnomalyCancellation/SM/NoGrav/Basic.lean index 2b61208..b8f0926 100644 --- a/HepLean/AnomalyCancellation/SM/NoGrav/Basic.lean +++ b/HepLean/AnomalyCancellation/SM/NoGrav/Basic.lean @@ -36,17 +36,17 @@ namespace SMNoGrav variable {n : ℕ} -lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by +lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 0 -lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by +lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 1 -lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by +lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by exact S.cubicSol /-- An element of `charges` which satisfies the linear ACCs @@ -84,7 +84,7 @@ def chargeToAF (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S (hc : accCube S = 0) : (SMNoGrav n).Sols := quadToAF (chargeToQuad S hSU2 hSU3) hc -/-- An element of `AnomalyFreeLinear` which satisfies the quadratic and cubic ACCs +/-- An element of `AnomalyFreeLinear` which satisfies the quadratic and cubic ACCs gives us a element of `AnomalyFree`. -/ def linearToAF (S : (SMNoGrav n).LinSols) (hc : accCube S.val = 0) : (SMNoGrav n).Sols := diff --git a/HepLean/AnomalyCancellation/SM/NoGrav/One/Lemmas.lean b/HepLean/AnomalyCancellation/SM/NoGrav/One/Lemmas.lean index c8d46eb..50de299 100644 --- a/HepLean/AnomalyCancellation/SM/NoGrav/One/Lemmas.lean +++ b/HepLean/AnomalyCancellation/SM/NoGrav/One/Lemmas.lean @@ -25,11 +25,11 @@ open SMACCs open BigOperators lemma E_zero_iff_Q_zero {S : (SMNoGrav 1).Sols} : Q S.val (0 : Fin 1) = 0 ↔ - E S.val (0 : Fin 1) = 0 := by + E S.val (0 : Fin 1) = 0 := by let S' := linearParameters.bijection.symm S.1.1 have hC := cubeSol S have hS' := congrArg (fun S => S.val) (linearParameters.bijection.right_inv S.1.1) - change S'.asCharges = S.val at hS' + change S'.asCharges = S.val at hS' rw [← hS'] at hC apply Iff.intro intro hQ @@ -37,7 +37,7 @@ lemma E_zero_iff_Q_zero {S : (SMNoGrav 1).Sols} : Q S.val (0 : Fin 1) = 0 ↔ intro hE exact S'.cubic_zero_E'_zero hC hE -lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) : +lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) : accGrav S.val = 0 := by rw [accGrav] simp only [SMSpecies_numberCharges, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue, @@ -59,12 +59,12 @@ lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0 have hC := cubeSol S have hS' := congrArg (fun S => S.val.val) (linearParametersQENeqZero.bijection.right_inv ⟨S.1.1, And.intro hQ hE⟩) - change _ = S.val at hS' + change _ = S.val at hS' rw [← hS'] at hC rw [← hS'] exact S'.grav_of_cubic hC FLTThree -/-- Any solution to the ACCs without gravity satisfies the gravitational ACC. -/ +/-- Any solution to the ACCs without gravity satisfies the gravitational ACC. -/ theorem accGravSatisfied {S : (SMNoGrav 1).Sols} (FLTThree : FermatLastTheoremWith ℚ 3) : accGrav S.val = 0 := by by_cases hQ : Q S.val (0 : Fin 1)= 0 diff --git a/HepLean/AnomalyCancellation/SM/NoGrav/One/LinearParameterization.lean b/HepLean/AnomalyCancellation/SM/NoGrav/One/LinearParameterization.lean index 126bf37..62c31f6 100644 --- a/HepLean/AnomalyCancellation/SM/NoGrav/One/LinearParameterization.lean +++ b/HepLean/AnomalyCancellation/SM/NoGrav/One/LinearParameterization.lean @@ -198,15 +198,15 @@ namespace linearParametersQENeqZero @[ext] lemma ext {S T : linearParametersQENeqZero} (hx : S.x = T.x) (hv : S.v = T.v) - (hw : S.w = T.w) : S = T := by + (hw : S.w = T.w) : S = T := by cases' S simp_all only /-- A map from `linearParametersQENeqZero` to `linearParameters`. -/ @[simps!] def toLinearParameters (S : linearParametersQENeqZero) : - {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} := - ⟨⟨S.x, 3 * S.x * (S.v - S.w) / (S.v + S.w), - 6 * S.x / (S.v + S.w)⟩, + {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} := + ⟨⟨S.x, 3 * S.x * (S.v - S.w) / (S.v + S.w), - 6 * S.x / (S.v + S.w)⟩, by apply And.intro S.hx simp only [neg_mul, ne_eq, div_eq_zero_iff, neg_eq_zero, mul_eq_zero, OfNat.ofNat_ne_zero, @@ -217,9 +217,9 @@ def toLinearParameters (S : linearParametersQENeqZero) : /-- A map from `linearParameters` to `linearParametersQENeqZero` in the special case when `Q'` and `E'` of the linear parameters are non-zero. -/ @[simps!] -def tolinearParametersQNeqZero (S : {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0}) : +def tolinearParametersQNeqZero (S : {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0}) : linearParametersQENeqZero := - ⟨S.1.Q', - (3 * S.1.Q' + S.1.Y) / S.1.E', - (3 * S.1.Q' - S.1.Y)/ S.1.E', S.2.1, + ⟨S.1.Q', - (3 * S.1.Q' + S.1.Y) / S.1.E', - (3 * S.1.Q' - S.1.Y)/ S.1.E', S.2.1, by simp only [ne_eq, neg_add_rev, neg_sub] field_simp @@ -231,7 +231,7 @@ def tolinearParametersQNeqZero (S : {S : linearParameters // S.Q' ≠ 0 ∧ S. with `Q'` and `E'` non-zero. -/ @[simps!] def bijectionLinearParameters : - linearParametersQENeqZero ≃ {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} where + linearParametersQENeqZero ≃ {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} where toFun := toLinearParameters invFun := tolinearParametersQNeqZero left_inv S := by @@ -260,7 +260,7 @@ def bijectionLinearParameters : /-- The bijection between `linearParametersQENeqZero` and `LinSols` with `Q` and `E` non-zero. -/ def bijection : linearParametersQENeqZero ≃ - {S : (SMNoGrav 1).LinSols // Q S.val (0 : Fin 1) ≠ 0 ∧ E S.val (0 : Fin 1) ≠ 0} := + {S : (SMNoGrav 1).LinSols // Q S.val (0 : Fin 1) ≠ 0 ∧ E S.val (0 : Fin 1) ≠ 0} := bijectionLinearParameters.trans (linearParameters.bijectionQEZero) lemma cubic (S : linearParametersQENeqZero) : @@ -297,7 +297,7 @@ lemma cubic_v_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1. have h' : (S.w + 1) * (1 * S.w * S.w + (-1) * S.w + 1) = 0 := by ring_nf exact add_eq_zero_iff_neg_eq.mpr (id (Eq.symm h)) - have h'' : (1 * S.w * S.w + (-1) * S.w + 1) ≠ 0 := by + have h'' : (1 * S.w * S.w + (-1) * S.w + 1) ≠ 0 := by refine quadratic_ne_zero_of_discrim_ne_sq ?_ S.w intro s by_contra hn @@ -315,7 +315,7 @@ lemma cube_w_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.v have h' : (S.v + 1) * (1 * S.v * S.v + (-1) * S.v + 1) = 0 := by ring_nf exact add_eq_zero_iff_neg_eq.mpr (id (Eq.symm h)) - have h'' : (1 * S.v * S.v + (-1) * S.v + 1) ≠ 0 := by + have h'' : (1 * S.v * S.v + (-1) * S.v + 1) ≠ 0 := by refine quadratic_ne_zero_of_discrim_ne_sq ?_ S.v intro s by_contra hn diff --git a/HepLean/AnomalyCancellation/SM/Permutations.lean b/HepLean/AnomalyCancellation/SM/Permutations.lean index 0dda836..35fdc19 100644 --- a/HepLean/AnomalyCancellation/SM/Permutations.lean +++ b/HepLean/AnomalyCancellation/SM/Permutations.lean @@ -16,7 +16,7 @@ We define the group of permutations for the SM charges with no RHN. universe v u open Nat -open Finset +open Finset namespace SM @@ -26,7 +26,7 @@ open BigOperators /-- The group of `Sₙ` permutations for each species. -/ @[simp] -def PermGroup (n : ℕ) := ∀ (_ : Fin 5), Equiv.Perm (Fin n) +def PermGroup (n : ℕ) := ∀ (_ : Fin 5), Equiv.Perm (Fin n) variable {n : ℕ} @@ -72,32 +72,32 @@ lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup n) (S : (SMCharges n).Cha exact Fintype.sum_equiv (f⁻¹ j) (fun x => ((fun a => a ^ m) ∘ (toSpecies j) S ∘ ⇑(f⁻¹ j)) x) (fun x => ((fun a => a ^ m) ∘ (toSpecies j) S) x) (congrFun rfl) -lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : +lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : accGrav (repCharges f S) = accGrav S := accGrav_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accSU2_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : +lemma accSU2_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : accSU2 (repCharges f S) = accSU2 S := accSU2_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accSU3_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : +lemma accSU3_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : accSU3 (repCharges f S) = accSU3 S := accSU3_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accYY_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : +lemma accYY_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : accYY (repCharges f S) = accYY S := accYY_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accQuad_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : +lemma accQuad_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : accQuad (repCharges f S) = accQuad S := accQuad_ext (toSpecies_sum_invariant 2 f S) -lemma accCube_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : +lemma accCube_invariant (f : PermGroup n) (S : (SMCharges n).Charges) : accCube (repCharges f S) = accCube S := accCube_ext (fun j => toSpecies_sum_invariant 3 f S j) diff --git a/HepLean/AnomalyCancellation/SMNu/Basic.lean b/HepLean/AnomalyCancellation/SMNu/Basic.lean index 625adf9..569b2b6 100644 --- a/HepLean/AnomalyCancellation/SMNu/Basic.lean +++ b/HepLean/AnomalyCancellation/SMNu/Basic.lean @@ -28,7 +28,7 @@ def SMνSpecies (n : ℕ) : ACCSystemCharges := ACCSystemChargesMk n namespace SMνCharges -variable {n : ℕ} +variable {n : ℕ} /-- An equivalence between `(SMνCharges n).charges` and `(Fin 6 → Fin n → ℚ)` splitting the charges into species.-/ @@ -54,12 +54,12 @@ lemma charges_eq_toSpecies_eq (S T : (SMνCharges n).Charges) : funext i exact h i -lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin n → ℚ) ) : +lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin n → ℚ) ) : (toSpecies i) (toSpeciesEquiv.symm f) = f i := by change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i = f i simp -lemma toSpecies_one (S : (SMνCharges 1).Charges) (j : Fin 6) : +lemma toSpecies_one (S : (SMνCharges 1).Charges) (j : Fin 6) : toSpecies j S ⟨0, by simp⟩ = S j := by match j with | 0 => rfl @@ -88,7 +88,7 @@ namespace SMνACCs open SMνCharges -variable {n : ℕ} +variable {n : ℕ} /-- The gravitational anomaly equation. -/ @[simp] @@ -97,7 +97,7 @@ def accGrav : (SMνCharges n).Charges →ₗ[ℚ] ℚ where map_add' S T := by simp only repeat rw [map_add] - simp [Pi.add_apply, mul_add] + simp [Pi.add_apply, mul_add] repeat erw [Finset.sum_add_distrib] ring map_smul' a S := by @@ -119,7 +119,7 @@ lemma accGrav_decomp (S : (SMνCharges n).Charges) : /-- Extensionality lemma for `accGrav`. -/ lemma accGrav_ext {S T : (SMνCharges n).Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accGrav S = accGrav T := by rw [accGrav_decomp, accGrav_decomp] repeat erw [hj] @@ -131,7 +131,7 @@ def accSU2 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where map_add' S T := by simp only repeat rw [map_add] - simp [Pi.add_apply, mul_add] + simp [Pi.add_apply, mul_add] repeat erw [Finset.sum_add_distrib] ring map_smul' a S := by @@ -152,7 +152,7 @@ lemma accSU2_decomp (S : (SMνCharges n).Charges) : /-- Extensionality lemma for `accSU2`. -/ lemma accSU2_ext {S T : (SMνCharges n).Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accSU2 S = accSU2 T := by rw [accSU2_decomp, accSU2_decomp] repeat erw [hj] @@ -164,7 +164,7 @@ def accSU3 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where map_add' S T := by simp only repeat rw [map_add] - simp [ Pi.add_apply, mul_add] + simp [ Pi.add_apply, mul_add] repeat erw [Finset.sum_add_distrib] ring map_smul' a S := by @@ -185,7 +185,7 @@ lemma accSU3_decomp (S : (SMνCharges n).Charges) : /-- Extensionality lemma for `accSU3`. -/ lemma accSU3_ext {S T : (SMνCharges n).Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accSU3 S = accSU3 T := by rw [accSU3_decomp, accSU3_decomp] repeat rw [hj] @@ -198,7 +198,7 @@ def accYY : (SMνCharges n).Charges →ₗ[ℚ] ℚ where map_add' S T := by simp only repeat rw [map_add] - simp [Pi.add_apply, mul_add] + simp [Pi.add_apply, mul_add] repeat erw [Finset.sum_add_distrib] ring map_smul' a S := by @@ -219,7 +219,7 @@ lemma accYY_decomp (S : (SMνCharges n).Charges) : /-- Extensionality lemma for `accYY`. -/ lemma accYY_ext {S T : (SMνCharges n).Charges} - (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : + (hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) : accYY S = accYY T := by rw [accYY_decomp, accYY_decomp] repeat rw [hj] @@ -259,8 +259,8 @@ def quadBiLin : BiLinearSymm (SMνCharges n).Charges := BiLinearSymm.mk₂ ring) lemma quadBiLin_decomp (S T : (SMνCharges n).Charges) : - quadBiLin S T = ∑ i, Q S i * Q T i - 2 * ∑ i, U S i * U T i + - ∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by + quadBiLin S T = ∑ i, Q S i * Q T i - 2 * ∑ i, U S i * U T i + + ∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by erw [← quadBiLin.toFun_eq_coe] rw [quadBiLin] simp only [quadBiLin, BiLinearSymm.mk₂, AddHom.toFun_eq_coe, AddHom.coe_mk, LinearMap.coe_mk] @@ -271,7 +271,7 @@ lemma quadBiLin_decomp (S T : (SMνCharges n).Charges) : /-- The quadratic anomaly cancellation condition. -/ @[simp] -def accQuad : HomogeneousQuadratic (SMνCharges n).Charges := +def accQuad : HomogeneousQuadratic (SMνCharges n).Charges := (@quadBiLin n).toHomogeneousQuad lemma accQuad_decomp (S : (SMνCharges n).Charges) : @@ -296,8 +296,8 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃ + 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i)) + 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i)) + 2 * ((L S.1 i) * (L S.2.1 i) * (L S.2.2 i)) - + ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i)) - + ((N S.1 i) * (N S.2.1 i) * (N S.2.2 i)))) + + ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i)) + + ((N S.1 i) * (N S.2.1 i) * (N S.2.2 i)))) (by intro a S T R simp only @@ -330,8 +330,8 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃ ring) lemma cubeTriLin_decomp (S T R : (SMνCharges n).Charges) : - cubeTriLin S T R = 6 * ∑ i, (Q S i * Q T i * Q R i) + 3 * ∑ i, (U S i * U T i * U R i) + - 3 * ∑ i, (D S i * D T i * D R i) + 2 * ∑ i, (L S i * L T i * L R i) + + cubeTriLin S T R = 6 * ∑ i, (Q S i * Q T i * Q R i) + 3 * ∑ i, (U S i * U T i * U R i) + + 3 * ∑ i, (D S i * D T i * D R i) + 2 * ∑ i, (L S i * L T i * L R i) + ∑ i, (E S i * E T i * E R i) + ∑ i, (N S i * N T i * N R i) := by erw [← cubeTriLin.toFun_eq_coe] rw [cubeTriLin] diff --git a/HepLean/AnomalyCancellation/SMNu/FamilyMaps.lean b/HepLean/AnomalyCancellation/SMNu/FamilyMaps.lean index 801ae0d..d25e26b 100644 --- a/HepLean/AnomalyCancellation/SMNu/FamilyMaps.lean +++ b/HepLean/AnomalyCancellation/SMNu/FamilyMaps.lean @@ -5,7 +5,7 @@ Authors: Joseph Tooby-Smith -/ import HepLean.AnomalyCancellation.SMNu.Basic /-! -# Family maps for the Standard Model for RHN ACCs +# Family maps for the Standard Model for RHN ACCs We define the a series of maps between the charges for different numbers of families. @@ -39,10 +39,10 @@ def chargesMapOfSpeciesMap {n m : ℕ} (f : (SMνSpecies n).Charges →ₗ[ℚ] lemma chargesMapOfSpeciesMap_toSpecies {n m : ℕ} (f : (SMνSpecies n).Charges →ₗ[ℚ] (SMνSpecies m).Charges) (S : (SMνCharges n).Charges) (j : Fin 6) : - toSpecies j (chargesMapOfSpeciesMap f S) = (LinearMap.comp f (toSpecies j)) S := by + toSpecies j (chargesMapOfSpeciesMap f S) = (LinearMap.comp f (toSpecies j)) S := by erw [toSMSpecies_toSpecies_inv] -/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/ +/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/ @[simps!] def speciesFamilyProj {m n : ℕ} (h : n ≤ m) : (SMνSpecies m).Charges →ₗ[ℚ] (SMνSpecies n).Charges where @@ -50,12 +50,12 @@ def speciesFamilyProj {m n : ℕ} (h : n ≤ m) : map_add' _ _ := rfl map_smul' _ _ := rfl -/-- The projection of the `m`-family charges onto the first `n`-family charges. -/ +/-- The projection of the `m`-family charges onto the first `n`-family charges. -/ def familyProjection {m n : ℕ} (h : n ≤ m) : (SMνCharges m).Charges →ₗ[ℚ] (SMνCharges n).Charges := chargesMapOfSpeciesMap (speciesFamilyProj h) /-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all -other charges zero. -/ +other charges zero. -/ @[simps!] def speciesEmbed (m n : ℕ) : (SMνSpecies m).Charges →ₗ[ℚ] (SMνSpecies n).Charges where @@ -80,7 +80,7 @@ def speciesEmbed (m n : ℕ) : exact Eq.symm (Rat.mul_zero a) /-- The embedding of the `m`-family charges onto the `n`-family charges, with all -other charges zero. -/ +other charges zero. -/ def familyEmbedding (m n : ℕ) : (SMνCharges m).Charges →ₗ[ℚ] (SMνCharges n).Charges := chargesMapOfSpeciesMap (speciesEmbed m n) @@ -118,7 +118,7 @@ lemma sum_familyUniversal {n : ℕ} (m : ℕ) (S : (SMνCharges 1).Charges) (j : intro i _ erw [toSpecies_familyUniversal] -lemma sum_familyUniversal_one {n : ℕ} (S : (SMνCharges 1).Charges) (j : Fin 6) : +lemma sum_familyUniversal_one {n : ℕ} (S : (SMνCharges 1).Charges) (j : Fin 6) : ∑ i, toSpecies j (familyUniversal n S) i = n * (toSpecies j S ⟨0, by simp⟩) := by simpa using @sum_familyUniversal n 1 S j @@ -134,7 +134,7 @@ lemma sum_familyUniversal_two {n : ℕ} (S : (SMνCharges 1).Charges) erw [toSpecies_familyUniversal] rfl -lemma sum_familyUniversal_three {n : ℕ} (S : (SMνCharges 1).Charges) +lemma sum_familyUniversal_three {n : ℕ} (S : (SMνCharges 1).Charges) (T L : (SMνCharges n).Charges) (j : Fin 6) : ∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i * toSpecies j L i) = (toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i * toSpecies j L i := by @@ -148,7 +148,7 @@ lemma sum_familyUniversal_three {n : ℕ} (S : (SMνCharges 1).Charges) ring lemma familyUniversal_accGrav (S : (SMνCharges 1).Charges) : - accGrav (familyUniversal n S) = n * (accGrav S) := by + accGrav (familyUniversal n S) = n * (accGrav S) := by rw [accGrav_decomp, accGrav_decomp] repeat rw [sum_familyUniversal_one] simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply, @@ -156,7 +156,7 @@ lemma familyUniversal_accGrav (S : (SMνCharges 1).Charges) : ring lemma familyUniversal_accSU2 (S : (SMνCharges 1).Charges) : - accSU2 (familyUniversal n S) = n * (accSU2 S) := by + accSU2 (familyUniversal n S) = n * (accSU2 S) := by rw [accSU2_decomp, accSU2_decomp] repeat rw [sum_familyUniversal_one] simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply, @@ -164,7 +164,7 @@ lemma familyUniversal_accSU2 (S : (SMνCharges 1).Charges) : ring lemma familyUniversal_accSU3 (S : (SMνCharges 1).Charges) : - accSU3 (familyUniversal n S) = n * (accSU3 S) := by + accSU3 (familyUniversal n S) = n * (accSU3 S) := by rw [accSU3_decomp, accSU3_decomp] repeat rw [sum_familyUniversal_one] simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply, @@ -172,7 +172,7 @@ lemma familyUniversal_accSU3 (S : (SMνCharges 1).Charges) : ring lemma familyUniversal_accYY (S : (SMνCharges 1).Charges) : - accYY (familyUniversal n S) = n * (accYY S) := by + accYY (familyUniversal n S) = n * (accYY S) := by rw [accYY_decomp, accYY_decomp] repeat rw [sum_familyUniversal_one] simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply, @@ -182,7 +182,7 @@ lemma familyUniversal_accYY (S : (SMνCharges 1).Charges) : lemma familyUniversal_quadBiLin (S : (SMνCharges 1).Charges) (T : (SMνCharges n).Charges) : quadBiLin (familyUniversal n S) T = S (0 : Fin 6) * ∑ i, Q T i - 2 * S (1 : Fin 6) * ∑ i, U T i + S (2 : Fin 6) *∑ i, D T i - - S (3 : Fin 6) * ∑ i, L T i + S (4 : Fin 6) * ∑ i, E T i := by + S (3 : Fin 6) * ∑ i, L T i + S (4 : Fin 6) * ∑ i, E T i := by rw [quadBiLin_decomp] repeat rw [sum_familyUniversal_two] repeat rw [toSpecies_one] @@ -191,7 +191,7 @@ lemma familyUniversal_quadBiLin (S : (SMνCharges 1).Charges) (T : (SMνCharges ring lemma familyUniversal_accQuad (S : (SMνCharges 1).Charges) : - accQuad (familyUniversal n S) = n * (accQuad S) := by + accQuad (familyUniversal n S) = n * (accQuad S) := by rw [accQuad_decomp] repeat erw [sum_familyUniversal] rw [accQuad_decomp] @@ -201,7 +201,7 @@ lemma familyUniversal_accQuad (S : (SMνCharges 1).Charges) : lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).Charges) (T R : (SMνCharges n).Charges) : cubeTriLin (familyUniversal n S) T R = 6 * S (0 : Fin 6) * ∑ i, (Q T i * Q R i) + - 3 * S (1 : Fin 6) * ∑ i, (U T i * U R i) + 3 * S (2 : Fin 6) * ∑ i, (D T i * D R i) + 3 * S (1 : Fin 6) * ∑ i, (U T i * U R i) + 3 * S (2 : Fin 6) * ∑ i, (D T i * D R i) + 2 * S (3 : Fin 6) * ∑ i, (L T i * L R i) + S (4 : Fin 6) * ∑ i, (E T i * E R i) + S (5 : Fin 6) * ∑ i, (N T i * N R i) := by rw [cubeTriLin_decomp] diff --git a/HepLean/AnomalyCancellation/SMNu/NoGrav/Basic.lean b/HepLean/AnomalyCancellation/SMNu/NoGrav/Basic.lean index 4f854b2..a1d09ca 100644 --- a/HepLean/AnomalyCancellation/SMNu/NoGrav/Basic.lean +++ b/HepLean/AnomalyCancellation/SMNu/NoGrav/Basic.lean @@ -38,17 +38,17 @@ namespace SMNoGrav variable {n : ℕ} -lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by +lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 0 -lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by +lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 1 -lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by +lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by exact S.cubicSol /-- An element of `charges` which satisfies the linear ACCs @@ -86,7 +86,7 @@ def chargeToAF (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S (hc : accCube S = 0) : (SMNoGrav n).Sols := quadToAF (chargeToQuad S hSU2 hSU3) hc -/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs +/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs gives us a element of `Sols`. -/ def linearToAF (S : (SMNoGrav n).LinSols) (hc : accCube S.val = 0) : (SMNoGrav n).Sols := diff --git a/HepLean/AnomalyCancellation/SMNu/Ordinary/Basic.lean b/HepLean/AnomalyCancellation/SMNu/Ordinary/Basic.lean index df41764..d2cb949 100644 --- a/HepLean/AnomalyCancellation/SMNu/Ordinary/Basic.lean +++ b/HepLean/AnomalyCancellation/SMNu/Ordinary/Basic.lean @@ -39,22 +39,22 @@ namespace SM variable {n : ℕ} -lemma gravSol (S : (SM n).LinSols) : accGrav S.val = 0 := by +lemma gravSol (S : (SM n).LinSols) : accGrav S.val = 0 := by have hS := S.linearSol simp at hS exact hS 0 -lemma SU2Sol (S : (SM n).LinSols) : accSU2 S.val = 0 := by +lemma SU2Sol (S : (SM n).LinSols) : accSU2 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 1 -lemma SU3Sol (S : (SM n).LinSols) : accSU3 S.val = 0 := by +lemma SU3Sol (S : (SM n).LinSols) : accSU3 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 2 -lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := by +lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := by exact S.cubicSol /-- An element of `charges` which satisfies the linear ACCs @@ -94,13 +94,13 @@ def chargeToAF (S : (SM n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0 (hSU3 : accSU3 S = 0) (hc : accCube S = 0) : (SM n).Sols := quadToAF (chargeToQuad S hGrav hSU2 hSU3) hc -/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs +/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs gives us a element of `Sols`. -/ def linearToAF (S : (SM n).LinSols) (hc : accCube S.val = 0) : (SM n).Sols := quadToAF (linearToQuad S) hc -/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/ +/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/ def perm (n : ℕ) : ACCSystemGroupAction (SM n) where group := PermGroup n groupInst := inferInstance diff --git a/HepLean/AnomalyCancellation/SMNu/Ordinary/DimSevenPlane.lean b/HepLean/AnomalyCancellation/SMNu/Ordinary/DimSevenPlane.lean index a8201a3..dc30736 100644 --- a/HepLean/AnomalyCancellation/SMNu/Ordinary/DimSevenPlane.lean +++ b/HepLean/AnomalyCancellation/SMNu/Ordinary/DimSevenPlane.lean @@ -32,7 +32,7 @@ def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i => ) lemma B₀_cubic (S T : (SM 3).Charges) : cubeTriLin B₀ S T = - 6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by + 6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by simp [Fin.sum_univ_three, B₀, Fin.divNat, Fin.modNat, finProdFinEquiv] ring @@ -45,7 +45,7 @@ def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i => ) lemma B₁_cubic (S T : (SM 3).Charges) : cubeTriLin B₁ S T = - 3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by + 3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by simp [Fin.sum_univ_three, B₁, Fin.divNat, Fin.modNat, finProdFinEquiv] ring @@ -58,7 +58,7 @@ def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i => ) lemma B₂_cubic (S T : (SM 3).Charges) : cubeTriLin B₂ S T = - 3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by + 3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by simp [Fin.sum_univ_three, B₂, Fin.divNat, Fin.modNat, finProdFinEquiv] ring @@ -71,7 +71,7 @@ def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i => ) lemma B₃_cubic (S T : (SM 3).Charges) : cubeTriLin B₃ S T = - 2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by + 2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by simp [Fin.sum_univ_three, B₃, Fin.divNat, Fin.modNat, finProdFinEquiv] ring_nf rfl @@ -85,7 +85,7 @@ def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i => ) lemma B₄_cubic (S T : (SM 3).Charges) : cubeTriLin B₄ S T = - (S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by + (S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by simp [Fin.sum_univ_three, B₄, Fin.divNat, Fin.modNat, finProdFinEquiv] ring_nf rfl @@ -99,7 +99,7 @@ def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i => ) lemma B₅_cubic (S T : (SM 3).Charges) : cubeTriLin B₅ S T = - (S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by + (S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by simp [Fin.sum_univ_three, B₅, Fin.divNat, Fin.modNat, finProdFinEquiv] ring_nf rfl @@ -113,7 +113,7 @@ def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i => ) lemma B₆_cubic (S T : (SM 3).Charges) : cubeTriLin B₆ S T = - 3* (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by + 3 * (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by simp [Fin.sum_univ_three, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv] ring_nf @@ -336,7 +336,7 @@ theorem basis_linear_independent : LinearIndependent ℚ B := by end PlaneSeven theorem seven_dim_plane_exists : ∃ (B : Fin 7 → (SM 3).Charges), - LinearIndependent ℚ B ∧ ∀ (f : Fin 7 → ℚ), (SM 3).IsSolution (∑ i, f i • B i) := by + LinearIndependent ℚ B ∧ ∀ (f : Fin 7 → ℚ), (SM 3).IsSolution (∑ i, f i • B i) := by use PlaneSeven.B apply And.intro exact PlaneSeven.basis_linear_independent diff --git a/HepLean/AnomalyCancellation/SMNu/Permutations.lean b/HepLean/AnomalyCancellation/SMNu/Permutations.lean index 3a1cb19..5a31565 100644 --- a/HepLean/AnomalyCancellation/SMNu/Permutations.lean +++ b/HepLean/AnomalyCancellation/SMNu/Permutations.lean @@ -16,7 +16,7 @@ We define the group of permutations for the SM charges with RHN. universe v u open Nat -open Finset +open Finset namespace SMRHN @@ -69,36 +69,36 @@ lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup n) (S : (SMνCharges n).C ∑ i, ((fun a => a ^ m) ∘ toSpecies j (repCharges f S)) i = ∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by erw [repCharges_toSpecies] - change ∑ i : Fin n, ((fun a => a ^ m) ∘ _) (⇑(f⁻¹ _) i) = ∑ i : Fin n, ((fun a => a ^ m) ∘ _) i + change ∑ i : Fin n, ((fun a => a ^ m) ∘ _) (⇑(f⁻¹ _) i) = ∑ i : Fin n, ((fun a => a ^ m) ∘ _) i refine Equiv.Perm.sum_comp _ _ _ ?_ simp only [PermGroup, Fin.isValue, Pi.inv_apply, ne_eq, coe_univ, Set.subset_univ] -lemma accGrav_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : +lemma accGrav_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : accGrav (repCharges f S) = accGrav S := accGrav_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accSU2_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : +lemma accSU2_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : accSU2 (repCharges f S) = accSU2 S := accSU2_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accSU3_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : +lemma accSU3_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : accSU3 (repCharges f S) = accSU3 S := accSU3_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accYY_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : +lemma accYY_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : accYY (repCharges f S) = accYY S := accYY_ext (by simpa using toSpecies_sum_invariant 1 f S) -lemma accQuad_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : +lemma accQuad_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : accQuad (repCharges f S) = accQuad S := accQuad_ext (toSpecies_sum_invariant 2 f S) -lemma accCube_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : +lemma accCube_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) : accCube (repCharges f S) = accCube S := accCube_ext (by simpa using toSpecies_sum_invariant 3 f S) diff --git a/HepLean/AnomalyCancellation/SMNu/PlusU1/Basic.lean b/HepLean/AnomalyCancellation/SMNu/PlusU1/Basic.lean index 5da3f0e..1689d1a 100644 --- a/HepLean/AnomalyCancellation/SMNu/PlusU1/Basic.lean +++ b/HepLean/AnomalyCancellation/SMNu/PlusU1/Basic.lean @@ -39,32 +39,32 @@ namespace PlusU1 variable {n : ℕ} -lemma gravSol (S : (PlusU1 n).LinSols) : accGrav S.val = 0 := by +lemma gravSol (S : (PlusU1 n).LinSols) : accGrav S.val = 0 := by have hS := S.linearSol simp at hS exact hS 0 -lemma SU2Sol (S : (PlusU1 n).LinSols) : accSU2 S.val = 0 := by +lemma SU2Sol (S : (PlusU1 n).LinSols) : accSU2 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 1 -lemma SU3Sol (S : (PlusU1 n).LinSols) : accSU3 S.val = 0 := by +lemma SU3Sol (S : (PlusU1 n).LinSols) : accSU3 S.val = 0 := by have hS := S.linearSol simp at hS exact hS 2 -lemma YYsol (S : (PlusU1 n).LinSols) : accYY S.val = 0 := by +lemma YYsol (S : (PlusU1 n).LinSols) : accYY S.val = 0 := by have hS := S.linearSol simp at hS exact hS 3 -lemma quadSol (S : (PlusU1 n).QuadSols) : accQuad S.val = 0 := by +lemma quadSol (S : (PlusU1 n).QuadSols) : accQuad S.val = 0 := by have hS := S.quadSol simp at hS exact hS 0 -lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by +lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by exact S.cubicSol /-- An element of `charges` which satisfies the linear ACCs @@ -110,13 +110,13 @@ def chargeToAF (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S (PlusU1 n).Sols := quadToAF (chargeToQuad S hGrav hSU2 hSU3 hYY hQ) hc -/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs +/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs gives us a element of `Sols`. -/ def linearToAF (S : (PlusU1 n).LinSols) (hQ : accQuad S.val = 0) (hc : accCube S.val = 0) : (PlusU1 n).Sols := quadToAF (linearToQuad S hQ) hc -/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/ +/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/ def perm (n : ℕ) : ACCSystemGroupAction (PlusU1 n) where group := PermGroup n groupInst := inferInstance diff --git a/HepLean/AnomalyCancellation/SMNu/PlusU1/HyperCharge.lean b/HepLean/AnomalyCancellation/SMNu/PlusU1/HyperCharge.lean index afb8aa9..4a09b31 100644 --- a/HepLean/AnomalyCancellation/SMNu/PlusU1/HyperCharge.lean +++ b/HepLean/AnomalyCancellation/SMNu/PlusU1/HyperCharge.lean @@ -26,11 +26,11 @@ def Y₁ : (PlusU1 1).Sols where val := fun i => match i with | (0 : Fin 6) => 1 - | (1 : Fin 6) => -4 - | (2 : Fin 6) => 2 - | (3 : Fin 6) => -3 - | (4 : Fin 6) => 6 - | (5 : Fin 6) => 0 + | (1 : Fin 6) => -4 + | (2 : Fin 6) => 2 + | (3 : Fin 6) => -3 + | (4 : Fin 6) => 6 + | (5 : Fin 6) => 0 linearSol := by intro i simp at i @@ -89,7 +89,7 @@ lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ℚ): addQuad S a 0 = a • S rfl lemma on_cubeTriLin (S : (PlusU1 n).Charges) : - cubeTriLin (Y n).val (Y n).val S = 6 * accYY S := by + cubeTriLin (Y n).val (Y n).val S = 6 * accYY S := by erw [familyUniversal_cubeTriLin'] rw [accYY_decomp] simp only [Fin.isValue, Y₁_val, mul_one, SMνSpecies_numberCharges, toSpecies_apply, mul_neg, @@ -103,7 +103,7 @@ lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) : simp lemma on_cubeTriLin' (S : (PlusU1 n).Charges) : - cubeTriLin (Y n).val S S = 6 * accQuad S := by + cubeTriLin (Y n).val S S = 6 * accQuad S := by erw [familyUniversal_cubeTriLin] rw [accQuad_decomp] simp only [Fin.isValue, Y₁_val, mul_one, SMνSpecies_numberCharges, toSpecies_apply, mul_neg, @@ -127,7 +127,7 @@ lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ℚ) : ring lemma add_AFQ_cube (S : (PlusU1 n).QuadSols) (a b : ℚ) : - accCube (a • S.val + b • (Y n).val) = a ^ 3 * accCube S.val := by + accCube (a • S.val + b • (Y n).val) = a ^ 3 * accCube S.val := by rw [add_AFL_cube] rw [cubeTriLin.swap₃] rw [on_cubeTriLin'_ALQ] diff --git a/HepLean/AnomalyCancellation/SMNu/PlusU1/PlaneNonSols.lean b/HepLean/AnomalyCancellation/SMNu/PlusU1/PlaneNonSols.lean index 54c3860..10f363c 100644 --- a/HepLean/AnomalyCancellation/SMNu/PlusU1/PlaneNonSols.lean +++ b/HepLean/AnomalyCancellation/SMNu/PlusU1/PlaneNonSols.lean @@ -50,7 +50,7 @@ def B₆ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, def B₇ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0] /-- A charge assignment forming one of the basis elements of the plane. -/ -def B₈ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0] +def B₈ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0] /-- A charge assignment forming one of the basis elements of the plane. -/ def B₉ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 2, 0] @@ -79,7 +79,7 @@ lemma Bi_Bj_quad {i j : Fin 11} (hi : i ≠ j) : quadBiLin (B i) (B j) = 0 := by all_goals simp at hi lemma Bi_sum_quad (i : Fin 11) (f : Fin 11 → ℚ) : - quadBiLin (B i) (∑ k, f k • B k) = f i * quadBiLin (B i) (B i) := by + quadBiLin (B i) (∑ k, f k • B k) = f i * quadBiLin (B i) (B i) := by rw [quadBiLin.map_sum₂] rw [Fintype.sum_eq_single i] rw [quadBiLin.map_smul₂] @@ -96,7 +96,7 @@ lemma quadCoeff_eq_bilinear (i : Fin 11) : quadCoeff i = quadBiLin (B i) (B i) : all_goals rfl lemma on_accQuad (f : Fin 11 → ℚ) : - accQuad (∑ i, f i • B i) = ∑ i, quadCoeff i * (f i)^2 := by + accQuad (∑ i, f i • B i) = ∑ i, quadCoeff i * (f i)^2 := by change quadBiLin _ _ = _ rw [quadBiLin.map_sum₁] apply Fintype.sum_congr @@ -105,7 +105,7 @@ lemma on_accQuad (f : Fin 11 → ℚ) : ring lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) - (k : Fin 11) : quadCoeff k * (f k)^2 = 0 := by + (k : Fin 11) : quadCoeff k * (f k)^2 = 0 := by obtain ⟨S, hS⟩ := hS have hQ := quadSol S.1 rw [hS, on_accQuad] at hQ @@ -169,7 +169,7 @@ lemma isSolution_sum_part (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ simp rfl -lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : +lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 10 = - 3 * f 9 := by have hx := isSolution_sum_part f hS obtain ⟨S, hS'⟩ := hS diff --git a/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSol.lean b/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSol.lean index e4965c1..04aa939 100644 --- a/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSol.lean +++ b/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSol.lean @@ -61,7 +61,7 @@ lemma accQuad_α₁_α₂_zero (S : (PlusU1 n).LinSols) (h1 : α₁ C S = 0) /-- The construction of a `QuadSol` from a `LinSols` in the generic case. -/ def genericToQuad (S : (PlusU1 n).LinSols) : (PlusU1 n).QuadSols := - linearToQuad ((α₁ C S) • S + α₂ S • C.1) (accQuad_α₁_α₂ C S) + linearToQuad ((α₁ C S) • S + α₂ S • C.1) (accQuad_α₁_α₂ C S) lemma genericToQuad_on_quad (S : (PlusU1 n).QuadSols) : genericToQuad C S.1 = (α₁ C S.1) • S := by @@ -80,7 +80,7 @@ def specialToQuad (S : (PlusU1 n).LinSols) (a b : ℚ) (h1 : α₁ C S = 0) (h2 : α₂ S = 0) : (PlusU1 n).QuadSols := linearToQuad (a • S + b • C.1) (accQuad_α₁_α₂_zero C S h1 h2 a b) -lemma special_on_quad (S : (PlusU1 n).QuadSols) (h1 : α₁ C S.1 = 0) : +lemma special_on_quad (S : (PlusU1 n).QuadSols) (h1 : α₁ C S.1 = 0) : specialToQuad C S.1 1 0 h1 (α₂_AFQ S) = S := by apply ACCSystemQuad.QuadSols.ext change (1 • S.val + 0 • C.val) = S.val @@ -123,9 +123,9 @@ lemma toQuadInv_special (S : (PlusU1 n).QuadSols) (h : α₁ C S.1 = 0) : rw [special_on_quad] lemma toQuadInv_generic (S : (PlusU1 n).QuadSols) (h : α₁ C S.1 ≠ 0) : - (toQuadInv C S).2.1 • genericToQuad C (toQuadInv C S).1 = S := by + (toQuadInv C S).2.1 • genericToQuad C (toQuadInv C S).1 = S := by simp only [toQuadInv_fst] - rw [show (toQuadInv C S).2.1 = (α₁ C S.1)⁻¹ by rw [toQuadInv, if_neg h]] + rw [show (toQuadInv C S).2.1 = (α₁ C S.1)⁻¹ by rw [toQuadInv, if_neg h]] rw [genericToQuad_neq_zero C S h] lemma toQuad_rightInverse : Function.RightInverse (@toQuadInv n C) (toQuad C) := by diff --git a/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSolToSol.lean b/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSolToSol.lean index 7488e75..1e6f96b 100644 --- a/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSolToSol.lean +++ b/HepLean/AnomalyCancellation/SMNu/PlusU1/QuadSolToSol.lean @@ -71,7 +71,7 @@ def special (S : (PlusU1 n).QuadSols) (a b : ℚ) (h1 : α₁ S = 0) (h2 : α₂ (PlusU1 n).Sols := quadToAF (BL.addQuad S a b) (cube_α₁_α₂_zero S a b h1 h2) -lemma special_on_AF (S : (PlusU1 n).Sols) (h1 : α₁ S.1 = 0) : +lemma special_on_AF (S : (PlusU1 n).Sols) (h1 : α₁ S.1 = 0) : special S.1 1 0 h1 (α₂_AF S) = S := by apply ACCSystem.Sols.ext change (BL.addQuad S.1 1 0).val = _ @@ -84,7 +84,7 @@ open QuadSolToSol /-- A map from `QuadSols × ℚ × ℚ` to `Sols` taking account of the special and generic cases. We will show that this map is a surjection. -/ def quadSolToSol {n : ℕ} : (PlusU1 n).QuadSols × ℚ × ℚ → (PlusU1 n).Sols := fun S => - if h1 : α₁ S.1 = 0 ∧ α₂ S.1 = 0 then + if h1 : α₁ S.1 = 0 ∧ α₂ S.1 = 0 then special S.1 S.2.1 S.2.2 h1.1 h1.2 else S.2.1 • generic S.1 @@ -98,7 +98,7 @@ def quadSolToSolInv {n : ℕ} : (PlusU1 n).Sols → (PlusU1 n).QuadSols × ℚ else (S.1, (α₁ S.1)⁻¹, 0) -lemma quadSolToSolInv_1 (S : (PlusU1 n).Sols) : +lemma quadSolToSolInv_1 (S : (PlusU1 n).Sols) : (quadSolToSolInv S).1 = S.1 := by simp [quadSolToSolInv] split @@ -110,24 +110,24 @@ lemma quadSolToSolInv_α₁_α₂_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 = 0) rw [quadSolToSolInv_1, α₂_AF S, h] simp -lemma quadSolToSolInv_α₁_α₂_neq_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) : - ¬ (α₁ (quadSolToSolInv S).1 = 0 ∧ α₂ (quadSolToSolInv S).1 = 0) := by +lemma quadSolToSolInv_α₁_α₂_neq_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) : + ¬ (α₁ (quadSolToSolInv S).1 = 0 ∧ α₂ (quadSolToSolInv S).1 = 0) := by rw [not_and, quadSolToSolInv_1, α₂_AF S] intro hn simp_all lemma quadSolToSolInv_special (S : (PlusU1 n).Sols) (h : α₁ S.1 = 0) : special (quadSolToSolInv S).1 (quadSolToSolInv S).2.1 (quadSolToSolInv S).2.2 - (quadSolToSolInv_α₁_α₂_zero S h).1 (quadSolToSolInv_α₁_α₂_zero S h).2 = S := by + (quadSolToSolInv_α₁_α₂_zero S h).1 (quadSolToSolInv_α₁_α₂_zero S h).2 = S := by simp [quadSolToSolInv_1] - rw [show (quadSolToSolInv S).2.1 = 1 by rw [quadSolToSolInv, if_pos h]] - rw [show (quadSolToSolInv S).2.2 = 0 by rw [quadSolToSolInv, if_pos h]] + rw [show (quadSolToSolInv S).2.1 = 1 by rw [quadSolToSolInv, if_pos h]] + rw [show (quadSolToSolInv S).2.2 = 0 by rw [quadSolToSolInv, if_pos h]] rw [special_on_AF] lemma quadSolToSolInv_generic (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) : - (quadSolToSolInv S).2.1 • generic (quadSolToSolInv S).1 = S := by + (quadSolToSolInv S).2.1 • generic (quadSolToSolInv S).1 = S := by simp [quadSolToSolInv_1] - rw [show (quadSolToSolInv S).2.1 = (α₁ S.1)⁻¹ by rw [quadSolToSolInv, if_neg h]] + rw [show (quadSolToSolInv S).2.1 = (α₁ S.1)⁻¹ by rw [quadSolToSolInv, if_neg h]] rw [generic_on_AF_α₁_ne_zero S h] lemma quadSolToSolInv_rightInverse : Function.RightInverse (@quadSolToSolInv n) quadSolToSol := by diff --git a/HepLean/BeyondTheStandardModel/TwoHDM/Basic.lean b/HepLean/BeyondTheStandardModel/TwoHDM/Basic.lean index f395d71..9c5c8fa 100644 --- a/HepLean/BeyondTheStandardModel/TwoHDM/Basic.lean +++ b/HepLean/BeyondTheStandardModel/TwoHDM/Basic.lean @@ -25,10 +25,10 @@ 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 + 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, Φ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 @@ -61,7 +61,7 @@ lemma swap_fields : /-- 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 + 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, diff --git a/HepLean/FeynmanDiagrams/Basic.lean b/HepLean/FeynmanDiagrams/Basic.lean index 9c41239..a5669ef 100644 --- a/HepLean/FeynmanDiagrams/Basic.lean +++ b/HepLean/FeynmanDiagrams/Basic.lean @@ -72,14 +72,14 @@ def toEdge {𝓔 𝓥 : Type} : Over (P.HalfEdgeLabel × 𝓔 × 𝓥) ⥤ Over /-- The functor from `Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` to `Over (P.HalfEdgeLabel)` induced by projections on products. -/ @[simps!] -def toHalfEdge {𝓔 𝓥 : Type} : Over (P.HalfEdgeLabel × 𝓔 × 𝓥) ⥤ Over P.HalfEdgeLabel := +def toHalfEdge {𝓔 𝓥 : Type} : Over (P.HalfEdgeLabel × 𝓔 × 𝓥) ⥤ Over P.HalfEdgeLabel := Over.map Prod.fst /-- The functor from `Over P.VertexLabel` to `Type` induced by pull-back along insertion of `v : P.VertexLabel`. -/ @[simps!] def preimageType' {𝓥 : Type} (v : 𝓥) : Over 𝓥 ⥤ Type where - obj := fun f => f.hom ⁻¹' {v} + obj := fun f => f.hom ⁻¹' {v} map {f g} F := fun x => ⟨F.left x.1, by have h := congrFun F.w x @@ -88,19 +88,19 @@ def preimageType' {𝓥 : Type} (v : 𝓥) : Over 𝓥 ⥤ Type where simpa [h] using x.2⟩ /-- The functor from `Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` to - `Over P.HalfEdgeLabel` induced by pull-back along insertion of `v : P.VertexLabel`. -/ -def preimageVertex {𝓔 𝓥 : Type} (v : 𝓥) : + `Over P.HalfEdgeLabel` induced by pull-back along insertion of `v : P.VertexLabel`. -/ +def preimageVertex {𝓔 𝓥 : Type} (v : 𝓥) : Over (P.HalfEdgeLabel × 𝓔 × 𝓥) ⥤ Over P.HalfEdgeLabel where - obj f := Over.mk (fun x => Prod.fst (f.hom x.1) : + obj f := Over.mk (fun x => Prod.fst (f.hom x.1) : (P.toVertex ⋙ preimageType' v).obj f ⟶ P.HalfEdgeLabel) map {f g} F := Over.homMk ((P.toVertex ⋙ preimageType' v).map F) (funext <| fun x => congrArg Prod.fst <| congrFun F.w x.1) /-- The functor from `Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` to - `Over P.HalfEdgeLabel` induced by pull-back along insertion of `v : P.EdgeLabel`. -/ + `Over P.HalfEdgeLabel` induced by pull-back along insertion of `v : P.EdgeLabel`. -/ def preimageEdge {𝓔 𝓥 : Type} (v : 𝓔) : - Over (P.HalfEdgeLabel × 𝓔 × 𝓥) ⥤ Over P.HalfEdgeLabel where - obj f := Over.mk (fun x => Prod.fst (f.hom x.1) : + Over (P.HalfEdgeLabel × 𝓔 × 𝓥) ⥤ Over P.HalfEdgeLabel where + obj f := Over.mk (fun x => Prod.fst (f.hom x.1) : (P.toEdge ⋙ preimageType' v).obj f ⟶ P.HalfEdgeLabel) map {f g} F := Over.homMk ((P.toEdge ⋙ preimageType' v).map F) (funext <| fun x => congrArg Prod.fst <| congrFun F.w x.1) @@ -117,7 +117,7 @@ properties. -/ -/-- A set of conditions on `PreFeynmanRule` for it to be considered finite. -/ +/-- A set of conditions on `PreFeynmanRule` for it to be considered finite. -/ class IsFinitePreFeynmanRule (P : PreFeynmanRule) where /-- The type of edge labels is decidable. -/ edgeLabelDecidable : DecidableEq P.EdgeLabel @@ -173,31 +173,31 @@ instance preimageEdgeDecidablePred {𝓔 𝓥 : Type} [DecidableEq 𝓔] (v : | isFalse h => isFalse h instance preimageVertexDecidable {𝓔 𝓥 : Type} (v : 𝓥) - (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] : + (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] : DecidableEq ((P.preimageVertex v).obj F).left := Subtype.instDecidableEq instance preimageEdgeDecidable {𝓔 𝓥 : Type} (v : 𝓔) - (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] : + (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] : DecidableEq ((P.preimageEdge v).obj F).left := Subtype.instDecidableEq -instance preimageVertexFintype {𝓔 𝓥 : Type} [DecidableEq 𝓥] - (v : 𝓥) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [h : Fintype F.left] : - Fintype ((P.preimageVertex v).obj F).left := @Subtype.fintype _ _ _ h +instance preimageVertexFintype {𝓔 𝓥 : Type} [DecidableEq 𝓥] + (v : 𝓥) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [h : Fintype F.left] : + Fintype ((P.preimageVertex v).obj F).left := @Subtype.fintype _ _ _ h instance preimageEdgeFintype {𝓔 𝓥 : Type} [DecidableEq 𝓔] - (v : 𝓔) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [h : Fintype F.left] : - Fintype ((P.preimageEdge v).obj F).left := @Subtype.fintype _ _ _ h + (v : 𝓔) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [h : Fintype F.left] : + Fintype ((P.preimageEdge v).obj F).left := @Subtype.fintype _ _ _ h -instance preimageVertexMapFintype [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} - [DecidableEq 𝓥] (v : 𝓥) (f : 𝓥 ⟶ P.VertexLabel) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) +instance preimageVertexMapFintype [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} + [DecidableEq 𝓥] (v : 𝓥) (f : 𝓥 ⟶ P.VertexLabel) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [Fintype F.left] : - Fintype ((P.vertexLabelMap (f v)).left → ((P.preimageVertex v).obj F).left) := + Fintype ((P.vertexLabelMap (f v)).left → ((P.preimageVertex v).obj F).left) := Pi.fintype -instance preimageEdgeMapFintype [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} - [DecidableEq 𝓔] (v : 𝓔) (f : 𝓔 ⟶ P.EdgeLabel) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) +instance preimageEdgeMapFintype [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} + [DecidableEq 𝓔] (v : 𝓔) (f : 𝓔 ⟶ P.EdgeLabel) (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [Fintype F.left] : - Fintype ((P.edgeLabelMap (f v)).left → ((P.preimageEdge v).obj F).left) := + Fintype ((P.edgeLabelMap (f v)).left → ((P.preimageEdge v).obj F).left) := Pi.fintype /-! @@ -214,11 +214,11 @@ We will show that for `IsFinitePreFeynmanRule` the condition of been external is def External {P : PreFeynmanRule} (v : P.VertexLabel) : Prop := IsIsomorphic (P.vertexLabelMap v).left (Fin 1) -lemma external_iff_exists_bijection {P : PreFeynmanRule} (v : P.VertexLabel) : +lemma external_iff_exists_bijection {P : PreFeynmanRule} (v : P.VertexLabel) : External v ↔ ∃ (κ : (P.vertexLabelMap v).left → Fin 1), Function.Bijective κ := by refine Iff.intro (fun h => ?_) (fun h => ?_) obtain ⟨κ, κm1, h1, h2⟩ := h - let f : (P.vertexLabelMap v).left ≅ (Fin 1) := ⟨κ, κm1, h1, h2⟩ + let f : (P.vertexLabelMap v).left ≅ (Fin 1) := ⟨κ, κm1, h1, h2⟩ exact ⟨f.hom, (isIso_iff_bijective f.hom).mp $ Iso.isIso_hom f⟩ obtain ⟨κ, h1⟩ := h let f : (P.vertexLabelMap v).left ⟶ (Fin 1) := κ @@ -248,7 +248,7 @@ def DiagramVertexProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × `F : Over (P.HalfEdgeLabel × P.EdgeLabel × P.VertexLabel)` for it to be a Feynman diagram. This condition corresponds to the edges of `F` having the correct half-edges associated with them. -/ -def DiagramEdgeProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) +def DiagramEdgeProp {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) (f : 𝓔 ⟶ P.EdgeLabel) := ∀ v, IsIsomorphic (P.edgeLabelMap (f v)) ((P.preimageEdge v).obj F) @@ -274,7 +274,7 @@ lemma diagramEdgeProp_iff {𝓔 𝓥 : Type} (F : Over (P.HalfEdgeLabel × 𝓔 ∧ ((P.preimageEdge v).obj F).hom ∘ κ = (P.edgeLabelMap (f v)).hom := by refine Iff.intro (fun h v => ?_) (fun h v => ?_) obtain ⟨κ, κm1, h1, h2⟩ := h v - let f := (Over.forget P.HalfEdgeLabel).mapIso ⟨κ, κm1, h1, h2⟩ + let f := (Over.forget P.HalfEdgeLabel).mapIso ⟨κ, κm1, h1, h2⟩ refine ⟨f.hom, (isIso_iff_bijective f.hom).mp $ Iso.isIso_hom f, κ.w⟩ obtain ⟨κ, h1, h2⟩ := h v let f : (P.edgeLabelMap (f v)) ⟶ ((P.preimageEdge v).obj F) := Over.homMk κ h2 @@ -289,17 +289,17 @@ instance diagramVertexPropDecidable @decidable_of_decidable_of_iff _ _ (@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ (fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _ - (fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _ ) _) + (fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _ ) _) (P.diagramVertexProp_iff F f).symm instance diagramEdgePropDecidable - [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} [Fintype 𝓔] [DecidableEq 𝓔] + [IsFinitePreFeynmanRule P] {𝓔 𝓥 : Type} [Fintype 𝓔] [DecidableEq 𝓔] (F : Over (P.HalfEdgeLabel × 𝓔 × 𝓥)) [DecidableEq F.left] [Fintype F.left] (f : 𝓔 ⟶ P.EdgeLabel) : Decidable (P.DiagramEdgeProp F f) := @decidable_of_decidable_of_iff _ _ (@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ (fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _ - (fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _ ) _) + (fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _ ) _) (P.diagramEdgeProp_iff F f).symm end PreFeynmanRule @@ -358,8 +358,8 @@ def 𝓱𝓔To𝓥 : Over F.𝓥 := P.toVertex.obj F.𝓱𝓔To𝓔𝓥 /-- The condition which must be satisfied by maps to form a Feynman diagram. -/ def Cond {𝓔 𝓥 𝓱𝓔 : Type} (𝓔𝓞 : 𝓔 → P.EdgeLabel) (𝓥𝓞 : 𝓥 → P.VertexLabel) - (𝓱𝓔To𝓔𝓥 : 𝓱𝓔 → P.HalfEdgeLabel × 𝓔 × 𝓥) : Prop := - P.DiagramEdgeProp (Over.mk 𝓱𝓔To𝓔𝓥) 𝓔𝓞 ∧ + (𝓱𝓔To𝓔𝓥 : 𝓱𝓔 → P.HalfEdgeLabel × 𝓔 × 𝓥) : Prop := + P.DiagramEdgeProp (Over.mk 𝓱𝓔To𝓔𝓥) 𝓔𝓞 ∧ P.DiagramVertexProp (Over.mk 𝓱𝓔To𝓔𝓥) 𝓥𝓞 lemma cond_self : Cond F.𝓔𝓞.hom F.𝓥𝓞.hom F.𝓱𝓔To𝓔𝓥.hom := @@ -441,7 +441,7 @@ instance {F : FeynmanDiagram P} [IsFiniteDiagram F] : DecidableEq F.𝓱𝓔 := IsFiniteDiagram.𝓱𝓔DecidableEq instance {F : FeynmanDiagram P} [IsFiniteDiagram F] (i : F.𝓱𝓔) (j : F.𝓔) : - Decidable (F.𝓱𝓔To𝓔.hom i = j) := IsFiniteDiagram.𝓔DecidableEq (F.𝓱𝓔To𝓔.hom i) j + Decidable (F.𝓱𝓔To𝓔.hom i = j) := IsFiniteDiagram.𝓔DecidableEq (F.𝓱𝓔To𝓔.hom i) j instance fintypeProdHalfEdgeLabel𝓔𝓥 {F : FeynmanDiagram P} [IsFinitePreFeynmanRule P] [IsFiniteDiagram F] : DecidableEq (P.HalfEdgeLabel × F.𝓔 × F.𝓥) := @@ -457,15 +457,15 @@ This will be used to define the category of Feynman diagrams. -/ /-- Given two maps `F.𝓔 ⟶ G.𝓔` and `F.𝓥 ⟶ G.𝓥` the corresponding map - `P.HalfEdgeLabel × F.𝓔 × F.𝓥 → P.HalfEdgeLabel × G.𝓔 × G.𝓥`. -/ + `P.HalfEdgeLabel × F.𝓔 × F.𝓥 → P.HalfEdgeLabel × G.𝓔 × G.𝓥`. -/ @[simps!] def edgeVertexMap {F G : FeynmanDiagram P} (𝓔 : F.𝓔 ⟶ G.𝓔) (𝓥 : F.𝓥 ⟶ G.𝓥) : - P.HalfEdgeLabel × F.𝓔 × F.𝓥 → P.HalfEdgeLabel × G.𝓔 × G.𝓥 := + P.HalfEdgeLabel × F.𝓔 × F.𝓥 → P.HalfEdgeLabel × G.𝓔 × G.𝓥 := fun x => ⟨x.1, 𝓔 x.2.1, 𝓥 x.2.2⟩ /-- Given equivalences `F.𝓔 ≃ G.𝓔` and `F.𝓥 ≃ G.𝓥`, the induced equivalence between `P.HalfEdgeLabel × F.𝓔 × F.𝓥` and `P.HalfEdgeLabel × G.𝓔 × G.𝓥 ` -/ -def edgeVertexEquiv {F G : FeynmanDiagram P} (𝓔 : F.𝓔 ≃ G.𝓔) (𝓥 : F.𝓥 ≃ G.𝓥) : +def edgeVertexEquiv {F G : FeynmanDiagram P} (𝓔 : F.𝓔 ≃ G.𝓔) (𝓥 : F.𝓥 ≃ G.𝓥) : P.HalfEdgeLabel × F.𝓔 × F.𝓥 ≃ P.HalfEdgeLabel × G.𝓔 × G.𝓥 where toFun := edgeVertexMap 𝓔.toFun 𝓥.toFun invFun := edgeVertexMap 𝓔.invFun 𝓥.invFun @@ -539,7 +539,7 @@ lemma ext {F G : FeynmanDiagram P} {f g : Hom F G} (h𝓔 : f.𝓔 = g.𝓔) def Cond {F G : FeynmanDiagram P} (𝓔 : F.𝓔 → G.𝓔) (𝓥 : F.𝓥 → G.𝓥) (𝓱𝓔 : F.𝓱𝓔 → G.𝓱𝓔) : Prop := (∀ x, G.𝓔𝓞.hom (𝓔 x) = F.𝓔𝓞.hom x) ∧ (∀ x, G.𝓥𝓞.hom (𝓥 x) = F.𝓥𝓞.hom x) ∧ - (∀ x, G.𝓱𝓔To𝓔𝓥.hom (𝓱𝓔 x) = edgeVertexMap 𝓔 𝓥 (F.𝓱𝓔To𝓔𝓥.hom x)) + (∀ x, G.𝓱𝓔To𝓔𝓥.hom (𝓱𝓔 x) = edgeVertexMap 𝓔 𝓥 (F.𝓱𝓔To𝓔𝓥.hom x)) lemma cond_satisfied {F G : FeynmanDiagram P} (f : Hom F G) : Cond f.𝓔 f.𝓥 f.𝓱𝓔 := @@ -564,9 +564,9 @@ instance {F G : FeynmanDiagram P} [IsFiniteDiagram F] [IsFinitePreFeynmanRule P] @Fintype.decidableForallFintype _ _ (fun _ => preFeynmanRuleDecEq𝓥 P _ _) _ instance {F G : FeynmanDiagram P} [IsFiniteDiagram F] [IsFiniteDiagram G] [IsFinitePreFeynmanRule P] - (𝓔 : F.𝓔 → G.𝓔) (𝓥 : F.𝓥 → G.𝓥) (𝓱𝓔 : F.𝓱𝓔 → G.𝓱𝓔) : - Decidable (∀ x, G.𝓱𝓔To𝓔𝓥.hom (𝓱𝓔 x) = edgeVertexMap 𝓔 𝓥 (F.𝓱𝓔To𝓔𝓥.hom x)) := - @Fintype.decidableForallFintype _ _ (fun _ => fintypeProdHalfEdgeLabel𝓔𝓥 _ _) _ + (𝓔 : F.𝓔 → G.𝓔) (𝓥 : F.𝓥 → G.𝓥) (𝓱𝓔 : F.𝓱𝓔 → G.𝓱𝓔) : + Decidable (∀ x, G.𝓱𝓔To𝓔𝓥.hom (𝓱𝓔 x) = edgeVertexMap 𝓔 𝓥 (F.𝓱𝓔To𝓔𝓥.hom x)) := + @Fintype.decidableForallFintype _ _ (fun _ => fintypeProdHalfEdgeLabel𝓔𝓥 _ _) _ instance {F G : FeynmanDiagram P} [IsFiniteDiagram F] [IsFiniteDiagram G] [IsFinitePreFeynmanRule P] (𝓔 : F.𝓔 → G.𝓔) (𝓥 : F.𝓥 → G.𝓥) (𝓱𝓔 : F.𝓱𝓔 → G.𝓱𝓔) : Decidable (Cond 𝓔 𝓥 𝓱𝓔) := @@ -678,7 +678,7 @@ instance [IsFinitePreFeynmanRule P] [IsFiniteDiagram F] : Fintype F.SymmetryType @[simp] def cardSymmetryFactor : Cardinal := Cardinal.mk (F.SymmetryType) -/-- The symmetry factor of a Finite Feynman diagram, as a natural number. -/ +/-- The symmetry factor of a Finite Feynman diagram, as a natural number. -/ @[simp] def symmetryFactor [IsFinitePreFeynmanRule P] [IsFiniteDiagram F] : ℕ := (Fintype.card F.SymmetryType) @@ -714,7 +714,7 @@ def AdjRelation : F.𝓥 → F.𝓥 → Prop := fun x y => instance [IsFiniteDiagram F] : DecidableRel F.AdjRelation := fun _ _ => @And.decidable _ _ _ $ - @Fintype.decidableExistsFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ ( + @Fintype.decidableExistsFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ ( fun _ => @And.decidable _ _ (instDecidableEq𝓔OfIsFiniteDiagram _ _) $ @And.decidable _ _ (instDecidableEq𝓥OfIsFiniteDiagram _ _) (instDecidableEq𝓥OfIsFiniteDiagram _ _)) _ ) _ diff --git a/HepLean/FeynmanDiagrams/Instances/ComplexScalar.lean b/HepLean/FeynmanDiagrams/Instances/ComplexScalar.lean index 6c081c2..fb9276b 100644 --- a/HepLean/FeynmanDiagrams/Instances/ComplexScalar.lean +++ b/HepLean/FeynmanDiagrams/Instances/ComplexScalar.lean @@ -42,22 +42,22 @@ instance (a : ℕ) : OfNat complexScalarFeynmanRules.VertexLabel a where ofNat := (a : Fin _) instance : IsFinitePreFeynmanRule complexScalarFeynmanRules where - edgeLabelDecidable := instDecidableEqFin _ + edgeLabelDecidable := instDecidableEqFin _ vertexLabelDecidable := instDecidableEqFin _ halfEdgeLabelDecidable := instDecidableEqFin _ vertexMapFintype := fun v => match v with | 0 => Fin.fintype _ - | 1 => Fin.fintype _ - | 2 => Fin.fintype _ + | 1 => Fin.fintype _ + | 2 => Fin.fintype _ edgeMapFintype := fun v => match v with | 0 => Fin.fintype _ vertexMapDecidable := fun v => match v with | 0 => instDecidableEqFin _ - | 1 => instDecidableEqFin _ - | 2 => instDecidableEqFin _ + | 1 => instDecidableEqFin _ + | 2 => instDecidableEqFin _ edgeMapDecidable := fun v => match v with | 0 => instDecidableEqFin _ diff --git a/HepLean/FeynmanDiagrams/Instances/Phi4.lean b/HepLean/FeynmanDiagrams/Instances/Phi4.lean index 0c780ad..2365063 100644 --- a/HepLean/FeynmanDiagrams/Instances/Phi4.lean +++ b/HepLean/FeynmanDiagrams/Instances/Phi4.lean @@ -44,20 +44,20 @@ instance (a : ℕ) : OfNat phi4PreFeynmanRules.VertexLabel a where ofNat := (a : Fin _) instance : IsFinitePreFeynmanRule phi4PreFeynmanRules where - edgeLabelDecidable := instDecidableEqFin _ + edgeLabelDecidable := instDecidableEqFin _ vertexLabelDecidable := instDecidableEqFin _ halfEdgeLabelDecidable := instDecidableEqFin _ vertexMapFintype := fun v => match v with | 0 => Fin.fintype _ - | 1 => Fin.fintype _ + | 1 => Fin.fintype _ edgeMapFintype := fun v => match v with | 0 => Fin.fintype _ vertexMapDecidable := fun v => match v with | 0 => instDecidableEqFin _ - | 1 => instDecidableEqFin _ + | 1 => instDecidableEqFin _ edgeMapDecidable := fun v => match v with | 0 => instDecidableEqFin _ diff --git a/HepLean/FeynmanDiagrams/Momentum.lean b/HepLean/FeynmanDiagrams/Momentum.lean index 51c58b0..27753d2 100644 --- a/HepLean/FeynmanDiagrams/Momentum.lean +++ b/HepLean/FeynmanDiagrams/Momentum.lean @@ -45,8 +45,8 @@ We define the direct sum of the edge and vertex momentum spaces. -/ /-- The type which assocaites to each half-edge a `1`-dimensional vector space. - Corresponding to that spanned by its momentum. -/ -def HalfEdgeMomenta : Type := F.𝓱𝓔 → ℝ + Corresponding to that spanned by its momentum. -/ +def HalfEdgeMomenta : Type := F.𝓱𝓔 → ℝ instance : AddCommGroup F.HalfEdgeMomenta := Pi.addCommGroup @@ -54,14 +54,14 @@ instance : Module ℝ F.HalfEdgeMomenta := Pi.module _ _ _ /-- An auxillary function used to define the Euclidean inner product on `F.HalfEdgeMomenta`. -/ def euclidInnerAux (x : F.HalfEdgeMomenta) : F.HalfEdgeMomenta →ₗ[ℝ] ℝ where - toFun y := ∑ i, (x i) * (y i) + toFun y := ∑ i, (x i) * (y i) map_add' z y := show (∑ i, (x i) * (z i + y i)) = (∑ i, x i * z i) + ∑ i, x i * (y i) by simp only [mul_add, Finset.sum_add_distrib] map_smul' c y := show (∑ i, x i * (c * y i)) = c * ∑ i, x i * y i by rw [Finset.mul_sum] - refine Finset.sum_congr rfl (fun _ _ => by ring) + refine Finset.sum_congr rfl (fun _ _ => by ring) lemma euclidInnerAux_symm (x y : F.HalfEdgeMomenta) : F.euclidInnerAux x y = F.euclidInnerAux y x := Finset.sum_congr rfl (fun _ _ => by ring) @@ -70,15 +70,15 @@ lemma euclidInnerAux_symm (x y : F.HalfEdgeMomenta) : def euclidInner : F.HalfEdgeMomenta →ₗ[ℝ] F.HalfEdgeMomenta →ₗ[ℝ] ℝ where toFun x := F.euclidInnerAux x map_add' x y := by - refine LinearMap.ext (fun z => ?_) + refine LinearMap.ext (fun z => ?_) simp only [euclidInnerAux_symm, map_add, LinearMap.add_apply] map_smul' c x := by - refine LinearMap.ext (fun z => ?_) + refine LinearMap.ext (fun z => ?_) simp only [euclidInnerAux_symm, LinearMapClass.map_smul, smul_eq_mul, RingHom.id_apply, LinearMap.smul_apply] /-- The type which assocaites to each ege a `1`-dimensional vector space. - Corresponding to that spanned by its total outflowing momentum. -/ + Corresponding to that spanned by its total outflowing momentum. -/ def EdgeMomenta : Type := F.𝓔 → ℝ instance : AddCommGroup F.EdgeMomenta := Pi.addCommGroup @@ -86,7 +86,7 @@ instance : AddCommGroup F.EdgeMomenta := Pi.addCommGroup instance : Module ℝ F.EdgeMomenta := Pi.module _ _ _ /-- The type which assocaites to each ege a `1`-dimensional vector space. - Corresponding to that spanned by its total inflowing momentum. -/ + Corresponding to that spanned by its total inflowing momentum. -/ def VertexMomenta : Type := F.𝓥 → ℝ instance : AddCommGroup F.VertexMomenta := Pi.addCommGroup @@ -96,8 +96,8 @@ instance : Module ℝ F.VertexMomenta := Pi.module _ _ _ /-- The map from `Fin 2` to `Type` landing on `EdgeMomenta` and `VertexMomenta`. -/ def EdgeVertexMomentaMap : Fin 2 → Type := fun i => match i with - | 0 => F.EdgeMomenta - | 1 => F.VertexMomenta + | 0 => F.EdgeMomenta + | 1 => F.VertexMomenta instance (i : Fin 2) : AddCommGroup (EdgeVertexMomentaMap F i) := match i with diff --git a/HepLean/FlavorPhysics/CKMMatrix/Basic.lean b/HepLean/FlavorPhysics/CKMMatrix/Basic.lean index ddcf854..991c8fa 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/Basic.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/Basic.lean @@ -27,7 +27,7 @@ noncomputable section leading diagonal. -/ @[simp] def phaseShiftMatrix (a b c : ℝ) : Matrix (Fin 3) (Fin 3) ℂ := - ![![cexp (I * a), 0, 0], ![0, cexp (I * b), 0], ![0, 0, cexp (I * c)]] + ![![cexp (I * a), 0, 0], ![0, cexp (I * b), 0], ![0, 0, cexp (I * c)]] lemma phaseShiftMatrix_one : phaseShiftMatrix 0 0 0 = 1 := by ext i j @@ -49,7 +49,7 @@ lemma phaseShiftMatrix_mul (a b c d e f : ℝ) : fin_cases i <;> fin_cases j <;> simp [Matrix.mul_apply, Fin.sum_univ_three] any_goals rw [mul_add, exp_add] - change cexp (I * ↑c) * 0 = 0 + change cexp (I * ↑c) * 0 = 0 simp /-- Given three real numbers `a b c` the unitary matrix with `exp (I * a)` etc on the @@ -161,7 +161,7 @@ lemma equiv (V : CKMMatrix) (a b c d e f : ℝ) : rfl lemma ud (V : CKMMatrix) (a b c d e f : ℝ) : - (phaseShiftApply V a b c d e f).1 0 0 = cexp (a * I + d * I) * V.1 0 0 := by + (phaseShiftApply V a b c d e f).1 0 0 = cexp (a * I + d * I) * V.1 0 0 := by simp only [Fin.isValue, phaseShiftApply_coe] rw [mul_apply, Fin.sum_univ_three] rw [mul_apply, mul_apply, mul_apply, Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three] @@ -173,7 +173,7 @@ lemma ud (V : CKMMatrix) (a b c d e f : ℝ) : ring_nf lemma us (V : CKMMatrix) (a b c d e f : ℝ) : - (phaseShiftApply V a b c d e f).1 0 1 = cexp (a * I + e * I) * V.1 0 1 := by + (phaseShiftApply V a b c d e f).1 0 1 = cexp (a * I + e * I) * V.1 0 1 := by simp only [Fin.isValue, phaseShiftApply_coe] rw [mul_apply, Fin.sum_univ_three] rw [mul_apply, mul_apply, mul_apply, Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three] @@ -184,7 +184,7 @@ lemma us (V : CKMMatrix) (a b c d e f : ℝ) : ring_nf lemma ub (V : CKMMatrix) (a b c d e f : ℝ) : - (phaseShiftApply V a b c d e f).1 0 2 = cexp (a * I + f * I) * V.1 0 2 := by + (phaseShiftApply V a b c d e f).1 0 2 = cexp (a * I + f * I) * V.1 0 2 := by simp only [Fin.isValue, phaseShiftApply_coe] rw [mul_apply, Fin.sum_univ_three] rw [mul_apply, mul_apply, mul_apply, Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three] @@ -236,7 +236,7 @@ lemma td (V : CKMMatrix) (a b c d e f : ℝ) : simp only [Fin.isValue, cons_val', cons_val_zero, empty_val', cons_val_fin_one, vecCons_const, cons_val_two, tail_val', head_val', cons_val_one, head_cons, tail_cons, head_fin_const, zero_mul, add_zero, mul_zero] - change (0 * _ + _ ) * _ + (0 * _ + _ ) * 0 = _ + change (0 * _ + _ ) * _ + (0 * _ + _ ) * 0 = _ simp only [Fin.isValue, zero_mul, zero_add, mul_zero, add_zero] rw [exp_add] ring_nf @@ -272,7 +272,7 @@ end phaseShiftApply def VAbs' (V : unitaryGroup (Fin 3) ℂ) (i j : Fin 3) : ℝ := Complex.abs (V i j) lemma VAbs'_equiv (i j : Fin 3) (V U : CKMMatrix) (h : V ≈ U) : - VAbs' V i j = VAbs' U i j := by + VAbs' V i j = VAbs' U i j := by simp only [VAbs'] obtain ⟨a, b, c, e, f, g, h⟩ := h rw [h] @@ -288,7 +288,7 @@ lemma VAbs'_equiv (i j : Fin 3) (V U : CKMMatrix) (h : V ≈ U) : all_goals change Complex.abs (0 * _ + _) = _ all_goals simp [Complex.abs_exp] -/-- The absolute value of the `(i,j)`th any representative of `⟦V⟧`. -/ +/-- The absolute value of the `(i,j)`th any representative of `⟦V⟧`. -/ def VAbs (i j : Fin 3) : Quotient CKMMatrixSetoid → ℝ := Quotient.lift (fun V => VAbs' V i j) (VAbs'_equiv i j) diff --git a/HepLean/FlavorPhysics/CKMMatrix/Invariants.lean b/HepLean/FlavorPhysics/CKMMatrix/Invariants.lean index 9b9b024..2eb9f47 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/Invariants.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/Invariants.lean @@ -28,10 +28,10 @@ namespace Invariant def jarlskogℂCKM (V : CKMMatrix) : ℂ := [V]us * [V]cb * conj [V]ub * conj [V]cs -lemma jarlskogℂCKM_equiv (V U : CKMMatrix) (h : V ≈ U) : +lemma jarlskogℂCKM_equiv (V U : CKMMatrix) (h : V ≈ U) : jarlskogℂCKM V = jarlskogℂCKM U := by obtain ⟨a, b, c, e, f, g, h⟩ := h - change V = phaseShiftApply U a b c e f g at h + change V = phaseShiftApply U a b c e f g at h rw [h] simp only [jarlskogℂCKM, Fin.isValue, phaseShiftApply.ub, phaseShiftApply.us, phaseShiftApply.cb, phaseShiftApply.cs] @@ -51,8 +51,8 @@ def jarlskogℂ : Quotient CKMMatrixSetoid → ℂ := /-- An invariant for CKM mtrices corresponding to the square of the absolute values of the `us`, `ub` and `cb` elements multipled together divied by `(VudAbs V ^ 2 + VusAbs V ^2)` . -/ -def VusVubVcdSq (V : Quotient CKMMatrixSetoid) : ℝ := - VusAbs V ^ 2 * VubAbs V ^ 2 * VcbAbs V ^2 / (VudAbs V ^ 2 + VusAbs V ^2) +def VusVubVcdSq (V : Quotient CKMMatrixSetoid) : ℝ := + VusAbs V ^ 2 * VubAbs V ^ 2 * VcbAbs V ^2 / (VudAbs V ^ 2 + VusAbs V ^2) /-- An invariant for CKM matrices. The argument of this invariant is `δ₁₃` in the standard parameterization. -/ diff --git a/HepLean/FlavorPhysics/CKMMatrix/PhaseFreedom.lean b/HepLean/FlavorPhysics/CKMMatrix/PhaseFreedom.lean index 29c9bbf..ac7d93b 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/PhaseFreedom.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/PhaseFreedom.lean @@ -126,7 +126,7 @@ lemma shift_cross_product_phase_zero {V : CKMMatrix} {τ : ℝ} have hτ0 := congrFun hτ 0 simp [tRow] at hτ0 rw [← hτ0] - rw [← mul_assoc, ← exp_add, h1] + rw [← mul_assoc, ← exp_add, h1] congr 2 simp only [ofReal_sub, ofReal_neg] ring @@ -166,15 +166,15 @@ def FstRowThdColRealCond (U : CKMMatrix) : Prop := [U]ud = VudAbs ⟦U⟧ ∧ [U the cross product of the conjugates of the `u`th and `c`th rows, and the `cd`th and `cs`th elements are real and related in a set way. -/ -def ubOnePhaseCond (U : CKMMatrix) : Prop := +def ubOnePhaseCond (U : CKMMatrix) : Prop := [U]ud = 0 ∧ [U]us = 0 ∧ [U]cb = 0 ∧ [U]ub = 1 ∧ [U]t = conj [U]u ×₃ conj [U]c ∧ [U]cd = - VcdAbs ⟦U⟧ ∧ [U]cs = √(1 - VcdAbs ⟦U⟧ ^ 2) lemma fstRowThdColRealCond_shift_solution {V : CKMMatrix} (h1 : a + d = - arg [V]ud) - (h2 : a + e = - arg [V]us) (h3 : b + f = - arg [V]cb) + (h2 : a + e = - arg [V]us) (h3 : b + f = - arg [V]cb) (h4 : c + f = - arg [V]tb) (h5 : τ = - a - b - c - d - e - f) : - b = - τ + arg [V]ud + arg [V]us + arg [V]tb + a ∧ - c = - τ + arg [V]cb + arg [V]ud + arg [V]us + a ∧ + b = - τ + arg [V]ud + arg [V]us + arg [V]tb + a ∧ + c = - τ + arg [V]cb + arg [V]ud + arg [V]us + a ∧ d = - arg [V]ud - a ∧ e = - arg [V]us - a ∧ f = τ - arg [V]ud - arg [V]us - arg [V]cb - arg [V]tb - a := by @@ -202,11 +202,11 @@ lemma fstRowThdColRealCond_shift_solution {V : CKMMatrix} (h1 : a + d = - arg [V lemma ubOnePhaseCond_shift_solution {V : CKMMatrix} (h1 : a + f = - arg [V]ub) (h2 : 0 = - a - b - c - d - e - f) - (h3 : b + d = Real.pi - arg [V]cd) (h5 : b + e = - arg [V]cs) : - c = - Real.pi + arg [V]cd + arg [V]cs + arg [V]ub + b ∧ - d = Real.pi - arg [V]cd - b ∧ - e = - arg [V]cs - b ∧ - f = - arg [V]ub - a := by + (h3 : b + d = Real.pi - arg [V]cd) (h5 : b + e = - arg [V]cs) : + c = - Real.pi + arg [V]cd + arg [V]cs + arg [V]ub + b ∧ + d = Real.pi - arg [V]cd - b ∧ + e = - arg [V]cs - b ∧ + f = - arg [V]ub - a := by have hf : f = - arg [V]ub - a := by linear_combination h1 subst hf @@ -227,21 +227,21 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) : obtain ⟨τ, hτ⟩ := V.uRow_cross_cRow_eq_tRow let U : CKMMatrix := phaseShiftApply V 0 - (- τ + arg [V]ud + arg [V]us + arg [V]tb ) - (- τ + arg [V]cb + arg [V]ud + arg [V]us ) + (- τ + arg [V]ud + arg [V]us + arg [V]tb ) + (- τ + arg [V]cb + arg [V]ud + arg [V]us ) (- arg [V]ud ) (- arg [V]us) (τ - arg [V]ud - arg [V]us - arg [V]cb - arg [V]tb) have hUV : Quotient.mk CKMMatrixSetoid U = ⟦V⟧ := by simp only [Quotient.eq] symm - exact phaseShiftApply.equiv _ _ _ _ _ _ _ + exact phaseShiftApply.equiv _ _ _ _ _ _ _ use U apply And.intro exact phaseShiftApply.equiv _ _ _ _ _ _ _ apply And.intro rw [hUV] - apply shift_ud_phase_zero _ _ _ _ _ _ _ + apply shift_ud_phase_zero _ _ _ _ _ _ _ ring apply And.intro rw [hUV] @@ -259,15 +259,15 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) : ring lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud ≠ 0 ∨ [V]us ≠ 0)) - (hV : FstRowThdColRealCond V) : + (hV : FstRowThdColRealCond V) : ∃ (U : CKMMatrix), V ≈ U ∧ ubOnePhaseCond U:= by let U : CKMMatrix := phaseShiftApply V 0 0 (- Real.pi + arg [V]cd + arg [V]cs + arg [V]ub) - (Real.pi - arg [V]cd ) (- arg [V]cs) (- arg [V]ub ) + (Real.pi - arg [V]cd ) (- arg [V]cs) (- arg [V]ub ) use U have hUV : Quotient.mk CKMMatrixSetoid U= ⟦V⟧ := by simp only [Quotient.eq] symm - exact phaseShiftApply.equiv _ _ _ _ _ _ _ + exact phaseShiftApply.equiv _ _ _ _ _ _ _ apply And.intro exact phaseShiftApply.equiv _ _ _ _ _ _ _ apply And.intro @@ -294,7 +294,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud exact h1 apply And.intro · have hU1 : [U]ub = VubAbs ⟦V⟧ := by - apply shift_ub_phase_zero _ _ _ _ _ _ _ + apply shift_ub_phase_zero _ _ _ _ _ _ _ ring rw [hU1] have h1:= (ud_us_neq_zero_iff_ub_neq_one V).mpr.mt hb @@ -307,7 +307,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud ring apply And.intro · rw [hUV] - apply shift_cd_phase_pi _ _ _ _ _ _ _ + apply shift_cd_phase_pi _ _ _ _ _ _ _ ring have hcs : [U]cs = VcsAbs ⟦U⟧ := by rw [hUV] @@ -338,7 +338,7 @@ lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ lemma cs_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) (hV : FstRowThdColRealCond V) : [V]cs = (VtbAbs ⟦V⟧ * VudAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) - + (- VubAbs ⟦V⟧ * VusAbs ⟦V⟧ * VcbAbs ⟦V⟧/ (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) + + (- VubAbs ⟦V⟧ * VusAbs ⟦V⟧ * VcbAbs ⟦V⟧/ (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) * cexp (- arg [V]ub * I) := by have hτ : [V]t = cexp ((0 : ℝ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by simp only [ofReal_zero, zero_mul, exp_zero, one_smul] diff --git a/HepLean/FlavorPhysics/CKMMatrix/Relations.lean b/HepLean/FlavorPhysics/CKMMatrix/Relations.lean index bbf9dff..5602103 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/Relations.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/Relations.lean @@ -64,7 +64,7 @@ lemma normSq_Vud_plus_normSq_Vus (V : CKMMatrix) : normSq [V]ud + normSq [V]us = 1 - normSq [V]ub := by linear_combination (fst_row_normalized_normSq V) -lemma VudAbs_sq_add_VusAbs_sq : VudAbs V ^ 2 + VusAbs V ^2 = 1 - VubAbs V ^2 := by +lemma VudAbs_sq_add_VusAbs_sq : VudAbs V ^ 2 + VusAbs V ^2 = 1 - VubAbs V ^2 := by linear_combination (VAbs_sum_sq_row_eq_one V 0) lemma ud_us_neq_zero_iff_ub_neq_one (V : CKMMatrix) : @@ -104,7 +104,7 @@ lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℝ {V : CKMMatrix} (hb : [V]ud ≠ 0 exact h2 h3 lemma VAbsub_neq_zero_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid} - (hV : VAbs 0 2 V ≠ 1) :(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by + (hV : VAbs 0 2 V ≠ 1) :(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by obtain ⟨V⟩ := V change VubAbs ⟦V⟧ ≠ 1 at hV simp only [VubAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk] at hV @@ -112,7 +112,7 @@ lemma VAbsub_neq_zero_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid} simpa [← Complex.sq_abs] using (normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hV) lemma VAbsub_neq_zero_sqrt_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid} - (hV : VAbs 0 2 V ≠ 1) : √(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by + (hV : VAbs 0 2 V ≠ 1) : √(VudAbs V ^ 2 + VusAbs V ^ 2) ≠ 0 := by obtain ⟨V⟩ := V rw [Real.sqrt_ne_zero (Left.add_nonneg (sq_nonneg _) (sq_nonneg _))] change VubAbs ⟦V⟧ ≠ 1 at hV @@ -120,7 +120,7 @@ lemma VAbsub_neq_zero_sqrt_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid} rw [← ud_us_neq_zero_iff_ub_neq_one V] at hV simpa [← Complex.sq_abs] using (normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hV) -lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℂ {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) : +lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℂ {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) : (normSq [V]ud : ℂ) + normSq [V]us ≠ 0 := by have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hb simp at h1 @@ -178,9 +178,9 @@ lemma VAbs_thd_eq_one_snd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3} (hV section crossProduct -lemma conj_Vtb_cross_product {V : CKMMatrix} {τ : ℝ} +lemma conj_Vtb_cross_product {V : CKMMatrix} {τ : ℝ} (hτ : [V]t = cexp (τ * I) • (conj [V]u ×₃ conj [V]c)) : - conj [V]tb = cexp (- τ * I) * ([V]cs * [V]ud - [V]us * [V]cd) := by + conj [V]tb = cexp (- τ * I) * ([V]cs * [V]ud - [V]us * [V]cd) := by have h1 := congrFun hτ 2 simp [crossProduct, tRow, uRow, cRow] at h1 apply congrArg conj at h1 @@ -201,7 +201,7 @@ lemma conj_Vtb_mul_Vud {V : CKMMatrix} {τ : ℝ} simp [exp_neg] have h1 := exp_ne_zero (τ * I) field_simp - have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]ud = [V]cs + have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]ud = [V]cs * [V]ud * conj [V]ud - [V]us * ([V]cd * conj [V]ud) := by ring rw [h2, V.Vcd_mul_conj_Vud] @@ -217,7 +217,7 @@ lemma conj_Vtb_mul_Vus {V : CKMMatrix} {τ : ℝ} simp [exp_neg] have h1 := exp_ne_zero (τ * I) field_simp - have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]us = ([V]cs + have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]us = ([V]cs * conj [V]us) * [V]ud - [V]us * [V]cd * conj [V]us := by ring rw [h2, V.Vcs_mul_conj_Vus] @@ -225,18 +225,18 @@ lemma conj_Vtb_mul_Vus {V : CKMMatrix} {τ : ℝ} simp only [Fin.isValue, neg_mul] ring -lemma cs_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0) +lemma cs_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0) {τ : ℝ} (hτ : [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c))) : - [V]cs = (- conj [V]ub * [V]us * [V]cb + - cexp (τ * I) * conj [V]tb * conj [V]ud) / (normSq [V]ud + normSq [V]us) := by + [V]cs = (- conj [V]ub * [V]us * [V]cb + + cexp (τ * I) * conj [V]tb * conj [V]ud) / (normSq [V]ud + normSq [V]us) := by have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℂ h rw [conj_Vtb_mul_Vud hτ] field_simp ring -lemma cd_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0) +lemma cd_of_ud_us_ub_cb_tb {V : CKMMatrix} (h : [V]ud ≠ 0 ∨ [V]us ≠ 0) {τ : ℝ} (hτ : [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c))) : - [V]cd = - (conj [V]ub * [V]ud * [V]cb + cexp (τ * I) * conj [V]tb * conj [V]us) / + [V]cd = - (conj [V]ub * [V]ud * [V]cb + cexp (τ * I) * conj [V]tb * conj [V]us) / (normSq [V]ud + normSq [V]us) := by have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℂ h rw [conj_Vtb_mul_Vus hτ] @@ -247,7 +247,7 @@ end rows section individual -lemma VAbs_ge_zero (i j : Fin 3) (V : Quotient CKMMatrixSetoid) : 0 ≤ VAbs i j V := by +lemma VAbs_ge_zero (i j : Fin 3) (V : Quotient CKMMatrixSetoid) : 0 ≤ VAbs i j V := by obtain ⟨V, hV⟩ := Quot.exists_rep V rw [← hV] exact Complex.abs.nonneg _ @@ -291,7 +291,7 @@ lemma thd_col_normalized_normSq (V : CKMMatrix) : repeat rw [Complex.sq_abs] at h1 exact h1 -lemma cb_eq_zero_of_ud_us_zero {V : CKMMatrix} (h : [V]ud = 0 ∧ [V]us = 0) : +lemma cb_eq_zero_of_ud_us_zero {V : CKMMatrix} (h : [V]ud = 0 ∧ [V]us = 0) : [V]cb = 0 := by have h1 := fst_row_normalized_abs V rw [← thd_col_normalized_abs V] at h1 diff --git a/HepLean/FlavorPhysics/CKMMatrix/Rows.lean b/HepLean/FlavorPhysics/CKMMatrix/Rows.lean index 39e7c15..b48d0b2 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/Rows.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/Rows.lean @@ -130,13 +130,13 @@ lemma tRow_cRow_orthog (V : CKMMatrix) : conj [V]t ⬝ᵥ [V]c = 0 := by rw [mul_comm (V.1 _ 0) _, mul_comm (V.1 _ 1) _, mul_comm (V.1 _ 2) _] at ht exact ht -lemma uRow_cross_cRow_conj (V : CKMMatrix) : conj (conj [V]u ×₃ conj [V]c) = [V]u ×₃ [V]c := by +lemma uRow_cross_cRow_conj (V : CKMMatrix) : conj (conj [V]u ×₃ conj [V]c) = [V]u ×₃ [V]c := by simp only [crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, LinearMap.mk₂_apply, Pi.conj_apply] funext i fin_cases i <;> simp -lemma cRow_cross_tRow_conj (V : CKMMatrix) : conj (conj [V]c ×₃ conj [V]t) = [V]c ×₃ [V]t := by +lemma cRow_cross_tRow_conj (V : CKMMatrix) : conj (conj [V]c ×₃ conj [V]t) = [V]c ×₃ [V]t := by simp only [crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, LinearMap.mk₂_apply, Pi.conj_apply] funext i @@ -169,9 +169,9 @@ lemma rows_linearly_independent (V : CKMMatrix) : LinearIndependent ℂ (rows V) intro g hg rw [Fin.sum_univ_three] at hg simp only [Fin.isValue, rows] at hg - have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg - have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg - have h2 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg + have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg + have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg + have h2 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply, smul_eq_mul, dotProduct_zero] at h0 h1 h2 rw [uRow_normalized, uRow_cRow_orthog, uRow_tRow_orthog] at h0 @@ -184,7 +184,7 @@ lemma rows_linearly_independent (V : CKMMatrix) : LinearIndependent ℂ (rows V) · exact h1 · exact h2 -lemma rows_card : Fintype.card (Fin 3) = FiniteDimensional.finrank ℂ (Fin 3 → ℂ) := by +lemma rows_card : Fintype.card (Fin 3) = FiniteDimensional.finrank ℂ (Fin 3 → ℂ) := by simp /-- The rows of a CKM matrix as a basis of `ℂ³`. -/ @@ -198,24 +198,24 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) : (conj [V]c ×₃ conj [V]t)) simp only [Fin.sum_univ_three, rowBasis, Fin.isValue, coe_basisOfLinearIndependentOfCardEqFinrank, rows] at hg - have h0 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg - have h1 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg + have h0 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg + have h1 := congrArg (fun X => conj [V]t ⬝ᵥ X) hg simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply, smul_eq_mul, dotProduct_zero] at h0 h1 rw [cRow_normalized, cRow_uRow_orthog, cRow_tRow_orthog, dot_self_cross] at h0 rw [tRow_normalized, tRow_uRow_orthog, tRow_cRow_orthog, dot_cross_self] at h1 simp only [Fin.isValue, mul_zero, mul_one, zero_add, add_zero] at h0 h1 hg simp only [h0, h1, zero_smul, add_zero] at hg - have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg + have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg simp only [Fin.isValue, dotProduct_smul, Pi.conj_apply, Pi.smul_apply, smul_eq_mul, _root_.map_mul, cRow_cross_tRow_normalized] at h2 - have h3 : conj (g 0 • [V]u) = conj (g 0) • conj [V]u := by + have h3 : conj (g 0 • [V]u) = conj (g 0) • conj [V]u := by funext i fin_cases i <;> simp simp only [h3, Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul, uRow_normalized, Fin.isValue, mul_one, mul_conj, ← Complex.sq_abs] at h2 simp at h2 - cases' h2 with h2 h2 + cases' h2 with h2 h2 swap have hx : Complex.abs (g 0) = -1 := by simp [← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one, h2] @@ -223,7 +223,7 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) : simp_all have h4 : (0 : ℝ) < 1 := by norm_num · exact False.elim (lt_iff_not_le.mp h4 h3) - have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by + have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by rw [← hg, smul_smul, inv_mul_cancel, one_smul] by_contra hn simp [hn] at h2 @@ -238,30 +238,30 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) : rw [hx, hτ] lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) : - ∃ (τ : ℝ), [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c)) := by + ∃ (τ : ℝ), [V]t = cexp (τ * I) • (conj ([V]u) ×₃ conj ([V]c)) := by obtain ⟨g, hg⟩ := (mem_span_range_iff_exists_fun ℂ).mp (Basis.mem_span (rowBasis V) (conj ([V]u) ×₃ conj ([V]c)) ) rw [Fin.sum_univ_three, rowBasis] at hg simp at hg - have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg - have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg + have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg + have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply, smul_eq_mul, dotProduct_zero] at h0 h1 rw [uRow_normalized, uRow_cRow_orthog, uRow_tRow_orthog, dot_self_cross] at h0 rw [cRow_normalized, cRow_uRow_orthog, cRow_tRow_orthog, dot_cross_self] at h1 simp only [Fin.isValue, mul_one, mul_zero, add_zero, zero_add] at h0 h1 simp only [Fin.isValue, h0, zero_smul, h1, add_zero, zero_add] at hg - have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg + have h2 := congrArg (fun X => conj X ⬝ᵥ X) hg simp only [Fin.isValue, dotProduct_smul, Pi.conj_apply, Pi.smul_apply, smul_eq_mul, _root_.map_mul] at h2 rw [uRow_cross_cRow_normalized] at h2 - have h3 : conj (g 2 • [V]t) = conj (g 2) • conj [V]t := by + have h3 : conj (g 2 • [V]t) = conj (g 2) • conj [V]t := by funext i fin_cases i <;> simp simp only [h3, Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul, tRow_normalized, Fin.isValue, mul_one, mul_conj, ← Complex.sq_abs, ofReal_pow, sq_eq_one_iff, ofReal_eq_one] at h2 - cases' h2 with h2 h2 + cases' h2 with h2 h2 swap have hx : Complex.abs (g 2) = -1 := by simp [h2, ← ofReal_inj, Fin.isValue, ofReal_neg, ofReal_one] @@ -269,7 +269,7 @@ lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) : simp_all have h4 : (0 : ℝ) < 1 := by norm_num exact False.elim (lt_iff_not_le.mp h4 h3) - have hx : [V]t = (g 2)⁻¹ • (conj ([V]u) ×₃ conj ([V]c)) := by + have hx : [V]t = (g 2)⁻¹ • (conj ([V]u) ×₃ conj ([V]c)) := by rw [← hg, @smul_smul, inv_mul_cancel, one_smul] by_contra hn simp [hn] at h2 @@ -303,30 +303,30 @@ open CKMMatrix variable (V : CKMMatrix) (a b c d e f : ℝ) /-- The cross product of the conjugate of the `u` and `c` rows of a CKM matrix. -/ -def ucCross : Fin 3 → ℂ := +def ucCross : Fin 3 → ℂ := conj [phaseShiftApply V a b c d e f]u ×₃ conj [phaseShiftApply V a b c d e f]c lemma ucCross_fst (V : CKMMatrix) : (ucCross V a b c d e f) 0 = - cexp ((- a * I) + (- b * I) + ( - e * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 0 := by + cexp ((- a * I) + (- b * I) + ( - e * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 0 := by simp [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, LinearMap.mk₂_apply, Pi.conj_apply, cons_val_zero, neg_mul, uRow, us, ub, cRow, cs, cb, exp_add, exp_sub, ← exp_conj, conj_ofReal] ring lemma ucCross_snd (V : CKMMatrix) : (ucCross V a b c d e f) 1 = - cexp ((- a * I) + (- b * I) + ( - d * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 1 := by + cexp ((- a * I) + (- b * I) + ( - d * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 1 := by simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add, exp_sub, ← exp_conj, conj_ofReal] ring lemma ucCross_thd (V : CKMMatrix) : (ucCross V a b c d e f) 2 = - cexp ((- a * I) + (- b * I) + ( - d * I) + (- e * I)) * (conj [V]u ×₃ conj [V]c) 2 := by + cexp ((- a * I) + (- b * I) + ( - d * I) + (- e * I)) * (conj [V]u ×₃ conj [V]c) 2 := by simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add, exp_sub, ← exp_conj, conj_ofReal] ring lemma uRow_mul (V : CKMMatrix) (a b c : ℝ) : - [phaseShiftApply V a b c 0 0 0]u = cexp (a * I) • [V]u := by + [phaseShiftApply V a b c 0 0 0]u = cexp (a * I) • [V]u := by funext i simp only [Pi.smul_apply, smul_eq_mul] fin_cases i <;> @@ -336,7 +336,7 @@ lemma uRow_mul (V : CKMMatrix) (a b c : ℝ) : simp [ub, uRow] lemma cRow_mul (V : CKMMatrix) (a b c : ℝ) : - [phaseShiftApply V a b c 0 0 0]c = cexp (b * I) • [V]c := by + [phaseShiftApply V a b c 0 0 0]c = cexp (b * I) • [V]c := by funext i simp only [Pi.smul_apply, smul_eq_mul] fin_cases i <;> @@ -346,7 +346,7 @@ lemma cRow_mul (V : CKMMatrix) (a b c : ℝ) : simp [cb, cRow] lemma tRow_mul (V : CKMMatrix) (a b c : ℝ) : - [phaseShiftApply V a b c 0 0 0]t = cexp (c * I) • [V]t := by + [phaseShiftApply V a b c 0 0 0]t = cexp (c * I) • [V]t := by funext i simp only [Pi.smul_apply, smul_eq_mul] fin_cases i <;> @@ -355,6 +355,6 @@ lemma tRow_mul (V : CKMMatrix) (a b c : ℝ) : simp [ts, tRow, ofReal_zero, zero_mul, add_zero, Fin.isValue, Fin.mk_one, cons_val_one, head_cons] simp [tb, tRow] -end phaseShiftApply +end phaseShiftApply end diff --git a/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/Basic.lean b/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/Basic.lean index 3beb917..291554a 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/Basic.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/Basic.lean @@ -25,7 +25,7 @@ noncomputable section /-- Given four reals `θ₁₂ θ₁₃ θ₂₃ δ₁₃` the standard paramaterization of the CKM matrix as a `3×3` complex matrix. -/ -def standParamAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : Matrix (Fin 3) (Fin 3) ℂ := +def standParamAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : Matrix (Fin 3) (Fin 3) ℂ := ![![Real.cos θ₁₂ * Real.cos θ₁₃, Real.sin θ₁₂ * Real.cos θ₁₃, Real.sin θ₁₃ * exp (-I * δ₁₃)], ![(-Real.sin θ₁₂ * Real.cos θ₂₃) - (Real.cos θ₁₂ * Real.sin θ₁₃ * Real.sin θ₂₃ * exp (I * δ₁₃)), Real.cos θ₁₂ * Real.cos θ₂₃ - Real.sin θ₁₂ * Real.sin θ₁₃ * Real.sin θ₂₃ * exp (I * δ₁₃), @@ -36,8 +36,8 @@ def standParamAsMatrix (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : Matrix (Fin open CKMMatrix -lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : - ((standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by +lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : + ((standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃)ᴴ * standParamAsMatrix θ₁₂ θ₁₃ θ₂₃ δ₁₃) = 1 := by funext j i simp only [standParamAsMatrix, neg_mul, Fin.isValue] rw [mul_apply] @@ -139,7 +139,7 @@ lemma cross_product_t (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) : ring lemma eq_rows (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ} (hu : [U]u = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]u) - (hc : [U]c = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) (hU : [U]t = conj [U]u ×₃ conj [U]c) : + (hc : [U]c = [standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃]c) (hU : [U]t = conj [U]u ×₃ conj [U]c) : U = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ := by apply ext_Rows hu hc rw [hU, cross_product_t, hu, hc] @@ -149,7 +149,7 @@ lemma eq_exp_of_phases (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ℝ) (h simp [standParam, standParamAsMatrix] apply CKMMatrix_ext simp only - rw [show exp (I * δ₁₃) = exp (I * δ₁₃') by rw [mul_comm, h, mul_comm]] + rw [show exp (I * δ₁₃) = exp (I * δ₁₃') by rw [mul_comm, h, mul_comm]] rw [show cexp (-(I * ↑δ₁₃)) = cexp (-(I * ↑δ₁₃')) by rw [exp_neg, exp_neg, mul_comm, h, mul_comm]] open Invariant in @@ -179,9 +179,9 @@ lemma VusVubVcdSq_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Rea open Invariant in lemma mulExpδ₁₃_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Real.sin θ₁₂) - (h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) : + (h2 : 0 ≤ Real.cos θ₁₃) (h3 : 0 ≤ Real.sin θ₂₃) (h4 : 0 ≤ Real.cos θ₁₂) : mulExpδ₁₃ ⟦standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃⟧ = - sin θ₁₂ * cos θ₁₃ ^ 2 * sin θ₂₃ * sin θ₁₃ * cos θ₁₂ * cos θ₂₃ * cexp (I * δ₁₃) := by + sin θ₁₂ * cos θ₁₃ ^ 2 * sin θ₂₃ * sin θ₁₃ * cos θ₁₂ * cos θ₂₃ * cexp (I * δ₁₃) := by rw [mulExpδ₁₃, VusVubVcdSq_eq _ _ _ _ h1 h2 h3 h4 ] simp only [jarlskogℂ, standParam, standParamAsMatrix, neg_mul, Quotient.lift_mk, jarlskogℂCKM, Fin.isValue, cons_val', cons_val_one, head_cons, diff --git a/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean b/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean index 523c70b..0c05432 100644 --- a/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean +++ b/HepLean/FlavorPhysics/CKMMatrix/StandardParameterization/StandardParameters.lean @@ -12,7 +12,7 @@ import Mathlib.Analysis.SpecialFunctions.Complex.Arg /-! # Standard parameters for the CKM Matrix -Given a CKM matrix `V` we can extract four real numbers `θ₁₂`, `θ₁₃`, `θ₂₃` and `δ₁₃`. +Given a CKM matrix `V` we can extract four real numbers `θ₁₂`, `θ₁₃`, `θ₂₃` and `δ₁₃`. These, when used in the standard parameterization return `V` up to equivalence. This leads to the theorem `standParam.exists_for_CKMatrix` which says that up to equivalence every @@ -26,15 +26,15 @@ open CKMMatrix noncomputable section /-- Given a CKM matrix `V` the real number corresponding to `sin θ₁₂` in the -standard parameterization. --/ +standard parameterization. --/ def S₁₂ (V : Quotient CKMMatrixSetoid) : ℝ := VusAbs V / (√ (VudAbs V ^ 2 + VusAbs V ^ 2)) /-- Given a CKM matrix `V` the real number corresponding to `sin θ₁₃` in the -standard parameterization. --/ +standard parameterization. --/ def S₁₃ (V : Quotient CKMMatrixSetoid) : ℝ := VubAbs V /-- Given a CKM matrix `V` the real number corresponding to `sin θ₂₃` in the -standard parameterization. --/ +standard parameterization. --/ def S₂₃ (V : Quotient CKMMatrixSetoid) : ℝ := if VubAbs V = 1 then VcdAbs V else VcbAbs V / √ (VudAbs V ^ 2 + VusAbs V ^ 2) @@ -56,7 +56,7 @@ standard parameterization. --/ def C₁₂ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₁₂ V) /-- Given a CKM matrix `V` the real number corresponding to `cos θ₁₃` in the -standard parameterization. --/ +standard parameterization. --/ def C₁₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₁₃ V) /-- Given a CKM matrix `V` the real number corresponding to `sin θ₂₃` in the @@ -64,7 +64,7 @@ standard parameterization. --/ def C₂₃ (V : Quotient CKMMatrixSetoid) : ℝ := Real.cos (θ₂₃ V) /-- Given a CKM matrix `V` the real number corresponding to the phase `δ₁₃` in the -standard parameterization. --/ +standard parameterization. --/ def δ₁₃ (V : Quotient CKMMatrixSetoid) : ℝ := arg (Invariant.mulExpδ₁₃ V) @@ -336,7 +336,7 @@ namespace standParam open Invariant lemma mulExpδ₁₃_on_param_δ₁₃ (V : CKMMatrix) (δ₁₃ : ℝ) : - mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = + mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧) * cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) * cexp (I * δ₁₃) := by refine mulExpδ₁₃_eq _ _ _ _ ?_ ?_ ?_ ?_ @@ -348,11 +348,11 @@ lemma mulExpδ₁₃_on_param_δ₁₃ (V : CKMMatrix) (δ₁₃ : ℝ) : exact Real.cos_arcsin_nonneg _ lemma mulExpδ₁₃_on_param_eq_zero_iff (V : CKMMatrix) (δ₁₃ : ℝ) : - mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = 0 ↔ + mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = 0 ↔ VudAbs ⟦V⟧ = 0 ∨ VubAbs ⟦V⟧ = 0 ∨ VusAbs ⟦V⟧ = 0 ∨ VcbAbs ⟦V⟧ = 0 ∨ VtbAbs ⟦V⟧ = 0 := by rw [VudAbs_eq_C₁₂_mul_C₁₃, VubAbs_eq_S₁₃, VusAbs_eq_S₁₂_mul_C₁₃, VcbAbs_eq_S₂₃_mul_C₁₃, VtbAbs_eq_C₂₃_mul_C₁₃, ← ofReal_inj, - ← ofReal_inj, ← ofReal_inj, ← ofReal_inj, ← ofReal_inj] + ← ofReal_inj, ← ofReal_inj, ← ofReal_inj, ← ofReal_inj] simp only [ofReal_mul] rw [← S₁₃_eq_ℂsin_θ₁₃, ← S₁₂_eq_ℂsin_θ₁₂, ← S₂₃_eq_ℂsin_θ₂₃, ← C₁₃_eq_ℂcos_θ₁₃, ← C₂₃_eq_ℂcos_θ₂₃,← C₁₂_eq_ℂcos_θ₁₂] @@ -364,7 +364,7 @@ lemma mulExpδ₁₃_on_param_eq_zero_iff (V : CKMMatrix) (δ₁₃ : ℝ) : aesop lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ℝ) : - Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) = + Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) = sin (θ₁₂ ⟦V⟧) * cos (θ₁₃ ⟦V⟧) ^ 2 * sin (θ₂₃ ⟦V⟧) * sin (θ₁₃ ⟦V⟧) * cos (θ₁₂ ⟦V⟧) * cos (θ₂₃ ⟦V⟧) := by rw [mulExpδ₁₃_on_param_δ₁₃] @@ -373,19 +373,19 @@ lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ℝ) : complexAbs_sin_θ₂₃, complexAbs_cos_θ₂₃] lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : ℝ) - (h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) : - cexp (arg ( mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ) * I) = + (h1 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ≠ 0 ) : + cexp (arg ( mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ ) * I) = cexp (δ₁₃ * I) := by have h1a := mulExpδ₁₃_on_param_δ₁₃ V δ₁₃ have habs := mulExpδ₁₃_on_param_abs V δ₁₃ - have h2 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = Complex.abs - (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * exp (δ₁₃ * I) := by + have h2 : mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ = Complex.abs + (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * exp (δ₁₃ * I) := by rw [habs, h1a] ring_nf nth_rewrite 1 [← abs_mul_exp_arg_mul_I (mulExpδ₁₃ - ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ )] at h2 + ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧ )] at h2 have habs_neq_zero : - (Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ℂ) ≠ 0 := by + (Complex.abs (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) : ℂ) ≠ 0 := by simp only [ne_eq, ofReal_eq_zero, map_eq_zero] exact h1 rw [← mul_right_inj' habs_neq_zero] @@ -393,7 +393,7 @@ lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : ℝ) lemma on_param_cos_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.cos (θ₁₃ ⟦V⟧) = 0) : standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by - have hS13 := congrArg ofReal (S₁₃_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h)) + have hS13 := congrArg ofReal (S₁₃_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h)) simp [← S₁₃_eq_ℂsin_θ₁₃] at hS13 have hC12 := congrArg ofReal (C₁₂_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h)) simp [← C₁₂_eq_ℂcos_θ₁₂] at hC12 @@ -638,8 +638,8 @@ theorem eq_standardParameterization_δ₃ (V : CKMMatrix) : V ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) (δ₁₃ ⟦V⟧) := by obtain ⟨δ₁₃', hδ₃⟩ := exists_δ₁₃ V have hSV := (Quotient.eq.mpr (hδ₃)) - by_cases h : Invariant.mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'⟧ ≠ 0 - have h2 := eq_exp_of_phases (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃' + by_cases h : Invariant.mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃'⟧ ≠ 0 + have h2 := eq_exp_of_phases (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃' (δ₁₃ ⟦V⟧) (by rw [← mulExpδ₁₃_on_param_neq_zero_arg V δ₁₃' h, ← hSV, δ₁₃, Invariant.mulExpδ₁₃]) rw [h2] at hδ₃ exact hδ₃ diff --git a/HepLean/Mathematics/LinearMaps.lean b/HepLean/Mathematics/LinearMaps.lean index aac7ebd..ad781ef 100644 --- a/HepLean/Mathematics/LinearMaps.lean +++ b/HepLean/Mathematics/LinearMaps.lean @@ -76,7 +76,7 @@ def mk₂ (f : V × V → ℚ) (map_smul : ∀ a S T, f (a • S, T) = a * f (S, rw [swap, map_smul] exact congrArg (HMul.hMul a) (swap T S) } - map_smul' := fun a S => LinearMap.ext fun T => map_smul a S T + map_smul' := fun a S => LinearMap.ext fun T => map_smul a S T map_add' := fun S1 S2 => LinearMap.ext fun T => map_add S1 S2 T swap' := swap @@ -87,7 +87,7 @@ lemma map_smul₁ (f : BiLinearSymm V) (a : ℚ) (S T : V) : f (a • S) T = a * lemma swap (f : BiLinearSymm V) (S T : V) : f S T = f T S := f.swap' S T -lemma map_smul₂ (f : BiLinearSymm V) (a : ℚ) (S : V) (T : V) : f S (a • T) = a * f S T := by +lemma map_smul₂ (f : BiLinearSymm V) (a : ℚ) (S : V) (T : V) : f S (a • T) = a * f S T := by rw [f.swap, f.map_smul₁, f.swap] lemma map_add₁ (f : BiLinearSymm V) (S1 S2 T : V) : f (S1 + S2) T = f S1 T + f S2 T := by @@ -108,7 +108,7 @@ def toLinear₁ (f : BiLinearSymm V) (T : V) : V →ₗ[ℚ] ℚ where lemma toLinear₁_apply (f : BiLinearSymm V) (S T : V) : f S T = f.toLinear₁ T S := rfl -lemma map_sum₁ {n : ℕ} (f : BiLinearSymm V) (S : Fin n → V) (T : V) : +lemma map_sum₁ {n : ℕ} (f : BiLinearSymm V) (S : Fin n → V) (T : V) : f (∑ i, S i) T = ∑ i, f (S i) T := by rw [f.toLinear₁_apply] rw [map_sum] @@ -173,7 +173,7 @@ namespace TriLinearSymm open BigOperators variable {V : Type} [AddCommMonoid V] [Module ℚ V] -instance instFun : FunLike (TriLinearSymm V) V (V →ₗ[ℚ] V →ₗ[ℚ] ℚ ) where +instance instFun : FunLike (TriLinearSymm V) V (V →ₗ[ℚ] V →ₗ[ℚ] ℚ ) where coe f := f.toFun coe_injective' f g h := by cases f @@ -251,7 +251,7 @@ lemma map_add₃ (f : TriLinearSymm V) (S T L1 L2 : V) : rw [f.swap₃, f.map_add₁, f.swap₃, f.swap₃ L2 T S] /-- Fixing the second and third input vectors, the resulting linear map. -/ -def toLinear₁ (f : TriLinearSymm V) (T L : V) : V →ₗ[ℚ] ℚ where +def toLinear₁ (f : TriLinearSymm V) (T L : V) : V →ₗ[ℚ] ℚ where toFun S := f S T L map_add' S1 S2 := by simp only [f.map_add₁] diff --git a/HepLean/Mathematics/SO3/Basic.lean b/HepLean/Mathematics/SO3/Basic.lean index 4df983e..b4bc14c 100644 --- a/HepLean/Mathematics/SO3/Basic.lean +++ b/HepLean/Mathematics/SO3/Basic.lean @@ -81,7 +81,7 @@ lemma toGL_injective : Function.Injective toGL := by def toProd : SO(3) →* (Matrix (Fin 3) (Fin 3) ℝ) × (Matrix (Fin 3) (Fin 3) ℝ)ᵐᵒᵖ := MonoidHom.comp (Units.embedProduct _) toGL -lemma toProd_eq_transpose : toProd A = (A.1, ⟨A.1ᵀ⟩) := by +lemma toProd_eq_transpose : toProd A = (A.1, ⟨A.1ᵀ⟩) := by simp only [toProd, Units.embedProduct, coe_units_inv, MulOpposite.op_inv, toGL, coe_inv, MonoidHom.coe_comp, MonoidHom.coe_mk, OneHom.coe_mk, Function.comp_apply, Prod.mk.injEq, true_and] @@ -145,8 +145,8 @@ instance : TopologicalGroup SO(3) := lemma det_minus_id (A : SO(3)) : det (A.1 - 1) = 0 := by have h1 : det (A.1 - 1) = - det (A.1 - 1) := calc - det (A.1 - 1) = det (A.1 - A.1 * A.1ᵀ) := by simp [A.2.2] - _ = det A.1 * det (1 - A.1ᵀ) := by rw [← det_mul, mul_sub, mul_one] + det (A.1 - 1) = det (A.1 - A.1 * A.1ᵀ) := by simp [A.2.2] + _ = det A.1 * det (1 - A.1ᵀ) := by rw [← det_mul, mul_sub, mul_one] _ = det (1 - A.1ᵀ):= by simp [A.2.1] _ = det (1 - A.1ᵀ)ᵀ := by rw [det_transpose] _ = det (1 - A.1) := by simp @@ -160,7 +160,7 @@ lemma det_id_minus (A : SO(3)) : det (1 - A.1) = 0 := by have h1 : det (1 - A.1) = - det (A.1 - 1) := by calc det (1 - A.1) = det (- (A.1 - 1)) := by simp - _ = (- 1) ^ 3 * det (A.1 - 1) := by simp only [det_neg, Fintype.card_fin, neg_mul, one_mul] + _ = (- 1) ^ 3 * det (A.1 - 1) := by simp only [det_neg, Fintype.card_fin, neg_mul, one_mul] _ = - det (A.1 - 1) := by simp [pow_three] rw [h1, det_minus_id] simp only [neg_zero] diff --git a/HepLean/SpaceTime/LorentzAlgebra/Basic.lean b/HepLean/SpaceTime/LorentzAlgebra/Basic.lean index fddd9ef..d609698 100644 --- a/HepLean/SpaceTime/LorentzAlgebra/Basic.lean +++ b/HepLean/SpaceTime/LorentzAlgebra/Basic.lean @@ -28,7 +28,7 @@ def lorentzAlgebra : LieSubalgebra ℝ (Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin namespace lorentzAlgebra open minkowskiMatrix -lemma transpose_eta (A : lorentzAlgebra) : A.1ᵀ * η = - η * A.1 := by +lemma transpose_eta (A : lorentzAlgebra) : A.1ᵀ * η = - η * A.1 := by have h1 := A.2 erw [mem_skewAdjointMatricesLieSubalgebra] at h1 simpa [LieAlgebra.Orthogonal.so', IsSkewAdjoint, IsAdjointPair] using h1 @@ -39,7 +39,7 @@ lemma mem_of_transpose_eta_eq_eta_mul_self {A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 simpa [LieAlgebra.Orthogonal.so', IsSkewAdjoint, IsAdjointPair] using h lemma mem_iff {A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ} : - A ∈ lorentzAlgebra ↔ Aᵀ * η = - η * A := + A ∈ lorentzAlgebra ↔ Aᵀ * η = - η * A := Iff.intro (fun h => transpose_eta ⟨A, h⟩) (fun h => mem_of_transpose_eta_eq_eta_mul_self h) lemma mem_iff' (A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) : @@ -52,7 +52,7 @@ lemma mem_iff' (A : Matrix (Fin 1 ⊕ Fin 3) (Fin 1 ⊕ Fin 3) ℝ) : rw [minkowskiMatrix.sq] all_goals noncomm_ring · nth_rewrite 2 [h] - trans (η * η) * Aᵀ * η + trans (η * η) * Aᵀ * η rw [minkowskiMatrix.sq] all_goals noncomm_ring @@ -80,7 +80,7 @@ end lorentzAlgebra @[simps!] instance lorentzVectorAsLieRingModule : LieRingModule lorentzAlgebra (LorentzVector 3) where - bracket Λ x := Λ.1.mulVec x + bracket Λ x := Λ.1.mulVec x add_lie Λ1 Λ2 x := by simp [add_mulVec] lie_add Λ x1 x2 := by diff --git a/HepLean/SpaceTime/LorentzGroup/Basic.lean b/HepLean/SpaceTime/LorentzGroup/Basic.lean index 5df1c17..818fd66 100644 --- a/HepLean/SpaceTime/LorentzGroup/Basic.lean +++ b/HepLean/SpaceTime/LorentzGroup/Basic.lean @@ -30,7 +30,7 @@ We start studying the properties of matrices which preserve `ηLin`. These matrices form the Lorentz group, which we will define in the next section at `lorentzGroup`. -/ -variable {d : ℕ} +variable {d : ℕ} open minkowskiMetric in /-- The Lorentz group is the subset of matrices which preserve the minkowski metric. -/ @@ -74,7 +74,7 @@ lemma mem_iff_on_right : Λ ∈ LorentzGroup d ↔ rw [← dual_mulVec_right, mulVec_mulVec] exact h x y -lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by +lemma mem_iff_dual_mul_self : Λ ∈ LorentzGroup d ↔ dual Λ * Λ = 1 := by rw [mem_iff_on_right, matrix_eq_id_iff] exact forall_comm @@ -145,7 +145,7 @@ namespace LorentzGroup open minkowskiMetric -variable {Λ Λ' : LorentzGroup d} +variable {Λ Λ' : LorentzGroup d} lemma coe_inv : (Λ⁻¹).1 = Λ.1⁻¹:= by refine (inv_eq_left_inv ?h).symm @@ -172,7 +172,7 @@ def toGL : LorentzGroup d →* GL (Fin 1 ⊕ Fin d) ℝ where map_one' := by simp rfl - map_mul' x y := by + map_mul' x y := by simp only [lorentzGroupIsGroup, _root_.mul_inv_rev, coe_inv] ext rfl diff --git a/HepLean/SpaceTime/LorentzGroup/Boosts.lean b/HepLean/SpaceTime/LorentzGroup/Boosts.lean index 15552ca..9933061 100644 --- a/HepLean/SpaceTime/LorentzGroup/Boosts.lean +++ b/HepLean/SpaceTime/LorentzGroup/Boosts.lean @@ -99,7 +99,7 @@ lemma toMatrix_mulVec (u v : FuturePointing d) (x : LorentzVector d) : open minkowskiMatrix LorentzVector in @[simp] -lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) : +lemma toMatrix_apply (u v : FuturePointing d) (μ ν : Fin 1 ⊕ Fin d) : (toMatrix u v) μ ν = η μ μ * (⟪e μ, e ν⟫ₘ + 2 * ⟪e ν, u⟫ₘ * ⟪e μ, v⟫ₘ - ⟪e μ, u + v⟫ₘ * ⟪e ν, u + v⟫ₘ / (1 + ⟪u, v.1.1⟫ₘ)) := by rw [matrix_apply_stdBasis (toMatrix u v) μ ν, toMatrix_mulVec] diff --git a/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean b/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean index 56f7615..ce60f06 100644 --- a/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean +++ b/HepLean/SpaceTime/LorentzGroup/Orthochronous.lean @@ -27,7 +27,7 @@ variable (Λ : LorentzGroup d) open LorentzVector open minkowskiMetric -/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/ +/-- A Lorentz transformation is `orthochronous` if its `0 0` element is non-negative. -/ def IsOrthochronous : Prop := 0 ≤ timeComp Λ lemma IsOrthochronous_iff_futurePointing : @@ -62,7 +62,7 @@ lemma not_orthochronous_iff_le_zero : linarith /-- The continuous map taking a Lorentz transformation to its `0 0` element. -/ -def timeCompCont : C(LorentzGroup d, ℝ) := ⟨fun Λ => timeComp Λ , +def timeCompCont : C(LorentzGroup d, ℝ) := ⟨fun Λ => timeComp Λ, Continuous.matrix_elem (continuous_iff_le_induced.mpr fun _ a => a) (Sum.inl 0) (Sum.inl 0)⟩ /-- An auxillary function used in the definition of `orthchroMapReal`. -/ @@ -75,7 +75,7 @@ lemma stepFunction_continuous : Continuous stepFunction := by <;> intro a ha rw [@Set.Iic_def, @frontier_Iic, @Set.mem_singleton_iff] at ha rw [ha] - simp [neg_lt_self_iff, zero_lt_one, ↓reduceIte] + simp [neg_lt_self_iff, zero_lt_one, ↓reduceIte] have h1 : ¬ (1 : ℝ) ≤ 0 := by simp exact Eq.symm (if_neg h1) rw [Set.Ici_def, @frontier_Ici, @Set.mem_singleton_iff] at ha diff --git a/HepLean/SpaceTime/LorentzGroup/Proper.lean b/HepLean/SpaceTime/LorentzGroup/Proper.lean index 3e76268..8d139ae 100644 --- a/HepLean/SpaceTime/LorentzGroup/Proper.lean +++ b/HepLean/SpaceTime/LorentzGroup/Proper.lean @@ -31,7 +31,7 @@ lemma det_eq_one_or_neg_one (Λ : 𝓛 d) : Λ.1.det = 1 ∨ Λ.1.det = -1 := by simp [det_mul, det_dual] at h1 exact mul_self_eq_one_iff.mp h1 -local notation "ℤ₂" => Multiplicative (ZMod 2) +local notation "ℤ₂" => Multiplicative (ZMod 2) instance : TopologicalSpace ℤ₂ := instTopologicalSpaceFin @@ -51,7 +51,7 @@ def coeForℤ₂ : C(({-1, 1} : Set ℝ), ℤ₂) where /-- The continuous map taking a Lorentz matrix to its determinant. -/ def detContinuous : C(𝓛 d, ℤ₂) := - ContinuousMap.comp coeForℤ₂ { + ContinuousMap.comp coeForℤ₂ { toFun := fun Λ => ⟨Λ.1.det, Or.symm (LorentzGroup.det_eq_one_or_neg_one _)⟩, continuous_toFun := by refine Continuous.subtype_mk ?_ _ @@ -65,7 +65,7 @@ lemma detContinuous_eq_iff_det_eq (Λ Λ' : LorentzGroup d) : intro h simp [detContinuous] at h cases' det_eq_one_or_neg_one Λ with h1 h1 - <;> cases' det_eq_one_or_neg_one Λ' with h2 h2 + <;> cases' det_eq_one_or_neg_one Λ' with h2 h2 <;> simp_all [h1, h2, h] rw [← toMul_zero, @Equiv.apply_eq_iff_eq] at h · change (0 : Fin 2) = (1 : Fin 2) at h diff --git a/HepLean/SpaceTime/LorentzTensor/Basic.lean b/HepLean/SpaceTime/LorentzTensor/Basic.lean index f55df30..fdf22bb 100644 --- a/HepLean/SpaceTime/LorentzTensor/Basic.lean +++ b/HepLean/SpaceTime/LorentzTensor/Basic.lean @@ -47,7 +47,7 @@ instance (d : ℕ) (μ : RealLorentzTensor.Colors) : Fintype (RealLorentzTensor. | RealLorentzTensor.Colors.down => instFintypeSum (Fin 1) (Fin d) /-- An `IndexValue` is a set of actual values an index can take. e.g. for a - 3-tensor (0, 1, 2). -/ + 3-tensor (0, 1, 2). -/ @[simp] def RealLorentzTensor.IndexValue {X : FintypeCat} (d : ℕ ) (c : X → RealLorentzTensor.Colors) : Type 0 := (x : X) → RealLorentzTensor.ColorsIndex d (c x) @@ -107,9 +107,9 @@ def congrColorsDual {μ ν : Colors} (h : μ = τ ν) : ColorsIndex d μ ≃ ColorsIndex d ν := (castColorsIndex h).trans dualColorsIndex.symm -lemma congrColorsDual_symm {μ ν : Colors} (h : μ = τ ν) : +lemma congrColorsDual_symm {μ ν : Colors} (h : μ = τ ν) : (congrColorsDual h).symm = - @congrColorsDual d _ _ ((Function.Involutive.eq_iff τ_involutive).mp h.symm) := by + @congrColorsDual d _ _ ((Function.Involutive.eq_iff τ_involutive).mp h.symm) := by match μ, ν with | Colors.up, Colors.down => rfl | Colors.down, Colors.up => rfl @@ -124,7 +124,7 @@ lemma color_eq_dual_symm {μ ν : Colors} (h : μ = τ ν) : ν = τ μ := -/ /-- An equivalence of Index values from an equality of color maps. -/ -def castIndexValue {X : FintypeCat} {T S : X → Colors} (h : T = S) : +def castIndexValue {X : FintypeCat} {T S : X → Colors} (h : T = S) : IndexValue d T ≃ IndexValue d S where toFun i := (fun μ => castColorsIndex (congrFun h μ) (i μ)) invFun i := (fun μ => (castColorsIndex (congrFun h μ)).symm (i μ)) @@ -133,7 +133,7 @@ def castIndexValue {X : FintypeCat} {T S : X → Colors} (h : T = S) : right_inv i := by simp -lemma indexValue_eq {T₁ T₂ : X → RealLorentzTensor.Colors} (d : ℕ) (h : T₁ = T₂) : +lemma indexValue_eq {T₁ T₂ : X → RealLorentzTensor.Colors} (d : ℕ) (h : T₁ = T₂) : IndexValue d T₁ = IndexValue d T₂ := pi_congr fun a => congrArg (ColorsIndex d) (congrFun h a) @@ -176,7 +176,7 @@ lemma ext' {T₁ T₂ : RealLorentzTensor d X} (h : T₁.color = T₂.color) /-- An equivalence between `X → Fin 1 ⊕ Fin d` and `Y → Fin 1 ⊕ Fin d` given an isomorphism between `X` and `Y`. -/ def congrSetIndexValue (d : ℕ) (f : X ≃ Y) (i : X → Colors) : - IndexValue d i ≃ IndexValue d (i ∘ f.symm) := + IndexValue d i ≃ IndexValue d (i ∘ f.symm) := Equiv.piCongrLeft' _ f /-- Given an equivalence of indexing sets, a map on Lorentz tensors. -/ @@ -259,7 +259,7 @@ def inrIndexValue {Tc : X → Colors} {Sc : Y → Colors} def sumCommIndexValue {X Y : FintypeCat} (Tc : X → Colors) (Sc : Y → Colors) : IndexValue d (sumElimIndexColor Tc Sc) ≃ IndexValue d (sumElimIndexColor Sc Tc) := (congrSetIndexValue d (sumCommFintypeCat X Y) (sumElimIndexColor Tc Sc)).trans - (castIndexValue ((sumElimIndexColor_symm Sc Tc).symm)) + (castIndexValue ((sumElimIndexColor_symm Sc Tc).symm)) lemma sumCommIndexValue_inlIndexValue {X Y : FintypeCat} {Tc : X → Colors} {Sc : Y → Colors} (i : IndexValue d (sumElimIndexColor Tc Sc)) : @@ -300,7 +300,7 @@ def unmarkedColor (T : Marked d X n) : X → Colors := T.color ∘ Sum.inl /-- The colors of marked indices. -/ -def markedColor (T : Marked d X n) : FintypeCat.of (Σ _ : Fin n, PUnit) → Colors := +def markedColor (T : Marked d X n) : FintypeCat.of (Σ _ : Fin n, PUnit) → Colors := T.color ∘ Sum.inr /-- The index values restricted to unmarked indices. -/ diff --git a/scripts/README.md b/scripts/README.md index 2b6e300..bb0b952 100644 --- a/scripts/README.md +++ b/scripts/README.md @@ -7,14 +7,16 @@ Taken from Mathlib's linters. These perform text-based linting of the code. These are to be slowly replaced with code written in Lean. -## double_line_lint +## hepLean_style_lint -Checks for double empty lines in HepLean. +Checks the following in HepLean +- There are no double empty lines. +- There are no non-initial double spaces. Passing this linter is currently not required to pass CI on github. Run using ``` -lake exe double_line_lint +lake exe hepLean_style_lint ``` ## check_file_imports.lean diff --git a/scripts/hepLean_style_lint.lean b/scripts/hepLean_style_lint.lean index 2cff93a..cf7b683 100644 --- a/scripts/hepLean_style_lint.lean +++ b/scripts/hepLean_style_lint.lean @@ -7,9 +7,11 @@ import Lean import Mathlib.Tactic.Linter.TextBased /-! -# Double line Lint +# HepLean style linter -This linter double empty lines in files. +A number of linters on HepLean to enforce a consistent style. + +There are currently not enforced at the GitHub action level. ## Note @@ -31,7 +33,7 @@ def doubleEmptyLineLinter : HepLeanTextLinter := fun lines ↦ Id.run do else none) errors.toArray -/-- Checks if there is a souble space in the line, which is not at the start. -/ +/-- Checks if there is a double space in the line, which is not at the start. -/ def doubleSpaceLinter : HepLeanTextLinter := fun lines ↦ Id.run do let enumLines := (lines.toList.enumFrom 1) let errors := enumLines.filterMap (fun (lno, l) ↦ From 8440f133609608475969a935b42a1be78d8e8b1d Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Fri, 12 Jul 2024 11:28:31 -0400 Subject: [PATCH 5/5] docs: Add punctuation --- HepLean/AnomalyCancellation/GroupActions.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/HepLean/AnomalyCancellation/GroupActions.lean b/HepLean/AnomalyCancellation/GroupActions.lean index be298c2..a6c1e30 100644 --- a/HepLean/AnomalyCancellation/GroupActions.lean +++ b/HepLean/AnomalyCancellation/GroupActions.lean @@ -22,7 +22,7 @@ From this we define the vector spaces of charges under which the anomaly equations are invariant. -/ structure ACCSystemGroupAction (χ : ACCSystem) where - /-- The underlying type of the group-/ + /-- The underlying type of the group. -/ group : Type /-- An instance given group the structure of a group. -/ groupInst : Group group