Merge branch 'Update-versions' into Tensors

This commit is contained in:
jstoobysmith 2024-07-12 16:45:09 -04:00
commit e393ac1b70
87 changed files with 269 additions and 216 deletions

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Mathematics.LinearMaps

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Tactic.FinCases

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic
@ -56,7 +56,7 @@ def proj (T : MSSMACC.LinSols) : MSSMACC.AnomalyFreePerp :=
lemma proj_val (T : MSSMACC.LinSols) :
(proj T).val = (dot B₃.val T.val - dot Y₃.val T.val) • Y₃.val +
( (dot Y₃.val T.val - 2 * dot B₃.val T.val)) • B₃.val +
(dot Y₃.val T.val - 2 * dot B₃.val T.val) • B₃.val +
dot Y₃.val B₃.val • T.val := by
rfl

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.MSSMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.Basic
@ -51,22 +51,19 @@ def accCubeTriLinSymm {n : } : TriLinearSymm (PureU1Charges n).Charges := Tri
rw [← Finset.sum_add_distrib]
apply Fintype.sum_congr
intro i
ring
)
ring)
(by
intro S L T
simp only [PureU1Charges_numberCharges]
apply Fintype.sum_congr
intro i
ring
)
ring)
(by
intro S L T
simp only [PureU1Charges_numberCharges]
apply Fintype.sum_congr
intro i
ring
)
ring)
/-- The cubic anomaly equation. -/
@[simp]

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Sort

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Sort

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic
@ -114,12 +114,12 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ )).LinSols}
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
linear_combination -(1 * h2)
have hll : f (Fin.castSucc (Fin.last (n ))) =
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

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Sort

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic
@ -303,8 +303,7 @@ lemma Prop_two (P : × → Prop) {S : (PureU1 n).LinSols}
{a b : Fin n} (hab : a ≠ b)
(h : ∀ (f : (FamilyPermutations n).group),
P ((((FamilyPermutations n).linSolRep f S).val a),
(((FamilyPermutations n).linSolRep f S).val b)
)) : ∀ (i j : Fin n) (_ : i ≠ j),
(((FamilyPermutations n).linSolRep f S).val b))) : ∀ (i j : Fin n) (_ : i ≠ j),
P (S.val i, S.val j) := by
intro i j hij
have h1 := h (permTwo hij hab).symm
@ -321,9 +320,8 @@ lemma Prop_three (P : × × → Prop) {S : (PureU1 n).LinSols}
(h : ∀ (f : (FamilyPermutations n).group),
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),
P (S.val i, (S.val j, S.val k)) := by
(((FamilyPermutations n).linSolRep f S).val c)))) : ∀ (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,

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Permutations

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Sort

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Tactic.FinCases
@ -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
@ -272,8 +272,7 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
intro i
repeat erw [map_smul]
simp [HSMul.hSMul, SMul.smul]
ring
)
ring)
(by
intro S T R L
simp only
@ -282,22 +281,19 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
intro i
repeat erw [map_add]
simp only [ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply, Fin.isValue]
ring
)
ring)
(by
intro S T L
simp only [SMSpecies_numberCharges, toSpecies_apply, Fin.isValue]
apply Fintype.sum_congr
intro i
ring
)
ring)
(by
intro S T L
simp only [SMSpecies_numberCharges, toSpecies_apply, Fin.isValue]
apply Fintype.sum_congr
intro i
ring
)
ring)
/-- The cubic acc. -/
@[simp]

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SM.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SM.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SM.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SM.Basic
@ -183,7 +183,8 @@ lemma grav (S : linearParameters) :
end linearParameters
/-- The parameters for solutions to the linear ACCs with the condition that Q and E are non-zero.-/
/-- The parameters for solutions to the linear ACCs with the condition that Q and E are
non-zero. -/
structure linearParametersQENeqZero where
/-- The parameter `x`. -/
x :

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SM.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Tactic.FinCases
@ -54,7 +54,7 @@ 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

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Ordinary.Basic
@ -28,8 +28,7 @@ def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 0, 0 => 1
| 0, 1 => - 1
| _, _ => 0
)
| _, _ => 0)
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
@ -41,8 +40,7 @@ def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 1, 0 => 1
| 1, 1 => - 1
| _, _ => 0
)
| _, _ => 0)
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
@ -54,8 +52,7 @@ def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 2, 0 => 1
| 2, 1 => - 1
| _, _ => 0
)
| _, _ => 0)
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
@ -67,8 +64,7 @@ def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 3, 0 => 1
| 3, 1 => - 1
| _, _ => 0
)
| _, _ => 0)
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
@ -81,8 +77,7 @@ def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 4, 0 => 1
| 4, 1 => - 1
| _, _ => 0
)
| _, _ => 0)
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
@ -95,8 +90,7 @@ def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 5, 0 => 1
| 5, 1 => - 1
| _, _ => 0
)
| _, _ => 0)
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
@ -109,8 +103,7 @@ def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 1, 2 => 1
| 2, 2 => -1
| _, _ => 0
)
| _, _ => 0)
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

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Ordinary.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.StandardModel.HiggsBoson.Potential

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Logic.Equiv.Fin

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FeynmanDiagrams.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FeynmanDiagrams.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FeynmanDiagrams.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.LinearAlgebra.UnitaryGroup

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FlavorPhysics.CKMMatrix.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FlavorPhysics.CKMMatrix.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FlavorPhysics.CKMMatrix.Basic
@ -158,7 +158,7 @@ lemma Vcd_mul_conj_Vud (V : CKMMatrix) :
lemma Vcs_mul_conj_Vus (V : CKMMatrix) :
[V]cs * conj [V]us = - [V]cd * conj [V]ud - [V]cb * conj [V]ub := by
linear_combination (V.fst_row_orthog_snd_row )
linear_combination V.fst_row_orthog_snd_row
end orthogonal

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FlavorPhysics.CKMMatrix.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FlavorPhysics.CKMMatrix.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.FlavorPhysics.CKMMatrix.Basic
@ -505,7 +505,7 @@ lemma eq_standParam_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0
rw [ud_us_neq_zero_iff_ub_neq_one] at hb
simp [VAbs, hb]
have h1 : ofReal (√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) *
↑√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2)) = ofReal ((VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) ) := by
↑√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2)) = ofReal (VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) := by
rw [Real.mul_self_sqrt ]
apply add_nonneg (sq_nonneg _) (sq_nonneg _)
simp at h1

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Tactic.Polyrith

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.LinearAlgebra.UnitaryGroup

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Analysis.Complex.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.MinkowskiMetric

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzAlgebra.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.MinkowskiMetric
@ -239,7 +239,7 @@ open LorentzVector
-/
/-- The first column of a lorentz matrix as a `NormOneLorentzVector`. -/
/-- The first column of a Lorentz matrix as a `NormOneLorentzVector`. -/
@[simps!]
def toNormOneLorentzVector (Λ : LorentzGroup d) : NormOneLorentzVector d :=
⟨Λ.1 *ᵥ timeVec, by rw [NormOneLorentzVector.mem_iff, Λ.2, minkowskiMetric.on_timeVec]⟩

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Proper

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.NormOne
@ -150,7 +150,7 @@ lemma mul_not_othchron_of_not_othchron_othchron {Λ Λ' : LorentzGroup d} (h :
rw [IsOrthochronous_iff_futurePointing] at h h'
exact NormOneLorentzVector.FuturePointing.metric_reflect_not_mem_mem h h'
/-- The homomorphism from `lorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
/-- The homomorphism from `LorentzGroup` to `ℤ₂` whose kernel are the Orthochronous elements. -/
def orthchroRep : LorentzGroup d →* ℤ₂ where
toFun := orthchroMap
map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous])

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic
@ -122,7 +122,7 @@ lemma IsProper_iff (Λ : LorentzGroup d) : IsProper Λ ↔ detRep Λ = 1 := by
rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq]
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
lemma id_IsProper : (@IsProper d) 1 := by
lemma id_IsProper : @IsProper d 1 := by
simp [IsProper]
lemma isProper_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) :

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
/-!

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.MinkowskiMetric

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzVector.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.LorentzGroup.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.SpaceTime.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.StandardModel.HiggsBoson.Basic

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.StandardModel.HiggsBoson.Basic
@ -95,7 +95,7 @@ lemma smooth_innerProd (φ1 φ2 : HiggsField) : Smooth 𝓘(, SpaceTime) 𝓘
/-- Given a `HiggsField`, the map `SpaceTime → ` obtained by taking the square norm of the
pointwise Higgs vector. -/
@[simp]
def normSq (φ : HiggsField) : SpaceTime → := fun x => ( ‖φ x‖ ^ 2)
def normSq (φ : HiggsField) : SpaceTime → := fun x => ‖φ x‖ ^ 2
/-- Notation for the norm squared of a Higgs field. -/
scoped[StandardModel.HiggsField] notation "‖" φ1 "‖_H ^ 2" => normSq φ1

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Algebra.QuadraticDiscriminant

View file

@ -1,6 +1,6 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Data.Complex.Exponential

View file

@ -26,7 +26,7 @@ name = "HepLean"
#[[require]]
#name = "LeanCopilot"
#git = "https://github.com/lean-dojo/LeanCopilot.git"
#rev = "v1.2.2"
#rev = "v1.4.1"
# lean_exe commands defined specifically for HepLean.
@ -45,3 +45,7 @@ srcDir = "scripts"
[[lean_exe]]
name = "find_TODOs"
srcDir = "scripts"
[[lean_exe]]
name = "mathlib_textLint_on_hepLean"
srcDir = "scripts"

View file

@ -46,6 +46,19 @@ def doubleSpaceLinter : HepLeanTextLinter := fun lines ↦ Id.run do
else none)
errors.toArray
/-- Substring linter. -/
def substringLinter (s : String) : HepLeanTextLinter := fun lines ↦ Id.run do
let enumLines := (lines.toList.enumFrom 1)
let errors := enumLines.filterMap (fun (lno, l) ↦
if String.containsSubstr l s then
let k := (Substring.findAllSubstr l s).toList.getLast?
let col := match k with
| none => 1
| some k => String.offsetOfPos l k.stopPos
some (s!" Found instance of substring {s}.", lno, col)
else none)
errors.toArray
structure HepLeanErrorContext where
/-- The underlying `message`. -/
error : String
@ -64,7 +77,8 @@ def hepLeanLintFile (path : FilePath) : IO Bool := do
let lines ← IO.FS.lines path
let allOutput := (Array.map (fun lint ↦
(Array.map (fun (e, n, c) ↦ HepLeanErrorContext.mk e n c path)) (lint lines)))
#[doubleEmptyLineLinter, doubleSpaceLinter]
#[doubleEmptyLineLinter, doubleSpaceLinter, substringLinter ".-/", substringLinter " )",
substringLinter "( ", substringLinter "=by", substringLinter " def "]
let errors := allOutput.flatten
printErrors errors
return errors.size > 0
@ -85,4 +99,6 @@ def main (_ : List String) : IO UInt32 := do
warned := true
if warned then
throw <| IO.userError s!"Errors found."
else
IO.println "No linting issues found."
return 0

View file

@ -0,0 +1,48 @@
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0.
Authors: Joseph Tooby-Smith
-/
import Mathlib.Tactic.Linter.TextBased
import Cli.Basic
/-!
# Text based style linters from Mathlib
This file is a copy of the `./scripts/lint_style.lean` executable from mathlib, adapted
to run text-based style linters from mathlib on HepLean.
That file is copyright Michael Rothgang, and is released under the Apache 2.0 license.
It is authored by Michael Rothgang.
-/
open Cli
/-- Implementation of the `lint_style` command line program. -/
def lintStyleCli (args : Cli.Parsed) : IO UInt32 := do
let errorStyle := match (args.hasFlag "github", args.hasFlag "update") with
| (true, _) => ErrorFormat.github
| (false, true) => ErrorFormat.exceptionsFile
| (false, false) => ErrorFormat.humanReadable
let mut numberErrorFiles : UInt32 := 0
for s in ["HepLean.lean"] do
let n ← lintAllFiles (System.mkFilePath [s]) errorStyle
numberErrorFiles := numberErrorFiles + n
return numberErrorFiles
/-- Setting up command line options and help text for `lake exe lint_style`. -/
-- so far, no help options or so: perhaps that is fine?
def heplean_lint_style : Cmd := `[Cli|
lint_style VIA lintStyleCli; ["0.0.1"]
"Run text-based style linters on every Lean file in HepLean (adapted from mathlib's lint_style).
Print errors about any unexpected style errors to standard output."
FLAGS:
github; "Print errors in a format suitable for github problem matchers\n\
otherwise, produce human-readable output"
update; "Print errors solely for the style exceptions file"
]
/-- The entry point to the `lake exe mathlib_textLint_on_hepLean` command. -/
def main (args : List String) : IO UInt32 := do heplean_lint_style.validate args

View file