Merge branch 'Update-versions' into Tensors
This commit is contained in:
commit
e393ac1b70
87 changed files with 269 additions and 216 deletions
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.Mathematics.LinearMaps
|
import HepLean.Mathematics.LinearMaps
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.Basic
|
import HepLean.AnomalyCancellation.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Tactic.FinCases
|
import Mathlib.Tactic.FinCases
|
||||||
|
@ -76,7 +76,7 @@ def toSMSpecies (i : Fin 6) : MSSMCharges.Charges →ₗ[ℚ] MSSMSpecies.Charge
|
||||||
|
|
||||||
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
|
(toSMSpecies i) (toSpecies.symm f) = f.1 i := by
|
||||||
change (Prod.fst ∘ toSpecies ∘ toSpecies.symm ) _ i= f.1 i
|
change (Prod.fst ∘ toSpecies ∘ toSpecies.symm) _ i= f.1 i
|
||||||
simp
|
simp
|
||||||
|
|
||||||
/-- The `Q` charges as a map `Fin 3 → ℚ`. -/
|
/-- The `Q` charges as a map `Fin 3 → ℚ`. -/
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
@ -56,7 +56,7 @@ def proj (T : MSSMACC.LinSols) : MSSMACC.AnomalyFreePerp :=
|
||||||
|
|
||||||
lemma proj_val (T : MSSMACC.LinSols) :
|
lemma proj_val (T : MSSMACC.LinSols) :
|
||||||
(proj T).val = (dot B₃.val T.val - dot Y₃.val T.val) • Y₃.val +
|
(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
|
dot Y₃.val B₃.val • T.val := by
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
|
@ -110,7 +110,7 @@ lemma quad_self_proj (T : MSSMACC.Sols) :
|
||||||
lemma quad_proj (T : MSSMACC.Sols) :
|
lemma quad_proj (T : MSSMACC.Sols) :
|
||||||
quadBiLin (proj T.1.1).val (proj T.1.1).val = 2 * dot Y₃.val B₃.val *
|
quadBiLin (proj T.1.1).val (proj T.1.1).val = 2 * dot Y₃.val B₃.val *
|
||||||
((dot B₃.val T.val - dot Y₃.val T.val) * quadBiLin Y₃.val T.val +
|
((dot B₃.val T.val - dot Y₃.val T.val) * quadBiLin Y₃.val T.val +
|
||||||
(dot Y₃.val T.val - 2 * dot B₃.val T.val) * quadBiLin B₃.val T.val ) := by
|
(dot Y₃.val T.val - 2 * dot B₃.val T.val) * quadBiLin B₃.val T.val) := by
|
||||||
nth_rewrite 1 [proj_val]
|
nth_rewrite 1 [proj_val]
|
||||||
repeat rw [quadBiLin.map_add₁]
|
repeat rw [quadBiLin.map_add₁]
|
||||||
repeat rw [quadBiLin.map_smul₁]
|
repeat rw [quadBiLin.map_smul₁]
|
||||||
|
@ -159,7 +159,7 @@ lemma cube_proj_proj_self (T : MSSMACC.Sols) :
|
||||||
cubeTriLin (proj T.1.1).val (proj T.1.1).val T.val =
|
cubeTriLin (proj T.1.1).val (proj T.1.1).val T.val =
|
||||||
2 * dot Y₃.val B₃.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
|
(dot Y₃.val T.val - 2 * dot B₃.val T.val) * cubeTriLin T.val T.val B₃.val) := by
|
||||||
rw [proj_val]
|
rw [proj_val]
|
||||||
rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂]
|
rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂]
|
||||||
erw [lineY₃B₃_doublePoint]
|
erw [lineY₃B₃_doublePoint]
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
@ -155,23 +155,23 @@ def α₁ (T : MSSMACC.AnomalyFreePerp) : ℚ :=
|
||||||
(3 * cubeTriLin T.val T.val B₃.val * quadBiLin T.val T.val -
|
(3 * cubeTriLin T.val T.val B₃.val * quadBiLin T.val T.val -
|
||||||
2 * cubeTriLin T.val T.val T.val * quadBiLin B₃.val T.val)
|
2 * cubeTriLin T.val T.val T.val * quadBiLin B₃.val T.val)
|
||||||
|
|
||||||
/-- A helper function to simplify following expressions. -/
|
/-- A helper function to simplify following expressions. -/
|
||||||
def α₂ (T : MSSMACC.AnomalyFreePerp) : ℚ :=
|
def α₂ (T : MSSMACC.AnomalyFreePerp) : ℚ :=
|
||||||
(2 * cubeTriLin T.val T.val T.val * quadBiLin Y₃.val T.val -
|
(2 * cubeTriLin T.val T.val T.val * quadBiLin Y₃.val T.val -
|
||||||
3 * cubeTriLin T.val T.val Y₃.val * quadBiLin T.val T.val)
|
3 * cubeTriLin T.val T.val Y₃.val * quadBiLin T.val T.val)
|
||||||
|
|
||||||
/-- A helper function to simplify following expressions. -/
|
/-- A helper function to simplify following expressions. -/
|
||||||
def α₃ (T : MSSMACC.AnomalyFreePerp) : ℚ :=
|
def α₃ (T : MSSMACC.AnomalyFreePerp) : ℚ :=
|
||||||
6 * ((cubeTriLin T.val T.val Y₃.val) * quadBiLin B₃.val T.val -
|
6 * ((cubeTriLin T.val T.val Y₃.val) * quadBiLin B₃.val T.val -
|
||||||
(cubeTriLin T.val T.val B₃.val) * quadBiLin Y₃.val T.val)
|
(cubeTriLin T.val T.val B₃.val) * quadBiLin Y₃.val T.val)
|
||||||
|
|
||||||
lemma lineQuad_cube (R : MSSMACC.AnomalyFreePerp) (c₁ c₂ c₃ : ℚ) :
|
lemma lineQuad_cube (R : MSSMACC.AnomalyFreePerp) (c₁ c₂ c₃ : ℚ) :
|
||||||
accCube (lineQuad R c₁ c₂ c₃).val =
|
accCube (lineQuad R c₁ c₂ c₃).val =
|
||||||
- 4 * ( c₁ * quadBiLin B₃.val R.val - c₂ * quadBiLin Y₃.val R.val) ^ 2 *
|
- 4 * (c₁ * quadBiLin B₃.val R.val - c₂ * quadBiLin Y₃.val R.val) ^ 2 *
|
||||||
( α₁ R * c₁ + α₂ R * c₂ + α₃ R * c₃ ) := by
|
(α₁ R * c₁ + α₂ R * c₂ + α₃ R * c₃) := by
|
||||||
rw [lineQuad_val]
|
rw [lineQuad_val]
|
||||||
rw [planeY₃B₃_cubic, α₁, α₂, α₃]
|
rw [planeY₃B₃_cubic, α₁, α₂, α₃]
|
||||||
ring
|
ring
|
||||||
|
|
||||||
/-- The line in the plane spanned by `Y₃`, `B₃` and `R` which is in the cubic. -/
|
/-- The line in the plane spanned by `Y₃`, `B₃` and `R` which is in the cubic. -/
|
||||||
def lineCube (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ℚ) :
|
def lineCube (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ℚ) :
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
@ -147,7 +147,7 @@ def InCubeSolProp (R : MSSMACC.Sols) : Prop :=
|
||||||
that solution satisfies `inCubeSolProp`. It appears in the definition of `inLineEqProj`. -/
|
that solution satisfies `inCubeSolProp`. It appears in the definition of `inLineEqProj`. -/
|
||||||
def cubicCoeff (T : MSSMACC.Sols) : ℚ :=
|
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 )
|
cubeTriLin T.val T.val B₃.val ^ 2)
|
||||||
|
|
||||||
lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
||||||
InCubeSolProp T ↔ cubicCoeff T = 0 := by
|
InCubeSolProp T ↔ cubicCoeff T = 0 := by
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
@ -60,7 +60,7 @@ lemma chargeMap_toSpecies (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6)
|
||||||
@[simp]
|
@[simp]
|
||||||
def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where
|
def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where
|
||||||
toFun f := chargeMap f⁻¹
|
toFun f := chargeMap f⁻¹
|
||||||
map_mul' f g :=by
|
map_mul' f g := by
|
||||||
simp only [PermGroup, mul_inv_rev]
|
simp only [PermGroup, mul_inv_rev]
|
||||||
apply LinearMap.ext
|
apply LinearMap.ext
|
||||||
intro S
|
intro S
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.Basic
|
import HepLean.AnomalyCancellation.Basic
|
||||||
|
@ -51,22 +51,19 @@ def accCubeTriLinSymm {n : ℕ} : TriLinearSymm (PureU1Charges n).Charges := Tri
|
||||||
rw [← Finset.sum_add_distrib]
|
rw [← Finset.sum_add_distrib]
|
||||||
apply Fintype.sum_congr
|
apply Fintype.sum_congr
|
||||||
intro i
|
intro i
|
||||||
ring
|
ring)
|
||||||
)
|
|
||||||
(by
|
(by
|
||||||
intro S L T
|
intro S L T
|
||||||
simp only [PureU1Charges_numberCharges]
|
simp only [PureU1Charges_numberCharges]
|
||||||
apply Fintype.sum_congr
|
apply Fintype.sum_congr
|
||||||
intro i
|
intro i
|
||||||
ring
|
ring)
|
||||||
)
|
|
||||||
(by
|
(by
|
||||||
intro S L T
|
intro S L T
|
||||||
simp only [PureU1Charges_numberCharges]
|
simp only [PureU1Charges_numberCharges]
|
||||||
apply Fintype.sum_congr
|
apply Fintype.sum_congr
|
||||||
intro i
|
intro i
|
||||||
ring
|
ring)
|
||||||
)
|
|
||||||
|
|
||||||
/-- The cubic anomaly equation. -/
|
/-- The cubic anomaly equation. -/
|
||||||
@[simp]
|
@[simp]
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -117,7 +117,7 @@ def coordinateMap : ((PureU1 n.succ).LinSols) ≃ₗ[ℚ] Fin n →₀ ℚ where
|
||||||
exact Rat.mul_zero (f k)
|
exact Rat.mul_zero (f k)
|
||||||
simp
|
simp
|
||||||
|
|
||||||
/-- The basis of `LinSols`.-/
|
/-- The basis of `LinSols`. -/
|
||||||
noncomputable
|
noncomputable
|
||||||
def asBasis : Basis (Fin n) ℚ ((PureU1 n.succ).LinSols) where
|
def asBasis : Basis (Fin n) ℚ ((PureU1 n.succ).LinSols) where
|
||||||
repr := coordinateMap
|
repr := coordinateMap
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Sort
|
import HepLean.AnomalyCancellation.PureU1.Sort
|
||||||
|
@ -191,7 +191,7 @@ lemma AFL_odd_noBoundary {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted
|
||||||
lemma AFL_odd_zero {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted A.val) :
|
lemma AFL_odd_zero {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted A.val) :
|
||||||
A.val (0 : Fin (2 * n + 1)) = 0 := by
|
A.val (0 : Fin (2 * n + 1)) = 0 := by
|
||||||
by_contra hn
|
by_contra hn
|
||||||
exact (AFL_odd_noBoundary h hn ) (AFL_hasBoundary h hn)
|
exact (AFL_odd_noBoundary h hn) (AFL_hasBoundary h hn)
|
||||||
|
|
||||||
theorem AFL_odd (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
|
A = 0 := by
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Sort
|
import HepLean.AnomalyCancellation.PureU1.Sort
|
||||||
|
@ -689,7 +689,7 @@ lemma P!_in_span (f : Fin n → ℚ) : P! f ∈ Submodule.span ℚ (Set.range ba
|
||||||
use f
|
use f
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
lemma smul_basis!AsCharges_in_span (S : (PureU1 (2 * n.succ )).LinSols) (j : Fin n) :
|
lemma smul_basis!AsCharges_in_span (S : (PureU1 (2 * n.succ)).LinSols) (j : Fin n) :
|
||||||
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∈
|
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∈
|
||||||
Submodule.span ℚ (Set.range basis!AsCharges) := by
|
Submodule.span ℚ (Set.range basis!AsCharges) := by
|
||||||
apply Submodule.smul_mem
|
apply Submodule.smul_mem
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -69,7 +69,7 @@ lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic
|
||||||
|
|
||||||
/-- 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 :=
|
def LineInCubicPerm (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=
|
||||||
∀ (M : (FamilyPermutations (2 * n.succ)).group ),
|
∀ (M : (FamilyPermutations (2 * n.succ)).group),
|
||||||
LineInCubic ((FamilyPermutations (2 * n.succ)).linSolRep M S)
|
LineInCubic ((FamilyPermutations (2 * n.succ)).linSolRep M S)
|
||||||
|
|
||||||
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
||||||
|
@ -104,7 +104,7 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols}
|
||||||
rw [accCubeTriLinSymm.map_add₃, h1, accCubeTriLinSymm.map_smul₃] at h2
|
rw [accCubeTriLinSymm.map_add₃, h1, accCubeTriLinSymm.map_smul₃] at h2
|
||||||
simpa using h2
|
simpa using h2
|
||||||
|
|
||||||
lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.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) :
|
(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))) * (2 * S.val δ!₄ +
|
||||||
|
@ -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 h2 := Pa_δ!₁ f g (Fin.last n)
|
||||||
have h3 := Pa_δ!₂ f g (Fin.last n)
|
have h3 := Pa_δ!₂ f g (Fin.last n)
|
||||||
simp at h1 h2 h3
|
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]
|
simp_all only [Fin.succ_last, neg_neg]
|
||||||
erw [hl] at h2
|
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)
|
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
|
- (Pa f g (δ!₂ (Fin.last n)) + Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄) := by
|
||||||
linear_combination h3 - 1 * hg
|
linear_combination h3 - 1 * hg
|
||||||
rw [← hS] at hl hll
|
rw [← hS] at hl hll
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -93,7 +93,7 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
||||||
rw [hS'.1, hS'.2] at hC
|
rw [hS'.1, hS'.2] at hC
|
||||||
exact hC' 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`. -/
|
||||||
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) ,
|
∀ (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
|
||||||
|
@ -125,7 +125,7 @@ theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : GenericCase S) :
|
||||||
rw [parameterization]
|
rw [parameterization]
|
||||||
apply ACCSystem.Sols.ext
|
apply ACCSystem.Sols.ext
|
||||||
rw [parameterizationAsLinear_val]
|
rw [parameterizationAsLinear_val]
|
||||||
change S.val = _ • ( _ • P g + _• P! f)
|
change S.val = _ • (_ • P g + _• P! f)
|
||||||
rw [anomalyFree_param _ _ hS]
|
rw [anomalyFree_param _ _ hS]
|
||||||
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
|
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
|
||||||
exact hS
|
exact hS
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -52,7 +52,7 @@ lemma lineInPlaneCond_perm {S : (PureU1 (n)).LinSols} (hS : LineInPlaneCond S)
|
||||||
not_false_eq_true]
|
not_false_eq_true]
|
||||||
|
|
||||||
lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : LineInPlaneCond S)
|
lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : LineInPlaneCond S)
|
||||||
(h : ¬ (S.val ((Fin.last n).castSucc))^2 = (S.val ((Fin.last n).succ))^2 ) :
|
(h : ¬ (S.val ((Fin.last n).castSucc))^2 = (S.val ((Fin.last n).succ))^2) :
|
||||||
(2 - n) * S.val (Fin.last (n + 1)) =
|
(2 - n) * S.val (Fin.last (n + 1)) =
|
||||||
- (2 - n)* S.val (Fin.castSucc (Fin.last n)) := by
|
- (2 - n)* S.val (Fin.castSucc (Fin.last n)) := by
|
||||||
erw [sq_eq_sq_iff_eq_or_eq_neg] at h
|
erw [sq_eq_sq_iff_eq_or_eq_neg] at h
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -43,7 +43,7 @@ lemma cube_for_linSol (S : (PureU1 3).LinSols) :
|
||||||
lemma three_sol_zero (S : (PureU1 3).Sols) : S.val (0 : Fin 3) = 0 ∨ S.val (1 : Fin 3) = 0
|
lemma three_sol_zero (S : (PureU1 3).Sols) : S.val (0 : Fin 3) = 0 ∨ S.val (1 : Fin 3) = 0
|
||||||
∨ S.val (2 : Fin 3) = 0 := (cube_for_linSol S.1.1).mpr S.cubicSol
|
∨ S.val (2 : Fin 3) = 0 := (cube_for_linSol S.1.1).mpr S.cubicSol
|
||||||
|
|
||||||
/-- Given a `LinSol` with a charge equal to zero a `Sol`.-/
|
/-- Given a `LinSol` with a charge equal to zero a `Sol`. -/
|
||||||
def solOfLinear (S : (PureU1 3).LinSols)
|
def solOfLinear (S : (PureU1 3).LinSols)
|
||||||
(hS : S.val (0 : Fin 3) = 0 ∨ S.val (1 : Fin 3) = 0 ∨ S.val (2 : Fin 3) = 0) :
|
(hS : S.val (0 : Fin 3) = 0 ∨ S.val (1 : Fin 3) = 0 ∨ S.val (2 : Fin 3) = 0) :
|
||||||
(PureU1 3).Sols :=
|
(PureU1 3).Sols :=
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Sort
|
import HepLean.AnomalyCancellation.PureU1.Sort
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -58,12 +58,12 @@ 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
|
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
|
||||||
intro g f hS
|
intro g f hS
|
||||||
linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1 ) -
|
linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1) -
|
||||||
(lineInCubic_expand h g f hS 1 2 ) / 6
|
(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 :=
|
def LineInCubicPerm (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
|
||||||
∀ (M : (FamilyPermutations (2 * n + 1)).group ),
|
∀ (M : (FamilyPermutations (2 * n + 1)).group),
|
||||||
LineInCubic ((FamilyPermutations (2 * n + 1)).linSolRep M S)
|
LineInCubic ((FamilyPermutations (2 * n + 1)).linSolRep M S)
|
||||||
|
|
||||||
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -124,7 +124,7 @@ theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
|
||||||
rw [parameterization]
|
rw [parameterization]
|
||||||
apply ACCSystem.Sols.ext
|
apply ACCSystem.Sols.ext
|
||||||
rw [parameterizationAsLinear_val]
|
rw [parameterizationAsLinear_val]
|
||||||
change S.val = _ • ( _ • P g + _• P! f)
|
change S.val = _ • (_ • P g + _• P! f)
|
||||||
rw [anomalyFree_param _ _ hS]
|
rw [anomalyFree_param _ _ hS]
|
||||||
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
|
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
|
||||||
exact hS
|
exact hS
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Basic
|
import HepLean.AnomalyCancellation.PureU1.Basic
|
||||||
|
@ -51,7 +51,7 @@ open PureU1Charges in
|
||||||
@[simp]
|
@[simp]
|
||||||
def permCharges {n : ℕ} : Representation ℚ (PermGroup n) (PureU1 n).Charges where
|
def permCharges {n : ℕ} : Representation ℚ (PermGroup n) (PureU1 n).Charges where
|
||||||
toFun f := chargeMap f⁻¹
|
toFun f := chargeMap f⁻¹
|
||||||
map_mul' f g :=by
|
map_mul' f g := by
|
||||||
simp only [PermGroup, mul_inv_rev]
|
simp only [PermGroup, mul_inv_rev]
|
||||||
apply LinearMap.ext
|
apply LinearMap.ext
|
||||||
intro S
|
intro S
|
||||||
|
@ -214,7 +214,7 @@ lemma permTwo_fst : (permTwo hij hij').toFun i' = i := by
|
||||||
lemma permTwo_snd : (permTwo hij hij').toFun j' = j := by
|
lemma permTwo_snd : (permTwo hij hij').toFun j' = j := by
|
||||||
rw [permTwo, permOfInjection]
|
rw [permTwo, permOfInjection]
|
||||||
have ht := Equiv.extendSubtype_apply_of_mem
|
have ht := Equiv.extendSubtype_apply_of_mem
|
||||||
((permTwoInj hij' ).toEquivRange.symm.trans
|
((permTwoInj hij').toEquivRange.symm.trans
|
||||||
(permTwoInj hij).toEquivRange) j' (permTwoInj_snd hij')
|
(permTwoInj hij).toEquivRange) j' (permTwoInj_snd hij')
|
||||||
simp at ht
|
simp at ht
|
||||||
simp [ht, permTwoInj_snd_apply]
|
simp [ht, permTwoInj_snd_apply]
|
||||||
|
@ -303,16 +303,15 @@ lemma Prop_two (P : ℚ × ℚ → Prop) {S : (PureU1 n).LinSols}
|
||||||
{a b : Fin n} (hab : a ≠ b)
|
{a b : Fin n} (hab : a ≠ b)
|
||||||
(h : ∀ (f : (FamilyPermutations n).group),
|
(h : ∀ (f : (FamilyPermutations n).group),
|
||||||
P ((((FamilyPermutations n).linSolRep f S).val a),
|
P ((((FamilyPermutations n).linSolRep f S).val a),
|
||||||
(((FamilyPermutations n).linSolRep f S).val b)
|
(((FamilyPermutations n).linSolRep f S).val b))) : ∀ (i j : Fin n) (_ : i ≠ j),
|
||||||
)) : ∀ (i j : Fin n) (_ : i ≠ j),
|
|
||||||
P (S.val i, S.val j) := by
|
P (S.val i, S.val j) := by
|
||||||
intro i j hij
|
intro i j hij
|
||||||
have h1 := h (permTwo hij hab ).symm
|
have h1 := h (permTwo hij hab).symm
|
||||||
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply] at h1
|
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply] at h1
|
||||||
simp at h1
|
simp at h1
|
||||||
change P
|
change P
|
||||||
(S.val ((permTwo hij hab ).toFun a),
|
(S.val ((permTwo hij hab).toFun a),
|
||||||
S.val ((permTwo hij hab ).toFun b)) at h1
|
S.val ((permTwo hij hab).toFun b)) at h1
|
||||||
erw [permTwo_fst,permTwo_snd] at h1
|
erw [permTwo_fst,permTwo_snd] at h1
|
||||||
exact h1
|
exact h1
|
||||||
|
|
||||||
|
@ -321,9 +320,8 @@ lemma Prop_three (P : ℚ × ℚ × ℚ → Prop) {S : (PureU1 n).LinSols}
|
||||||
(h : ∀ (f : (FamilyPermutations n).group),
|
(h : ∀ (f : (FamilyPermutations n).group),
|
||||||
P ((((FamilyPermutations n).linSolRep f S).val a),(
|
P ((((FamilyPermutations n).linSolRep f S).val a),(
|
||||||
(((FamilyPermutations n).linSolRep f S).val b),
|
(((FamilyPermutations n).linSolRep f S).val b),
|
||||||
(((FamilyPermutations n).linSolRep f S).val c)
|
(((FamilyPermutations n).linSolRep f S).val c)))) : ∀ (i j k : Fin n)
|
||||||
))) : ∀ (i j k : Fin n) (_ : i ≠ j) (_ : j ≠ k) (_ : i ≠ k),
|
(_ : i ≠ j) (_ : j ≠ k) (_ : i ≠ k), P (S.val i, (S.val j, S.val k)) := by
|
||||||
P (S.val i, (S.val j, S.val k)) := by
|
|
||||||
intro i j k hij hjk hik
|
intro i j k hij hjk hik
|
||||||
have h1 := h (permThree hij hjk hik hab hbc hac).symm
|
have h1 := h (permThree hij hjk hik hab hbc hac).symm
|
||||||
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply,
|
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply,
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Permutations
|
import HepLean.AnomalyCancellation.PureU1.Permutations
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.PureU1.Sort
|
import HepLean.AnomalyCancellation.PureU1.Sort
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Tactic.FinCases
|
import Mathlib.Tactic.FinCases
|
||||||
|
@ -54,9 +54,9 @@ lemma charges_eq_toSpecies_eq (S T : (SMCharges n).Charges) :
|
||||||
apply toSpeciesEquiv.injective
|
apply toSpeciesEquiv.injective
|
||||||
exact (Set.eqOn_univ (toSpeciesEquiv S) (toSpeciesEquiv T)).mp fun ⦃x⦄ _ => h x
|
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
|
(toSpecies i) (toSpeciesEquiv.symm f) = f i := by
|
||||||
change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i= f i
|
change (toSpeciesEquiv ∘ toSpeciesEquiv.symm) _ i= f i
|
||||||
simp
|
simp
|
||||||
|
|
||||||
/-- The `Q` charges as a map `Fin n → ℚ`. -/
|
/-- The `Q` charges as a map `Fin n → ℚ`. -/
|
||||||
|
@ -272,8 +272,7 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
|
||||||
intro i
|
intro i
|
||||||
repeat erw [map_smul]
|
repeat erw [map_smul]
|
||||||
simp [HSMul.hSMul, SMul.smul]
|
simp [HSMul.hSMul, SMul.smul]
|
||||||
ring
|
ring)
|
||||||
)
|
|
||||||
(by
|
(by
|
||||||
intro S T R L
|
intro S T R L
|
||||||
simp only
|
simp only
|
||||||
|
@ -282,22 +281,19 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
|
||||||
intro i
|
intro i
|
||||||
repeat erw [map_add]
|
repeat erw [map_add]
|
||||||
simp only [ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply, Fin.isValue]
|
simp only [ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply, Fin.isValue]
|
||||||
ring
|
ring)
|
||||||
)
|
|
||||||
(by
|
(by
|
||||||
intro S T L
|
intro S T L
|
||||||
simp only [SMSpecies_numberCharges, toSpecies_apply, Fin.isValue]
|
simp only [SMSpecies_numberCharges, toSpecies_apply, Fin.isValue]
|
||||||
apply Fintype.sum_congr
|
apply Fintype.sum_congr
|
||||||
intro i
|
intro i
|
||||||
ring
|
ring)
|
||||||
)
|
|
||||||
(by
|
(by
|
||||||
intro S T L
|
intro S T L
|
||||||
simp only [SMSpecies_numberCharges, toSpecies_apply, Fin.isValue]
|
simp only [SMSpecies_numberCharges, toSpecies_apply, Fin.isValue]
|
||||||
apply Fintype.sum_congr
|
apply Fintype.sum_congr
|
||||||
intro i
|
intro i
|
||||||
ring
|
ring)
|
||||||
)
|
|
||||||
|
|
||||||
/-- The cubic acc. -/
|
/-- The cubic acc. -/
|
||||||
@[simp]
|
@[simp]
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SM.Basic
|
import HepLean.AnomalyCancellation.SM.Basic
|
||||||
|
@ -100,7 +100,7 @@ def speciesFamilyUniversial (n : ℕ) :
|
||||||
|
|
||||||
/-- The embedding of the `1`-family charges into the `n`-family charges in
|
/-- The embedding of the `1`-family charges into the `n`-family charges in
|
||||||
a universal manor. -/
|
a universal manor. -/
|
||||||
def familyUniversal ( n : ℕ) : (SMCharges 1).Charges →ₗ[ℚ] (SMCharges n).Charges :=
|
def familyUniversal (n : ℕ) : (SMCharges 1).Charges →ₗ[ℚ] (SMCharges n).Charges :=
|
||||||
chargesMapOfSpeciesMap (speciesFamilyUniversial n)
|
chargesMapOfSpeciesMap (speciesFamilyUniversial n)
|
||||||
|
|
||||||
end SM
|
end SM
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SM.Basic
|
import HepLean.AnomalyCancellation.SM.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SM.Basic
|
import HepLean.AnomalyCancellation.SM.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SM.Basic
|
import HepLean.AnomalyCancellation.SM.Basic
|
||||||
|
@ -183,7 +183,8 @@ lemma grav (S : linearParameters) :
|
||||||
|
|
||||||
end 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
|
structure linearParametersQENeqZero where
|
||||||
/-- The parameter `x`. -/
|
/-- The parameter `x`. -/
|
||||||
x : ℚ
|
x : ℚ
|
||||||
|
@ -291,7 +292,7 @@ lemma cubic_v_or_w_zero (S : linearParametersQENeqZero) (h : accCube (bijection
|
||||||
exact h2 h
|
exact h2 h
|
||||||
|
|
||||||
lemma cubic_v_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.val = 0)
|
lemma cubic_v_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.val = 0)
|
||||||
(hv : S.v = 0 ) : S.w = -1 := by
|
(hv : S.v = 0) : S.w = -1 := by
|
||||||
rw [S.cubic, hv] at h
|
rw [S.cubic, hv] at h
|
||||||
simp at h
|
simp at h
|
||||||
have h' : (S.w + 1) * (1 * S.w * S.w + (-1) * S.w + 1) = 0 := by
|
have h' : (S.w + 1) * (1 * S.w * S.w + (-1) * S.w + 1) = 0 := by
|
||||||
|
@ -309,7 +310,7 @@ lemma cubic_v_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.
|
||||||
exact eq_neg_of_add_eq_zero_left h'
|
exact eq_neg_of_add_eq_zero_left h'
|
||||||
|
|
||||||
lemma cube_w_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.val = 0)
|
lemma cube_w_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.val = 0)
|
||||||
(hw : S.w = 0 ) : S.v = -1 := by
|
(hw : S.w = 0) : S.v = -1 := by
|
||||||
rw [S.cubic, hw] at h
|
rw [S.cubic, hw] at h
|
||||||
simp at h
|
simp at h
|
||||||
have h' : (S.v + 1) * (1 * S.v * S.v + (-1) * S.v + 1) = 0 := by
|
have h' : (S.v + 1) * (1 * S.v * S.v + (-1) * S.v + 1) = 0 := by
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SM.Basic
|
import HepLean.AnomalyCancellation.SM.Basic
|
||||||
|
@ -36,7 +36,7 @@ instance : Group (PermGroup n) := Pi.group
|
||||||
/-- The image of an element of `permGroup n` under the representation on charges. -/
|
/-- The image of an element of `permGroup n` under the representation on charges. -/
|
||||||
@[simps!]
|
@[simps!]
|
||||||
def chargeMap (f : PermGroup n) : (SMCharges n).Charges →ₗ[ℚ] (SMCharges n).Charges where
|
def chargeMap (f : PermGroup n) : (SMCharges n).Charges →ₗ[ℚ] (SMCharges n).Charges where
|
||||||
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i )
|
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i)
|
||||||
map_add' _ _ := rfl
|
map_add' _ _ := rfl
|
||||||
map_smul' _ _ := rfl
|
map_smul' _ _ := rfl
|
||||||
|
|
||||||
|
@ -44,7 +44,7 @@ def chargeMap (f : PermGroup n) : (SMCharges n).Charges →ₗ[ℚ] (SMCharges n
|
||||||
@[simp]
|
@[simp]
|
||||||
def repCharges {n : ℕ} : Representation ℚ (PermGroup n) (SMCharges n).Charges where
|
def repCharges {n : ℕ} : Representation ℚ (PermGroup n) (SMCharges n).Charges where
|
||||||
toFun f := chargeMap f⁻¹
|
toFun f := chargeMap f⁻¹
|
||||||
map_mul' f g :=by
|
map_mul' f g := by
|
||||||
simp only [PermGroup, mul_inv_rev]
|
simp only [PermGroup, mul_inv_rev]
|
||||||
apply LinearMap.ext
|
apply LinearMap.ext
|
||||||
intro S
|
intro S
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Tactic.FinCases
|
import Mathlib.Tactic.FinCases
|
||||||
|
@ -31,7 +31,7 @@ namespace SMνCharges
|
||||||
variable {n : ℕ}
|
variable {n : ℕ}
|
||||||
|
|
||||||
/-- An equivalence between `(SMνCharges n).charges` and `(Fin 6 → Fin n → ℚ)`
|
/-- An equivalence between `(SMνCharges n).charges` and `(Fin 6 → Fin n → ℚ)`
|
||||||
splitting the charges into species.-/
|
splitting the charges into species. -/
|
||||||
@[simps!]
|
@[simps!]
|
||||||
def toSpeciesEquiv : (SMνCharges n).Charges ≃ (Fin 6 → Fin n → ℚ) :=
|
def toSpeciesEquiv : (SMνCharges n).Charges ≃ (Fin 6 → Fin n → ℚ) :=
|
||||||
((Equiv.curry _ _ _).symm.trans ((@finProdFinEquiv 6 n).arrowCongr (Equiv.refl ℚ))).symm
|
((Equiv.curry _ _ _).symm.trans ((@finProdFinEquiv 6 n).arrowCongr (Equiv.refl ℚ))).symm
|
||||||
|
@ -54,9 +54,9 @@ lemma charges_eq_toSpecies_eq (S T : (SMνCharges n).Charges) :
|
||||||
funext i
|
funext i
|
||||||
exact h 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
|
(toSpecies i) (toSpeciesEquiv.symm f) = f i := by
|
||||||
change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i = f i
|
change (toSpeciesEquiv ∘ toSpeciesEquiv.symm) _ i = f i
|
||||||
simp
|
simp
|
||||||
|
|
||||||
lemma toSpecies_one (S : (SMνCharges 1).Charges) (j : Fin 6) :
|
lemma toSpecies_one (S : (SMνCharges 1).Charges) (j : Fin 6) :
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.Basic
|
import HepLean.AnomalyCancellation.SMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.Basic
|
import HepLean.AnomalyCancellation.SMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.Basic
|
import HepLean.AnomalyCancellation.SMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.Ordinary.Basic
|
import HepLean.AnomalyCancellation.SMNu.Ordinary.Basic
|
||||||
|
@ -24,12 +24,11 @@ open BigOperators
|
||||||
namespace PlaneSeven
|
namespace PlaneSeven
|
||||||
|
|
||||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||||
def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
||||||
match s, i with
|
match s, i with
|
||||||
| 0, 0 => 1
|
| 0, 0 => 1
|
||||||
| 0, 1 => - 1
|
| 0, 1 => - 1
|
||||||
| _, _ => 0
|
| _, _ => 0)
|
||||||
)
|
|
||||||
|
|
||||||
lemma B₀_cubic (S T : (SM 3).Charges) : cubeTriLin B₀ S T =
|
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
|
||||||
|
@ -37,12 +36,11 @@ lemma B₀_cubic (S T : (SM 3).Charges) : cubeTriLin B₀ S T =
|
||||||
ring
|
ring
|
||||||
|
|
||||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||||
def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
||||||
match s, i with
|
match s, i with
|
||||||
| 1, 0 => 1
|
| 1, 0 => 1
|
||||||
| 1, 1 => - 1
|
| 1, 1 => - 1
|
||||||
| _, _ => 0
|
| _, _ => 0)
|
||||||
)
|
|
||||||
|
|
||||||
lemma B₁_cubic (S T : (SM 3).Charges) : cubeTriLin B₁ S T =
|
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
|
||||||
|
@ -50,12 +48,11 @@ lemma B₁_cubic (S T : (SM 3).Charges) : cubeTriLin B₁ S T =
|
||||||
ring
|
ring
|
||||||
|
|
||||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||||
def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
||||||
match s, i with
|
match s, i with
|
||||||
| 2, 0 => 1
|
| 2, 0 => 1
|
||||||
| 2, 1 => - 1
|
| 2, 1 => - 1
|
||||||
| _, _ => 0
|
| _, _ => 0)
|
||||||
)
|
|
||||||
|
|
||||||
lemma B₂_cubic (S T : (SM 3).Charges) : cubeTriLin B₂ S T =
|
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
|
||||||
|
@ -63,12 +60,11 @@ lemma B₂_cubic (S T : (SM 3).Charges) : cubeTriLin B₂ S T =
|
||||||
ring
|
ring
|
||||||
|
|
||||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||||
def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
||||||
match s, i with
|
match s, i with
|
||||||
| 3, 0 => 1
|
| 3, 0 => 1
|
||||||
| 3, 1 => - 1
|
| 3, 1 => - 1
|
||||||
| _, _ => 0
|
| _, _ => 0)
|
||||||
)
|
|
||||||
|
|
||||||
lemma B₃_cubic (S T : (SM 3).Charges) : cubeTriLin B₃ S T =
|
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
|
||||||
|
@ -77,12 +73,11 @@ lemma B₃_cubic (S T : (SM 3).Charges) : cubeTriLin B₃ S T =
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||||
def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
||||||
match s, i with
|
match s, i with
|
||||||
| 4, 0 => 1
|
| 4, 0 => 1
|
||||||
| 4, 1 => - 1
|
| 4, 1 => - 1
|
||||||
| _, _ => 0
|
| _, _ => 0)
|
||||||
)
|
|
||||||
|
|
||||||
lemma B₄_cubic (S T : (SM 3).Charges) : cubeTriLin B₄ S T =
|
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
|
||||||
|
@ -91,12 +86,11 @@ lemma B₄_cubic (S T : (SM 3).Charges) : cubeTriLin B₄ S T =
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||||
def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
||||||
match s, i with
|
match s, i with
|
||||||
| 5, 0 => 1
|
| 5, 0 => 1
|
||||||
| 5, 1 => - 1
|
| 5, 1 => - 1
|
||||||
| _, _ => 0
|
| _, _ => 0)
|
||||||
)
|
|
||||||
|
|
||||||
lemma B₅_cubic (S T : (SM 3).Charges) : cubeTriLin B₅ S T =
|
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
|
||||||
|
@ -105,12 +99,11 @@ lemma B₅_cubic (S T : (SM 3).Charges) : cubeTriLin B₅ S T =
|
||||||
rfl
|
rfl
|
||||||
|
|
||||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||||
def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
||||||
match s, i with
|
match s, i with
|
||||||
| 1, 2 => 1
|
| 1, 2 => 1
|
||||||
| 2, 2 => -1
|
| 2, 2 => -1
|
||||||
| _, _ => 0
|
| _, _ => 0)
|
||||||
)
|
|
||||||
|
|
||||||
lemma B₆_cubic (S T : (SM 3).Charges) : cubeTriLin B₆ S T =
|
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
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.Ordinary.Basic
|
import HepLean.AnomalyCancellation.SMNu.Ordinary.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.Basic
|
import HepLean.AnomalyCancellation.SMNu.Basic
|
||||||
|
@ -36,7 +36,7 @@ instance : Group (PermGroup n) := Pi.group
|
||||||
/-- The image of an element of `permGroup n` under the representation on charges. -/
|
/-- The image of an element of `permGroup n` under the representation on charges. -/
|
||||||
@[simps!]
|
@[simps!]
|
||||||
def chargeMap (f : PermGroup n) : (SMνCharges n).Charges →ₗ[ℚ] (SMνCharges n).Charges where
|
def chargeMap (f : PermGroup n) : (SMνCharges n).Charges →ₗ[ℚ] (SMνCharges n).Charges where
|
||||||
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i )
|
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i)
|
||||||
map_add' _ _ := rfl
|
map_add' _ _ := rfl
|
||||||
map_smul' _ _ := rfl
|
map_smul' _ _ := rfl
|
||||||
|
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.Basic
|
import HepLean.AnomalyCancellation.SMNu.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.StandardModel.HiggsBoson.Potential
|
import HepLean.StandardModel.HiggsBoson.Potential
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Logic.Equiv.Fin
|
import Mathlib.Logic.Equiv.Fin
|
||||||
|
@ -289,7 +289,7 @@ instance diagramVertexPropDecidable
|
||||||
@decidable_of_decidable_of_iff _ _
|
@decidable_of_decidable_of_iff _ _
|
||||||
(@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _
|
(@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _
|
||||||
(fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _
|
(fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _
|
||||||
(fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _ ) _)
|
(fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _) _)
|
||||||
(P.diagramVertexProp_iff F f).symm
|
(P.diagramVertexProp_iff F f).symm
|
||||||
|
|
||||||
instance diagramEdgePropDecidable
|
instance diagramEdgePropDecidable
|
||||||
|
@ -299,7 +299,7 @@ instance diagramEdgePropDecidable
|
||||||
@decidable_of_decidable_of_iff _ _
|
@decidable_of_decidable_of_iff _ _
|
||||||
(@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _
|
(@Fintype.decidableForallFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _
|
||||||
(fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _
|
(fun _ => @And.decidable _ _ _ (@Fintype.decidablePiFintype _ _
|
||||||
(fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _ ) _)
|
(fun _ => P.preFeynmanRuleDecEq𝓱𝓔) _ _ _)) _) _)
|
||||||
(P.diagramEdgeProp_iff F f).symm
|
(P.diagramEdgeProp_iff F f).symm
|
||||||
|
|
||||||
end PreFeynmanRule
|
end PreFeynmanRule
|
||||||
|
@ -717,7 +717,7 @@ instance [IsFiniteDiagram F] : DecidableRel F.AdjRelation := fun _ _ =>
|
||||||
@Fintype.decidableExistsFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ (
|
@Fintype.decidableExistsFintype _ _ (fun _ => @Fintype.decidableExistsFintype _ _ (
|
||||||
fun _ => @And.decidable _ _ (instDecidableEq𝓔OfIsFiniteDiagram _ _) $
|
fun _ => @And.decidable _ _ (instDecidableEq𝓔OfIsFiniteDiagram _ _) $
|
||||||
@And.decidable _ _ (instDecidableEq𝓥OfIsFiniteDiagram _ _)
|
@And.decidable _ _ (instDecidableEq𝓥OfIsFiniteDiagram _ _)
|
||||||
(instDecidableEq𝓥OfIsFiniteDiagram _ _)) _ ) _
|
(instDecidableEq𝓥OfIsFiniteDiagram _ _)) _) _
|
||||||
|
|
||||||
/-- From a Feynman diagram the simple graph showing those vertices which are connected. -/
|
/-- From a Feynman diagram the simple graph showing those vertices which are connected. -/
|
||||||
def toSimpleGraph : SimpleGraph F.𝓥 where
|
def toSimpleGraph : SimpleGraph F.𝓥 where
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FeynmanDiagrams.Basic
|
import HepLean.FeynmanDiagrams.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FeynmanDiagrams.Basic
|
import HepLean.FeynmanDiagrams.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FeynmanDiagrams.Basic
|
import HepLean.FeynmanDiagrams.Basic
|
||||||
|
@ -109,7 +109,7 @@ instance (i : Fin 2) : Module ℝ (EdgeVertexMomentaMap F i) :=
|
||||||
| 0 => instModuleRealEdgeMomenta F
|
| 0 => instModuleRealEdgeMomenta F
|
||||||
| 1 => instModuleRealVertexMomenta F
|
| 1 => instModuleRealVertexMomenta F
|
||||||
|
|
||||||
/-- The direct sum of `EdgeMomenta` and `VertexMomenta`.-/
|
/-- The direct sum of `EdgeMomenta` and `VertexMomenta`. -/
|
||||||
def EdgeVertexMomenta : Type := DirectSum (Fin 2) (EdgeVertexMomentaMap F)
|
def EdgeVertexMomenta : Type := DirectSum (Fin 2) (EdgeVertexMomentaMap F)
|
||||||
|
|
||||||
instance : AddCommGroup F.EdgeVertexMomenta := DirectSum.instAddCommGroup _
|
instance : AddCommGroup F.EdgeVertexMomenta := DirectSum.instAddCommGroup _
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.LinearAlgebra.UnitaryGroup
|
import Mathlib.LinearAlgebra.UnitaryGroup
|
||||||
|
@ -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,
|
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,
|
cons_val_two, tail_val', head_val', cons_val_one, head_cons, tail_cons, head_fin_const,
|
||||||
zero_mul, add_zero, mul_zero]
|
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]
|
simp only [Fin.isValue, zero_mul, zero_add, mul_zero, add_zero]
|
||||||
rw [exp_add]
|
rw [exp_add]
|
||||||
ring_nf
|
ring_nf
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
||||||
|
@ -227,9 +227,9 @@ lemma fstRowThdColRealCond_holds_up_to_equiv (V : CKMMatrix) :
|
||||||
obtain ⟨τ, hτ⟩ := V.uRow_cross_cRow_eq_tRow
|
obtain ⟨τ, hτ⟩ := V.uRow_cross_cRow_eq_tRow
|
||||||
let U : CKMMatrix := phaseShiftApply V
|
let U : CKMMatrix := phaseShiftApply V
|
||||||
0
|
0
|
||||||
(- τ + arg [V]ud + arg [V]us + arg [V]tb )
|
(- τ + arg [V]ud + arg [V]us + arg [V]tb)
|
||||||
(- τ + arg [V]cb + arg [V]ud + arg [V]us )
|
(- τ + arg [V]cb + arg [V]ud + arg [V]us)
|
||||||
(- arg [V]ud )
|
(- arg [V]ud)
|
||||||
(- arg [V]us)
|
(- arg [V]us)
|
||||||
(τ - arg [V]ud - arg [V]us - arg [V]cb - arg [V]tb)
|
(τ - arg [V]ud - arg [V]us - arg [V]cb - arg [V]tb)
|
||||||
have hUV : Quotient.mk CKMMatrixSetoid U = ⟦V⟧ := by
|
have hUV : Quotient.mk CKMMatrixSetoid U = ⟦V⟧ := by
|
||||||
|
@ -262,7 +262,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
|
||||||
(hV : FstRowThdColRealCond V) :
|
(hV : FstRowThdColRealCond V) :
|
||||||
∃ (U : CKMMatrix), V ≈ U ∧ ubOnePhaseCond U:= by
|
∃ (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)
|
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
|
use U
|
||||||
have hUV : Quotient.mk CKMMatrixSetoid U= ⟦V⟧ := by
|
have hUV : Quotient.mk CKMMatrixSetoid U= ⟦V⟧ := by
|
||||||
simp only [Quotient.eq]
|
simp only [Quotient.eq]
|
||||||
|
@ -318,7 +318,7 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
|
||||||
lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0)
|
||||||
(hV : FstRowThdColRealCond V) :
|
(hV : FstRowThdColRealCond V) :
|
||||||
[V]cd = (- VtbAbs ⟦V⟧ * VusAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) +
|
[V]cd = (- VtbAbs ⟦V⟧ * VusAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2)) +
|
||||||
(- VubAbs ⟦V⟧ * VudAbs ⟦V⟧ * VcbAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2 ))
|
(- VubAbs ⟦V⟧ * VudAbs ⟦V⟧ * VcbAbs ⟦V⟧ / (VudAbs ⟦V⟧ ^2 + VusAbs ⟦V⟧ ^2))
|
||||||
* cexp (- arg [V]ub * I) := by
|
* cexp (- arg [V]ub * I) := by
|
||||||
have hτ : [V]t = cexp ((0 : ℝ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := 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]
|
simp only [ofReal_zero, zero_mul, exp_zero, one_smul]
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
||||||
|
@ -84,7 +84,7 @@ lemma ud_us_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
|
||||||
simp_all
|
simp_all
|
||||||
have h1 := Complex.abs.nonneg [V]ub
|
have h1 := Complex.abs.nonneg [V]ub
|
||||||
rw [h2] at h1
|
rw [h2] at h1
|
||||||
have h2 : ¬ 0 ≤ ( -1 : ℝ) := by simp
|
have h2 : ¬ 0 ≤ (-1 : ℝ) := by simp
|
||||||
exact h2 h1
|
exact h2 h1
|
||||||
|
|
||||||
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) :
|
||||||
|
@ -100,7 +100,7 @@ lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℝ {V : CKMMatrix} (hb : [V]ud ≠ 0
|
||||||
exact hb h2
|
exact hb h2
|
||||||
have h3 := Complex.abs.nonneg [V]ub
|
have h3 := Complex.abs.nonneg [V]ub
|
||||||
rw [h2] at h3
|
rw [h2] at h3
|
||||||
have h2 : ¬ 0 ≤ ( -1 : ℝ) := by simp
|
have h2 : ¬ 0 ≤ (-1 : ℝ) := by simp
|
||||||
exact h2 h3
|
exact h2 h3
|
||||||
|
|
||||||
lemma VAbsub_neq_zero_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
lemma VAbsub_neq_zero_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
||||||
|
@ -154,11 +154,11 @@ lemma fst_row_orthog_thd_row (V : CKMMatrix) :
|
||||||
|
|
||||||
lemma Vcd_mul_conj_Vud (V : CKMMatrix) :
|
lemma Vcd_mul_conj_Vud (V : CKMMatrix) :
|
||||||
[V]cd * conj [V]ud = -[V]cs * conj [V]us - [V]cb * conj [V]ub := by
|
[V]cd * conj [V]ud = -[V]cs * conj [V]us - [V]cb * conj [V]ub := by
|
||||||
linear_combination (V.fst_row_orthog_snd_row )
|
linear_combination (V.fst_row_orthog_snd_row)
|
||||||
|
|
||||||
lemma Vcs_mul_conj_Vus (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
|
[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
|
end orthogonal
|
||||||
|
|
||||||
|
@ -344,7 +344,7 @@ lemma cb_tb_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
|
||||||
simp_all
|
simp_all
|
||||||
have h1 := Complex.abs.nonneg [V]ub
|
have h1 := Complex.abs.nonneg [V]ub
|
||||||
rw [h2] at h1
|
rw [h2] at h1
|
||||||
have h2 : ¬ 0 ≤ ( -1 : ℝ) := by simp
|
have h2 : ¬ 0 ≤ (-1 : ℝ) := by simp
|
||||||
exact h2 h1
|
exact h2 h1
|
||||||
|
|
||||||
lemma VAbs_fst_col_eq_one_snd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3}
|
lemma VAbs_fst_col_eq_one_snd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3}
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
||||||
|
@ -240,7 +240,7 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
|
||||||
lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
|
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)
|
obtain ⟨g, hg⟩ := (mem_span_range_iff_exists_fun ℂ).mp (Basis.mem_span (rowBasis V)
|
||||||
(conj ([V]u) ×₃ conj ([V]c)) )
|
(conj ([V]u) ×₃ conj ([V]c)))
|
||||||
rw [Fin.sum_univ_three, rowBasis] at hg
|
rw [Fin.sum_univ_three, rowBasis] at hg
|
||||||
simp at hg
|
simp at hg
|
||||||
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
||||||
|
@ -307,20 +307,20 @@ def ucCross : Fin 3 → ℂ :=
|
||||||
conj [phaseShiftApply V a b c d e f]u ×₃ conj [phaseShiftApply V a b c d e f]c
|
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 =
|
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,
|
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,
|
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]
|
exp_add, exp_sub, ← exp_conj, conj_ofReal]
|
||||||
ring
|
ring
|
||||||
|
|
||||||
lemma ucCross_snd (V : CKMMatrix) : (ucCross V a b c d e f) 1 =
|
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,
|
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add,
|
||||||
exp_sub, ← exp_conj, conj_ofReal]
|
exp_sub, ← exp_conj, conj_ofReal]
|
||||||
ring
|
ring
|
||||||
|
|
||||||
lemma ucCross_thd (V : CKMMatrix) : (ucCross V a b c d e f) 2 =
|
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,
|
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add, exp_sub,
|
||||||
← exp_conj, conj_ofReal]
|
← exp_conj, conj_ofReal]
|
||||||
ring
|
ring
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
import HepLean.FlavorPhysics.CKMMatrix.Basic
|
||||||
|
@ -373,8 +373,8 @@ lemma mulExpδ₁₃_on_param_abs (V : CKMMatrix) (δ₁₃ : ℝ) :
|
||||||
complexAbs_sin_θ₂₃, complexAbs_cos_θ₂₃]
|
complexAbs_sin_θ₂₃, complexAbs_cos_θ₂₃]
|
||||||
|
|
||||||
lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : ℝ)
|
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 (arg (mulExpδ₁₃ ⟦standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃⟧) * I) =
|
||||||
cexp (δ₁₃ * I) := by
|
cexp (δ₁₃ * I) := by
|
||||||
have h1a := mulExpδ₁₃_on_param_δ₁₃ V δ₁₃
|
have h1a := mulExpδ₁₃_on_param_δ₁₃ V δ₁₃
|
||||||
have habs := mulExpδ₁₃_on_param_abs V δ₁₃
|
have habs := mulExpδ₁₃_on_param_abs V δ₁₃
|
||||||
|
@ -383,7 +383,7 @@ lemma mulExpδ₁₃_on_param_neq_zero_arg (V : CKMMatrix) (δ₁₃ : ℝ)
|
||||||
rw [habs, h1a]
|
rw [habs, h1a]
|
||||||
ring_nf
|
ring_nf
|
||||||
nth_rewrite 1 [← abs_mul_exp_arg_mul_I (mulExpδ₁₃
|
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 :
|
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]
|
simp only [ne_eq, ofReal_eq_zero, map_eq_zero]
|
||||||
|
@ -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
|
rw [ud_us_neq_zero_iff_ub_neq_one] at hb
|
||||||
simp [VAbs, hb]
|
simp [VAbs, hb]
|
||||||
have h1 : ofReal (√(VAbs 0 0 ⟦V⟧ ^ 2 + VAbs 0 1 ⟦V⟧ ^ 2) *
|
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 ]
|
rw [Real.mul_self_sqrt ]
|
||||||
apply add_nonneg (sq_nonneg _) (sq_nonneg _)
|
apply add_nonneg (sq_nonneg _) (sq_nonneg _)
|
||||||
simp at h1
|
simp at h1
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Tactic.Polyrith
|
import Mathlib.Tactic.Polyrith
|
||||||
|
@ -70,7 +70,7 @@ def mk₂ (f : V × V → ℚ) (map_smul : ∀ a S T, f (a • S, T) = a * f (S,
|
||||||
simp only
|
simp only
|
||||||
rw [swap, map_add]
|
rw [swap, map_add]
|
||||||
exact Mathlib.Tactic.LinearCombination.add_pf (swap T1 S) (swap T2 S)
|
exact Mathlib.Tactic.LinearCombination.add_pf (swap T1 S) (swap T2 S)
|
||||||
map_smul' :=by
|
map_smul' := by
|
||||||
intro a T
|
intro a T
|
||||||
simp only [eq_ratCast, Rat.cast_eq_id, id_eq, smul_eq_mul]
|
simp only [eq_ratCast, Rat.cast_eq_id, id_eq, smul_eq_mul]
|
||||||
rw [swap, map_smul]
|
rw [swap, map_smul]
|
||||||
|
@ -173,7 +173,7 @@ namespace TriLinearSymm
|
||||||
open BigOperators
|
open BigOperators
|
||||||
variable {V : Type} [AddCommMonoid V] [Module ℚ V]
|
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 f := f.toFun
|
||||||
coe_injective' f g h := by
|
coe_injective' f g h := by
|
||||||
cases f
|
cases f
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.LinearAlgebra.UnitaryGroup
|
import Mathlib.LinearAlgebra.UnitaryGroup
|
||||||
|
@ -183,12 +183,12 @@ def toEnd (A : SO(3)) : End ℝ (EuclideanSpace ℝ (Fin 3)) :=
|
||||||
|
|
||||||
lemma one_is_eigenvalue (A : SO(3)) : A.toEnd.HasEigenvalue 1 := by
|
lemma one_is_eigenvalue (A : SO(3)) : A.toEnd.HasEigenvalue 1 := by
|
||||||
rw [End.hasEigenvalue_iff_mem_spectrum]
|
rw [End.hasEigenvalue_iff_mem_spectrum]
|
||||||
erw [AlgEquiv.spectrum_eq (Matrix.toLinAlgEquiv (EuclideanSpace.basisFun (Fin 3) ℝ).toBasis ) A.1]
|
erw [AlgEquiv.spectrum_eq (Matrix.toLinAlgEquiv (EuclideanSpace.basisFun (Fin 3) ℝ).toBasis) A.1]
|
||||||
exact one_in_spectrum A
|
exact one_in_spectrum A
|
||||||
|
|
||||||
lemma exists_stationary_vec (A : SO(3)) :
|
lemma exists_stationary_vec (A : SO(3)) :
|
||||||
∃ (v : EuclideanSpace ℝ (Fin 3)),
|
∃ (v : EuclideanSpace ℝ (Fin 3)),
|
||||||
Orthonormal ℝ (({0} : Set (Fin 3)).restrict (fun _ => v ))
|
Orthonormal ℝ (({0} : Set (Fin 3)).restrict (fun _ => v))
|
||||||
∧ A.toEnd v = v := by
|
∧ A.toEnd v = v := by
|
||||||
obtain ⟨v, hv⟩ := End.HasEigenvalue.exists_hasEigenvector $ one_is_eigenvalue A
|
obtain ⟨v, hv⟩ := End.HasEigenvalue.exists_hasEigenvector $ one_is_eigenvalue A
|
||||||
have hvn : ‖v‖ ≠ 0 := norm_ne_zero_iff.mpr hv.2
|
have hvn : ‖v‖ ≠ 0 := norm_ne_zero_iff.mpr hv.2
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Data.Complex.Exponential
|
import Mathlib.Data.Complex.Exponential
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Analysis.Complex.Basic
|
import Mathlib.Analysis.Complex.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.MinkowskiMetric
|
import HepLean.SpaceTime.MinkowskiMetric
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzAlgebra.Basic
|
import HepLean.SpaceTime.LorentzAlgebra.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.MinkowskiMetric
|
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!]
|
@[simps!]
|
||||||
def toNormOneLorentzVector (Λ : LorentzGroup d) : NormOneLorentzVector d :=
|
def toNormOneLorentzVector (Λ : LorentzGroup d) : NormOneLorentzVector d :=
|
||||||
⟨Λ.1 *ᵥ timeVec, by rw [NormOneLorentzVector.mem_iff, Λ.2, minkowskiMetric.on_timeVec]⟩
|
⟨Λ.1 *ᵥ timeVec, by rw [NormOneLorentzVector.mem_iff, Λ.2, minkowskiMetric.on_timeVec]⟩
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzGroup.Proper
|
import HepLean.SpaceTime.LorentzGroup.Proper
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzVector.NormOne
|
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'
|
rw [IsOrthochronous_iff_futurePointing] at h h'
|
||||||
exact NormOneLorentzVector.FuturePointing.metric_reflect_not_mem_mem 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
|
def orthchroRep : LorentzGroup d →* ℤ₂ where
|
||||||
toFun := orthchroMap
|
toFun := orthchroMap
|
||||||
map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous])
|
map_one' := orthchroMap_IsOrthochronous (by simp [IsOrthochronous])
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzGroup.Basic
|
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]
|
rw [detRep_apply, detRep_apply, detContinuous_eq_iff_det_eq]
|
||||||
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
|
simp only [IsProper, lorentzGroupIsGroup_one_coe, det_one]
|
||||||
|
|
||||||
lemma id_IsProper : (@IsProper d) 1 := by
|
lemma id_IsProper : @IsProper d 1 := by
|
||||||
simp [IsProper]
|
simp [IsProper]
|
||||||
|
|
||||||
lemma isProper_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) :
|
lemma isProper_on_connected_component {Λ Λ' : LorentzGroup d} (h : Λ' ∈ connectedComponent Λ) :
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
/-!
|
/-!
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzGroup.Basic
|
import HepLean.SpaceTime.LorentzGroup.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.MinkowskiMetric
|
import HepLean.SpaceTime.MinkowskiMetric
|
||||||
|
@ -75,12 +75,12 @@ noncomputable def toSelfAdjointMatrix :
|
||||||
ext i j
|
ext i j
|
||||||
fin_cases i <;> fin_cases j <;>
|
fin_cases i <;> fin_cases j <;>
|
||||||
field_simp [fromSelfAdjointMatrix', toMatrix, conj_ofReal]
|
field_simp [fromSelfAdjointMatrix', toMatrix, conj_ofReal]
|
||||||
exact conj_eq_iff_re.mp (congrArg (fun M => M 0 0) $ selfAdjoint.mem_iff.mp x.2 )
|
exact conj_eq_iff_re.mp (congrArg (fun M => M 0 0) $ selfAdjoint.mem_iff.mp x.2)
|
||||||
have h01 := congrArg (fun M => M 0 1) $ selfAdjoint.mem_iff.mp x.2
|
have h01 := congrArg (fun M => M 0 1) $ selfAdjoint.mem_iff.mp x.2
|
||||||
simp only [Fin.isValue, star_apply, RCLike.star_def] at h01
|
simp only [Fin.isValue, star_apply, RCLike.star_def] at h01
|
||||||
rw [← h01, RCLike.conj_eq_re_sub_im]
|
rw [← h01, RCLike.conj_eq_re_sub_im]
|
||||||
rfl
|
rfl
|
||||||
exact conj_eq_iff_re.mp (congrArg (fun M => M 1 1) $ selfAdjoint.mem_iff.mp x.2 )
|
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
|
ext i j : 2
|
||||||
simp only [toSelfAdjointMatrix'_coe, add_apply, ofReal_add, of_apply, cons_val', empty_val',
|
simp only [toSelfAdjointMatrix'_coe, add_apply, ofReal_add, of_apply, cons_val', empty_val',
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Data.Complex.Exponential
|
import Mathlib.Data.Complex.Exponential
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzVector.Basic
|
import HepLean.SpaceTime.LorentzVector.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzVector.Basic
|
import HepLean.SpaceTime.LorentzVector.Basic
|
||||||
|
@ -190,7 +190,7 @@ lemma self_spaceReflection_eq_zero_iff : ⟪v, v.spaceReflection⟫ₘ = 0 ↔ v
|
||||||
/-- The metric tensor is non-degenerate. -/
|
/-- The metric tensor is non-degenerate. -/
|
||||||
lemma nondegenerate : (∀ w, ⟪w, v⟫ₘ = 0) ↔ v = 0 := by
|
lemma nondegenerate : (∀ w, ⟪w, v⟫ₘ = 0) ↔ v = 0 := by
|
||||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||||
· exact (self_spaceReflection_eq_zero_iff _ ).mp ((symm _ _).trans $ h v.spaceReflection)
|
· exact (self_spaceReflection_eq_zero_iff _).mp ((symm _ _).trans $ h v.spaceReflection)
|
||||||
· simp [h]
|
· simp [h]
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.LorentzGroup.Basic
|
import HepLean.SpaceTime.LorentzGroup.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Data.Complex.Exponential
|
import Mathlib.Data.Complex.Exponential
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.SpaceTime.Basic
|
import HepLean.SpaceTime.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.StandardModel.HiggsBoson.Basic
|
import HepLean.StandardModel.HiggsBoson.Basic
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import HepLean.StandardModel.HiggsBoson.Basic
|
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
|
/-- Given a `HiggsField`, the map `SpaceTime → ℝ` obtained by taking the square norm of the
|
||||||
pointwise Higgs vector. -/
|
pointwise Higgs vector. -/
|
||||||
@[simp]
|
@[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. -/
|
/-- Notation for the norm squared of a Higgs field. -/
|
||||||
scoped[StandardModel.HiggsField] notation "‖" φ1 "‖_H ^ 2" => normSq φ1
|
scoped[StandardModel.HiggsField] notation "‖" φ1 "‖_H ^ 2" => normSq φ1
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Algebra.QuadraticDiscriminant
|
import Mathlib.Algebra.QuadraticDiscriminant
|
||||||
|
@ -34,7 +34,7 @@ open SpaceTime
|
||||||
|
|
||||||
/-- The Higgs potential of the form `- μ² * |φ|² + 𝓵 * |φ|⁴`. -/
|
/-- The Higgs potential of the form `- μ² * |φ|² + 𝓵 * |φ|⁴`. -/
|
||||||
@[simp]
|
@[simp]
|
||||||
def potential (μ2 𝓵 : ℝ ) (φ : HiggsField) (x : SpaceTime) : ℝ :=
|
def potential (μ2 𝓵 : ℝ) (φ : HiggsField) (x : SpaceTime) : ℝ :=
|
||||||
- μ2 * ‖φ‖_H ^ 2 x + 𝓵 * ‖φ‖_H ^ 2 x * ‖φ‖_H ^ 2 x
|
- μ2 * ‖φ‖_H ^ 2 x + 𝓵 * ‖φ‖_H ^ 2 x * ‖φ‖_H ^ 2 x
|
||||||
|
|
||||||
/-!
|
/-!
|
||||||
|
@ -94,7 +94,7 @@ lemma snd_term_nonneg (φ : HiggsField) (x : SpaceTime) :
|
||||||
and_self]
|
and_self]
|
||||||
|
|
||||||
lemma as_quad (μ2 𝓵 : ℝ) (φ : HiggsField) (x : SpaceTime) :
|
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]
|
simp only [normSq, neg_mul, potential, neg_add_rev, neg_neg]
|
||||||
ring
|
ring
|
||||||
|
|
||||||
|
@ -237,7 +237,7 @@ lemma IsMinOn_iff_of_μSq_nonpos {μ2 : ℝ} (hμ2 : μ2 ≤ 0) :
|
||||||
have h0 := isMinOn_univ_iff.mp h 0
|
have h0 := isMinOn_univ_iff.mp h 0
|
||||||
have h1 := bounded_below_of_μSq_nonpos h𝓵 hμ2 φ x
|
have h1 := bounded_below_of_μSq_nonpos h𝓵 hμ2 φ x
|
||||||
simp only at h0
|
simp only at h0
|
||||||
rw [(eq_bound_iff_of_μSq_nonpos h𝓵 hμ2 0 0 ).mpr (by rfl)] at h0
|
rw [(eq_bound_iff_of_μSq_nonpos h𝓵 hμ2 0 0).mpr (by rfl)] at h0
|
||||||
exact (Real.partialOrder.le_antisymm _ _ h1 h0).symm
|
exact (Real.partialOrder.le_antisymm _ _ h1 h0).symm
|
||||||
· exact eq_bound_IsMinOn_of_μSq_nonpos h𝓵 hμ2 φ x
|
· exact eq_bound_IsMinOn_of_μSq_nonpos h𝓵 hμ2 φ x
|
||||||
|
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
/-
|
/-
|
||||||
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
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
|
Authors: Joseph Tooby-Smith
|
||||||
-/
|
-/
|
||||||
import Mathlib.Data.Complex.Exponential
|
import Mathlib.Data.Complex.Exponential
|
||||||
|
|
|
@ -26,7 +26,7 @@ name = "HepLean"
|
||||||
#[[require]]
|
#[[require]]
|
||||||
#name = "LeanCopilot"
|
#name = "LeanCopilot"
|
||||||
#git = "https://github.com/lean-dojo/LeanCopilot.git"
|
#git = "https://github.com/lean-dojo/LeanCopilot.git"
|
||||||
#rev = "v1.2.2"
|
#rev = "v1.4.1"
|
||||||
|
|
||||||
# lean_exe commands defined specifically for HepLean.
|
# lean_exe commands defined specifically for HepLean.
|
||||||
|
|
||||||
|
@ -45,3 +45,7 @@ srcDir = "scripts"
|
||||||
[[lean_exe]]
|
[[lean_exe]]
|
||||||
name = "find_TODOs"
|
name = "find_TODOs"
|
||||||
srcDir = "scripts"
|
srcDir = "scripts"
|
||||||
|
|
||||||
|
[[lean_exe]]
|
||||||
|
name = "mathlib_textLint_on_hepLean"
|
||||||
|
srcDir = "scripts"
|
|
@ -46,6 +46,19 @@ def doubleSpaceLinter : HepLeanTextLinter := fun lines ↦ Id.run do
|
||||||
else none)
|
else none)
|
||||||
errors.toArray
|
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
|
structure HepLeanErrorContext where
|
||||||
/-- The underlying `message`. -/
|
/-- The underlying `message`. -/
|
||||||
error : String
|
error : String
|
||||||
|
@ -64,7 +77,8 @@ def hepLeanLintFile (path : FilePath) : IO Bool := do
|
||||||
let lines ← IO.FS.lines path
|
let lines ← IO.FS.lines path
|
||||||
let allOutput := (Array.map (fun lint ↦
|
let allOutput := (Array.map (fun lint ↦
|
||||||
(Array.map (fun (e, n, c) ↦ HepLeanErrorContext.mk e n c path)) (lint lines)))
|
(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
|
let errors := allOutput.flatten
|
||||||
printErrors errors
|
printErrors errors
|
||||||
return errors.size > 0
|
return errors.size > 0
|
||||||
|
@ -72,7 +86,7 @@ def hepLeanLintFile (path : FilePath) : IO Bool := do
|
||||||
def main (_ : List String) : IO UInt32 := do
|
def main (_ : List String) : IO UInt32 := do
|
||||||
initSearchPath (← findSysroot)
|
initSearchPath (← findSysroot)
|
||||||
let mods : Name := `HepLean
|
let mods : Name := `HepLean
|
||||||
let imp : Import := {module := mods}
|
let imp : Import := {module := mods}
|
||||||
let mFile ← findOLean imp.module
|
let mFile ← findOLean imp.module
|
||||||
unless (← mFile.pathExists) do
|
unless (← mFile.pathExists) do
|
||||||
throw <| IO.userError s!"object file '{mFile}' of module {imp.module} does not exist"
|
throw <| IO.userError s!"object file '{mFile}' of module {imp.module} does not exist"
|
||||||
|
@ -85,4 +99,6 @@ def main (_ : List String) : IO UInt32 := do
|
||||||
warned := true
|
warned := true
|
||||||
if warned then
|
if warned then
|
||||||
throw <| IO.userError s!"Errors found."
|
throw <| IO.userError s!"Errors found."
|
||||||
|
else
|
||||||
|
IO.println "No linting issues found."
|
||||||
return 0
|
return 0
|
||||||
|
|
48
scripts/mathlib_textLint_on_hepLean.lean
Normal file
48
scripts/mathlib_textLint_on_hepLean.lean
Normal 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
|
0
scripts/nolints-style.txt
Normal file
0
scripts/nolints-style.txt
Normal file
Loading…
Add table
Add a link
Reference in a new issue