reactor: Removal of double spaces
This commit is contained in:
parent
ce92e1d649
commit
13f62a50eb
64 changed files with 550 additions and 546 deletions
|
@ -197,7 +197,7 @@ instance quadSolsMulAction (χ : ACCSystemQuad) : MulAction ℚ χ.QuadSols wher
|
|||
exact one_smul _ _
|
||||
|
||||
/-- The inclusion of quadratic solutions into linear solutions. -/
|
||||
def quadSolsInclLinSols (χ : ACCSystemQuad) : χ.QuadSols →[ℚ] χ.LinSols where
|
||||
def quadSolsInclLinSols (χ : ACCSystemQuad) : χ.QuadSols →[ℚ] χ.LinSols where
|
||||
toFun := QuadSols.toLinSols
|
||||
map_smul' _ _ := rfl
|
||||
|
||||
|
@ -229,7 +229,7 @@ structure Sols (χ : ACCSystem) extends χ.QuadSols where
|
|||
/-- Two solutions are equal if the underlying charges are equal. -/
|
||||
lemma Sols.ext {χ : ACCSystem} {S T : χ.Sols} (h : S.val = T.val) :
|
||||
S = T := by
|
||||
have h := ACCSystemQuad.QuadSols.ext h
|
||||
have h := ACCSystemQuad.QuadSols.ext h
|
||||
cases' S
|
||||
simp_all only
|
||||
|
||||
|
@ -251,8 +251,8 @@ instance solsMulAction (χ : ACCSystem) : MulAction ℚ χ.Sols where
|
|||
exact one_smul _ _
|
||||
|
||||
/-- The inclusion of `Sols` into `QuadSols`. -/
|
||||
def solsInclQuadSols (χ : ACCSystem) : χ.Sols →[ℚ] χ.QuadSols where
|
||||
toFun := Sols.toQuadSols
|
||||
def solsInclQuadSols (χ : ACCSystem) : χ.Sols →[ℚ] χ.QuadSols where
|
||||
toFun := Sols.toQuadSols
|
||||
map_smul' _ _ := rfl
|
||||
|
||||
/-- The inclusion of `Sols` into `LinSols`. -/
|
||||
|
|
|
@ -104,7 +104,7 @@ lemma linSolRep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction
|
|||
G.linSolRep g (χ.quadSolsInclLinSols S) := rfl
|
||||
|
||||
lemma rep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
(S : χ.QuadSols) : χ.quadSolsIncl (G.quadSolAction.toFun S g) =
|
||||
(S : χ.QuadSols) : χ.quadSolsIncl (G.quadSolAction.toFun S g) =
|
||||
G.rep g (χ.quadSolsIncl S) := rfl
|
||||
|
||||
/-- The group action acting on solutions to the anomaly cancellation conditions. -/
|
||||
|
@ -126,7 +126,7 @@ instance solAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) : MulAction G.
|
|||
rfl
|
||||
|
||||
lemma quadSolAction_solAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
(S : χ.Sols) : χ.solsInclQuadSols (G.solAction.toFun S g) =
|
||||
(S : χ.Sols) : χ.solsInclQuadSols (G.solAction.toFun S g) =
|
||||
G.quadSolAction.toFun (χ.solsInclQuadSols S) g := rfl
|
||||
|
||||
lemma linSolRep_solAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group)
|
||||
|
|
|
@ -34,7 +34,7 @@ def B₃AsCharge : MSSMACC.Charges := toSpecies.symm
|
|||
| 0, 2 => - 1
|
||||
| 1, 0 => - 1
|
||||
| 1, 1 => - 1
|
||||
| 1, 2 => 1
|
||||
| 1, 2 => 1
|
||||
| 2, 0 => - 1
|
||||
| 2, 1 => - 1
|
||||
| 2, 2 => 1
|
||||
|
|
|
@ -43,7 +43,7 @@ def toSMPlusH : MSSMCharges.Charges ≃ (Fin 18 ⊕ Fin 2 → ℚ) :=
|
|||
@[simps!]
|
||||
def splitSMPlusH : (Fin 18 ⊕ Fin 2 → ℚ) ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) where
|
||||
toFun f := (f ∘ Sum.inl , f ∘ Sum.inr)
|
||||
invFun f := Sum.elim f.1 f.2
|
||||
invFun f := Sum.elim f.1 f.2
|
||||
left_inv f := Sum.elim_comp_inl_inr f
|
||||
right_inv _ := rfl
|
||||
|
||||
|
@ -107,7 +107,7 @@ def Hu : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
map_smul' _ _ := by rfl
|
||||
|
||||
lemma charges_eq_toSpecies_eq (S T : MSSMCharges.Charges) :
|
||||
S = T ↔ (∀ i, toSMSpecies i S = toSMSpecies i T) ∧ Hd S = Hd T ∧ Hu S = Hu T := by
|
||||
S = T ↔ (∀ i, toSMSpecies i S = toSMSpecies i T) ∧ Hd S = Hd T ∧ Hu S = Hu T := by
|
||||
apply Iff.intro
|
||||
intro h
|
||||
rw [h]
|
||||
|
@ -159,7 +159,7 @@ def accGrav : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accGrav`. -/
|
||||
lemma accGrav_ext {S T : MSSMCharges.Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
|
||||
(hd : Hd S = Hd T) (hu : Hu S = Hu T) :
|
||||
accGrav S = accGrav T := by
|
||||
simp only [accGrav, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue,
|
||||
|
@ -192,7 +192,7 @@ def accSU2 : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accSU2`. -/
|
||||
lemma accSU2_ext {S T : MSSMCharges.Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
|
||||
(hd : Hd S = Hd T) (hu : Hu S = Hu T) :
|
||||
accSU2 S = accSU2 T := by
|
||||
simp only [accSU2, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue,
|
||||
|
@ -224,7 +224,7 @@ def accSU3 : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accSU3`. -/
|
||||
lemma accSU3_ext {S T : MSSMCharges.Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i) :
|
||||
accSU3 S = accSU3 T := by
|
||||
simp only [accSU3, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, LinearMap.coe_mk,
|
||||
AddHom.coe_mk]
|
||||
|
@ -256,7 +256,7 @@ def accYY : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accGrav`. -/
|
||||
lemma accYY_ext {S T : MSSMCharges.Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSMSpecies j) S i = ∑ i, (toSMSpecies j) T i)
|
||||
(hd : Hd S = Hd T) (hu : Hu S = Hu T) :
|
||||
accYY S = accYY T := by
|
||||
simp only [accYY, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue,
|
||||
|
@ -269,9 +269,9 @@ lemma accYY_ext {S T : MSSMCharges.Charges}
|
|||
|
||||
/-- The symmetric bilinear function used to define the quadratic ACC. -/
|
||||
@[simps!]
|
||||
def quadBiLin : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ (
|
||||
fun (S, T) => ∑ i, (Q S i * Q T i + (- 2) * (U S i * U T i) +
|
||||
D S i * D T i + (- 1) * (L S i * L T i) + E S i * E T i) +
|
||||
def quadBiLin : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ (
|
||||
fun (S, T) => ∑ i, (Q S i * Q T i + (- 2) * (U S i * U T i) +
|
||||
D S i * D T i + (- 1) * (L S i * L T i) + E S i * E T i) +
|
||||
(- Hd S * Hd T + Hu S * Hu T))
|
||||
(by
|
||||
intro a S T
|
||||
|
@ -346,7 +346,7 @@ def cubeTriLinToFun
|
|||
+ (2 * Hd S.1 * Hd S.2.1 * Hd S.2.2
|
||||
+ 2 * Hu S.1 * Hu S.2.1 * Hu S.2.2)
|
||||
|
||||
lemma cubeTriLinToFun_map_smul₁ (a : ℚ) (S T R : MSSMCharges.Charges) :
|
||||
lemma cubeTriLinToFun_map_smul₁ (a : ℚ) (S T R : MSSMCharges.Charges) :
|
||||
cubeTriLinToFun (a • S, T, R) = a * cubeTriLinToFun (S, T, R) := by
|
||||
simp only [cubeTriLinToFun]
|
||||
rw [mul_add]
|
||||
|
@ -366,7 +366,7 @@ lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.Charges) :
|
|||
rw [add_assoc, ← add_assoc (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _ _]
|
||||
rw [add_comm (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _]
|
||||
rw [add_assoc]
|
||||
rw [← add_assoc _ _ (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L +
|
||||
rw [← add_assoc _ _ (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L +
|
||||
(2 * Hd T * Hd R * Hd L + 2 * Hu T * Hu R * Hu L))]
|
||||
congr 1
|
||||
rw [← Finset.sum_add_distrib]
|
||||
|
@ -415,7 +415,7 @@ def accCube : HomogeneousCubic MSSMCharges.Charges := cubeTriLin.toCubic
|
|||
lemma accCube_ext {S T : MSSMCharges.Charges}
|
||||
(h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSMSpecies j S) i =
|
||||
∑ i, ((fun a => a^3) ∘ toSMSpecies j T) i)
|
||||
(hd : Hd S = Hd T) (hu : Hu S = Hu T) :
|
||||
(hd : Hd S = Hd T) (hu : Hu S = Hu T) :
|
||||
accCube S = accCube T := by
|
||||
simp only [HomogeneousCubic, accCube, cubeTriLin, TriLinearSymm.toCubic_apply,
|
||||
TriLinearSymm.mk₃_toFun_apply_apply, cubeTriLinToFun]
|
||||
|
@ -452,7 +452,7 @@ def MSSMACC : ACCSystem where
|
|||
namespace MSSMACC
|
||||
open MSSMCharges
|
||||
|
||||
lemma quadSol (S : MSSMACC.QuadSols) : accQuad S.val = 0 := by
|
||||
lemma quadSol (S : MSSMACC.QuadSols) : accQuad S.val = 0 := by
|
||||
have hS := S.quadSol
|
||||
simp only [MSSMACC_numberQuadratic, HomogeneousQuadratic, MSSMACC_quadraticACCs] at hS
|
||||
exact hS 0
|
||||
|
@ -516,9 +516,9 @@ lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols)
|
|||
|
||||
/-- The dot product on the vector space of charges. -/
|
||||
@[simps!]
|
||||
def dot : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂
|
||||
(fun S => ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i +
|
||||
D S.1 i * D S.2 i + L S.1 i * L S.2 i + E S.1 i * E S.2 i
|
||||
def dot : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂
|
||||
(fun S => ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i +
|
||||
D S.1 i * D S.2 i + L S.1 i * L S.2 i + E S.1 i * E S.2 i
|
||||
+ N S.1 i * N S.2 i) + Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2)
|
||||
(by
|
||||
intro a S T
|
||||
|
|
|
@ -11,7 +11,7 @@ import Mathlib.Tactic.Polyrith
|
|||
# The line through B₃ and Y₃
|
||||
|
||||
We give properties of lines through `B₃` and `Y₃`. We show that every point on this line
|
||||
is a solution to the quadratic `lineY₃B₃Charges_quad` and a double point of the cubic
|
||||
is a solution to the quadratic `lineY₃B₃Charges_quad` and a double point of the cubic
|
||||
`lineY₃B₃_doublePoint`.
|
||||
|
||||
# References
|
||||
|
@ -52,7 +52,7 @@ lemma lineY₃B₃Charges_cubic (a b : ℚ) : accCube (lineY₃B₃Charges a b).
|
|||
rw [cubeTriLin.toCubic.map_smul]
|
||||
rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₃]
|
||||
rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₃]
|
||||
erw [Y₃.cubicSol, B₃.cubicSol]
|
||||
erw [Y₃.cubicSol, B₃.cubicSol]
|
||||
rw [show cubeTriLin Y₃.val Y₃.val B₃.val = 0 by rfl]
|
||||
rw [show cubeTriLin B₃.val B₃.val Y₃.val = 0 by rfl]
|
||||
simp
|
||||
|
|
|
@ -158,7 +158,7 @@ lemma cube_proj_proj_B₃ (T : MSSMACC.LinSols) :
|
|||
lemma cube_proj_proj_self (T : MSSMACC.Sols) :
|
||||
cubeTriLin (proj T.1.1).val (proj T.1.1).val T.val =
|
||||
2 * dot Y₃.val B₃.val *
|
||||
((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin T.val T.val Y₃.val +
|
||||
((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin T.val T.val Y₃.val +
|
||||
( dot Y₃.val T.val- 2 * dot B₃.val T.val) * cubeTriLin T.val T.val B₃.val) := by
|
||||
rw [proj_val]
|
||||
rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂]
|
||||
|
@ -168,13 +168,13 @@ lemma cube_proj_proj_self (T : MSSMACC.Sols) :
|
|||
repeat rw [cubeTriLin.map_add₂]
|
||||
repeat rw [cubeTriLin.map_smul₂]
|
||||
erw [T.cubicSol]
|
||||
rw [cubeTriLin.swap₁ Y₃.val T.val T.val, cubeTriLin.swap₂ T.val Y₃.val T.val]
|
||||
rw [cubeTriLin.swap₁ B₃.val T.val T.val, cubeTriLin.swap₂ T.val B₃.val T.val]
|
||||
rw [cubeTriLin.swap₁ Y₃.val T.val T.val, cubeTriLin.swap₂ T.val Y₃.val T.val]
|
||||
rw [cubeTriLin.swap₁ B₃.val T.val T.val, cubeTriLin.swap₂ T.val B₃.val T.val]
|
||||
ring
|
||||
|
||||
lemma cube_proj (T : MSSMACC.Sols) :
|
||||
cubeTriLin (proj T.1.1).val (proj T.1.1).val (proj T.1.1).val =
|
||||
3 * dot Y₃.val B₃.val ^ 2 *
|
||||
3 * dot Y₃.val B₃.val ^ 2 *
|
||||
((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin T.val T.val Y₃.val +
|
||||
(dot Y₃.val T.val - 2 * dot B₃.val T.val) * cubeTriLin T.val T.val B₃.val) := by
|
||||
nth_rewrite 3 [proj_val]
|
||||
|
|
|
@ -42,7 +42,7 @@ lemma planeY₃B₃_smul (R : MSSMACC.AnomalyFreePerp) (a b c d : ℚ) :
|
|||
rw [smul_add, smul_add]
|
||||
rw [smul_smul, smul_smul, smul_smul]
|
||||
|
||||
lemma planeY₃B₃_eq (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (h : a = a' ∧ b = b' ∧ c = c') :
|
||||
lemma planeY₃B₃_eq (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (h : a = a' ∧ b = b' ∧ c = c') :
|
||||
(planeY₃B₃ R a b c) = (planeY₃B₃ R a' b' c') := by
|
||||
rw [h.1, h.2.1, h.2.2]
|
||||
|
||||
|
@ -82,7 +82,7 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (hR' : R
|
|||
have h2 := congrArg (fun S => S i) h1i
|
||||
change _ = 0 at h2
|
||||
simp [HSMul.hSMul] at h2
|
||||
have hc : c + -c' = 0 := by
|
||||
have hc : c + -c' = 0 := by
|
||||
cases h2 <;> rename_i h2
|
||||
exact h2
|
||||
exact (hi h2).elim
|
||||
|
@ -105,8 +105,8 @@ lemma planeY₃B₃_quad (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) :
|
|||
|
||||
lemma planeY₃B₃_cubic (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) :
|
||||
accCube (planeY₃B₃ R a b c).val = c ^ 2 *
|
||||
(3 * a * cubeTriLin R.val R.val Y₃.val
|
||||
+ 3 * b * cubeTriLin R.val R.val B₃.val + c * cubeTriLin R.val R.val R.val) := by
|
||||
(3 * a * cubeTriLin R.val R.val Y₃.val
|
||||
+ 3 * b * cubeTriLin R.val R.val B₃.val + c * cubeTriLin R.val R.val R.val) := by
|
||||
rw [planeY₃B₃_val]
|
||||
erw [TriLinearSymm.toCubic_add]
|
||||
erw [lineY₃B₃Charges_cubic]
|
||||
|
@ -178,7 +178,7 @@ def lineCube (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ℚ) :
|
|||
MSSMACC.LinSols :=
|
||||
planeY₃B₃ R
|
||||
(a₂ * cubeTriLin R.val R.val R.val - 3 * a₃ * cubeTriLin R.val R.val B₃.val)
|
||||
(3 * a₃ * cubeTriLin R.val R.val Y₃.val - a₁ * cubeTriLin R.val R.val R.val)
|
||||
(3 * a₃ * cubeTriLin R.val R.val Y₃.val - a₁ * cubeTriLin R.val R.val R.val)
|
||||
(3 * (a₁ * cubeTriLin R.val R.val B₃.val - a₂ * cubeTriLin R.val R.val Y₃.val))
|
||||
|
||||
lemma lineCube_smul (R : MSSMACC.AnomalyFreePerp) (a b c d : ℚ) :
|
||||
|
@ -205,7 +205,7 @@ lemma lineCube_quad (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ℚ) :
|
|||
|
||||
section proj
|
||||
|
||||
lemma α₃_proj (T : MSSMACC.Sols) : α₃ (proj T.1.1) =
|
||||
lemma α₃_proj (T : MSSMACC.Sols) : α₃ (proj T.1.1) =
|
||||
6 * dot Y₃.val B₃.val ^ 3 * (
|
||||
cubeTriLin T.val T.val Y₃.val * quadBiLin B₃.val T.val -
|
||||
cubeTriLin T.val T.val B₃.val * quadBiLin Y₃.val T.val) := by
|
||||
|
@ -214,13 +214,13 @@ lemma α₃_proj (T : MSSMACC.Sols) : α₃ (proj T.1.1) =
|
|||
ring
|
||||
|
||||
lemma α₂_proj (T : MSSMACC.Sols) : α₂ (proj T.1.1) =
|
||||
- α₃ (proj T.1.1) * (dot Y₃.val T.val - 2 * dot B₃.val T.val) := by
|
||||
- α₃ (proj T.1.1) * (dot Y₃.val T.val - 2 * dot B₃.val T.val) := by
|
||||
rw [α₃_proj, α₂]
|
||||
rw [cube_proj_proj_Y₃, quad_Y₃_proj, quad_proj, cube_proj]
|
||||
ring
|
||||
|
||||
lemma α₁_proj (T : MSSMACC.Sols) : α₁ (proj T.1.1) =
|
||||
- α₃ (proj T.1.1) * (dot B₃.val T.val - dot Y₃.val T.val) := by
|
||||
- α₃ (proj T.1.1) * (dot B₃.val T.val - dot Y₃.val T.val) := by
|
||||
rw [α₃_proj, α₁]
|
||||
rw [cube_proj_proj_B₃, quad_B₃_proj, quad_proj, cube_proj]
|
||||
ring
|
||||
|
|
|
@ -146,7 +146,7 @@ def InCubeSolProp (R : MSSMACC.Sols) : Prop :=
|
|||
/-- A rational which has two properties. It is zero for a solution `T` if and only if
|
||||
that solution satisfies `inCubeSolProp`. It appears in the definition of `inLineEqProj`. -/
|
||||
def cubicCoeff (T : MSSMACC.Sols) : ℚ :=
|
||||
3 * (dot Y₃.val B₃.val) ^ 3 * (cubeTriLin T.val T.val Y₃.val ^ 2 +
|
||||
3 * (dot Y₃.val B₃.val) ^ 3 * (cubeTriLin T.val T.val Y₃.val ^ 2 +
|
||||
cubeTriLin T.val T.val B₃.val ^ 2 )
|
||||
|
||||
lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
||||
|
@ -243,7 +243,7 @@ def toSolNSProj (T : MSSMACC.Sols) : MSSMACC.AnomalyFreePerp × ℚ × ℚ ×
|
|||
lemma toSolNS_proj (T : NotInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val := by
|
||||
apply ACCSystem.Sols.ext
|
||||
rw [toSolNS, toSolNSProj]
|
||||
change (lineEqCoeff T.val)⁻¹ • (toSolNSQuad _).1.1 = _
|
||||
change (lineEqCoeff T.val)⁻¹ • (toSolNSQuad _).1.1 = _
|
||||
rw [toSolNSQuad_eq_planeY₃B₃_on_α]
|
||||
rw [planeY₃B₃_val]
|
||||
rw [Y₃_plus_B₃_plus_proj]
|
||||
|
@ -254,7 +254,7 @@ lemma toSolNS_proj (T : NotInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val :=
|
|||
rw [lineEqCoeff]
|
||||
ring
|
||||
rw [h1]
|
||||
have h1 := (lineEqPropSol_iff_lineEqCoeff_zero T.val).mpr.mt T.prop
|
||||
have h1 := (lineEqPropSol_iff_lineEqCoeff_zero T.val).mpr.mt T.prop
|
||||
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h1]
|
||||
simp
|
||||
|
||||
|
@ -302,7 +302,7 @@ lemma inLineEqToSol_proj (T : InLineEqSol) : inLineEqToSol (inLineEqProj T) = T.
|
|||
simp
|
||||
|
||||
/-- Given a element of `inQuad × ℚ × ℚ × ℚ`, a solution to the ACCs. -/
|
||||
def inQuadToSol : InQuad × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a₁, a₂, a₃) =>
|
||||
def inQuadToSol : InQuad × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a₁, a₂, a₃) =>
|
||||
AnomalyFreeMk' (lineCube R.val.val a₁ a₂ a₃)
|
||||
(by
|
||||
erw [planeY₃B₃_quad]
|
||||
|
@ -391,8 +391,8 @@ lemma inQuadCubeToSol_proj (T : InQuadCubeSol) :
|
|||
|
||||
/-- Given an element of `MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ` a solution. We will
|
||||
show that this map is a surjection. -/
|
||||
def toSol : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a, b, c) =>
|
||||
if h₃ : LineEqProp R ∧ InQuadProp R ∧ InCubeProp R then
|
||||
def toSol : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a, b, c) =>
|
||||
if h₃ : LineEqProp R ∧ InQuadProp R ∧ InCubeProp R then
|
||||
inQuadCubeToSol (⟨⟨⟨R, h₃.1⟩, h₃.2.1⟩, h₃.2.2⟩, a, b, c)
|
||||
else
|
||||
if h₂ : LineEqProp R ∧ InQuadProp R then
|
||||
|
|
|
@ -17,7 +17,7 @@ and its action on the MSSM.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace MSSM
|
||||
|
||||
|
@ -30,7 +30,7 @@ open BigOperators
|
|||
def PermGroup := Fin 6 → Equiv.Perm (Fin 3)
|
||||
|
||||
@[simp]
|
||||
instance : Group PermGroup := Pi.group
|
||||
instance : Group PermGroup := Pi.group
|
||||
|
||||
/-- The image of an element of `permGroup` under the representation on charges. -/
|
||||
@[simps!]
|
||||
|
@ -65,7 +65,7 @@ def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where
|
|||
apply LinearMap.ext
|
||||
intro S
|
||||
rw [charges_eq_toSpecies_eq]
|
||||
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
intro i
|
||||
simp only [ Pi.mul_apply, Pi.inv_apply, Equiv.Perm.coe_mul, LinearMap.mul_apply]
|
||||
rw [chargeMap_toSpecies, chargeMap_toSpecies]
|
||||
|
@ -76,7 +76,7 @@ def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where
|
|||
apply LinearMap.ext
|
||||
intro S
|
||||
rw [charges_eq_toSpecies_eq]
|
||||
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
||||
intro i
|
||||
erw [toSMSpecies_toSpecies_inv]
|
||||
rfl
|
||||
|
@ -91,46 +91,46 @@ lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup) (S : MSSMCharges.Charges
|
|||
rw [repCharges_toSMSpecies]
|
||||
exact Equiv.sum_comp (f⁻¹ j) ((fun a => a ^ m) ∘ (toSMSpecies j) S)
|
||||
|
||||
lemma Hd_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma Hd_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
Hd (repCharges f S) = Hd S := rfl
|
||||
|
||||
lemma Hu_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma Hu_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
Hu (repCharges f S) = Hu S := rfl
|
||||
|
||||
lemma accGrav_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma accGrav_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
accGrav (repCharges f S) = accGrav S :=
|
||||
accGrav_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
(Hd_invariant f S)
|
||||
(Hu_invariant f S)
|
||||
|
||||
lemma accSU2_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma accSU2_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
accSU2 (repCharges f S) = accSU2 S :=
|
||||
accSU2_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
(Hd_invariant f S)
|
||||
(Hu_invariant f S)
|
||||
|
||||
lemma accSU3_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma accSU3_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
accSU3 (repCharges f S) = accSU3 S :=
|
||||
accSU3_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accYY_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma accYY_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
accYY (repCharges f S) = accYY S :=
|
||||
accYY_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
(Hd_invariant f S)
|
||||
(Hu_invariant f S)
|
||||
|
||||
lemma accQuad_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma accQuad_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
accQuad (repCharges f S) = accQuad S :=
|
||||
accQuad_ext
|
||||
(toSpecies_sum_invariant 2 f S)
|
||||
(Hd_invariant f S)
|
||||
(Hu_invariant f S)
|
||||
|
||||
lemma accCube_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
lemma accCube_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
||||
accCube (repCharges f S) = accCube S :=
|
||||
accCube_ext
|
||||
(fun j => toSpecies_sum_invariant 3 f S j)
|
||||
|
|
|
@ -16,7 +16,7 @@ We define the anomaly cancellation conditions for a pure U(1) gauge theory with
|
|||
universe v u
|
||||
open Nat
|
||||
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace PureU1
|
||||
open BigOperators
|
||||
|
@ -37,7 +37,7 @@ def accGrav (n : ℕ) : ((PureU1Charges n).Charges →ₗ[ℚ] ℚ) where
|
|||
/-- The symmetric trilinear form used to define the cubic anomaly. -/
|
||||
@[simps!]
|
||||
def accCubeTriLinSymm {n : ℕ} : TriLinearSymm (PureU1Charges n).Charges := TriLinearSymm.mk₃
|
||||
(fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i)
|
||||
(fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i)
|
||||
(by
|
||||
intro a S L T
|
||||
simp [HSMul.hSMul]
|
||||
|
@ -70,13 +70,13 @@ def accCubeTriLinSymm {n : ℕ} : TriLinearSymm (PureU1Charges n).Charges := Tri
|
|||
|
||||
/-- The cubic anomaly equation. -/
|
||||
@[simp]
|
||||
def accCube (n : ℕ) : HomogeneousCubic ((PureU1Charges n).Charges) :=
|
||||
def accCube (n : ℕ) : HomogeneousCubic ((PureU1Charges n).Charges) :=
|
||||
(accCubeTriLinSymm).toCubic
|
||||
|
||||
lemma accCube_explicit (n : ℕ) (S : (PureU1Charges n).Charges) :
|
||||
accCube n S = ∑ i : Fin n, S i ^ 3:= by
|
||||
rw [accCube, TriLinearSymm.toCubic]
|
||||
change accCubeTriLinSymm S S S = _
|
||||
change accCubeTriLinSymm S S S = _
|
||||
rw [accCubeTriLinSymm]
|
||||
simp only [PureU1Charges_numberCharges, TriLinearSymm.mk₃_toFun_apply_apply]
|
||||
apply Finset.sum_congr
|
||||
|
@ -99,7 +99,7 @@ def PureU1 (n : ℕ) : ACCSystem where
|
|||
|
||||
/-- An equivalence of vector spaces of charges when the number of fermions is equal. -/
|
||||
def pureU1EqCharges {n m : ℕ} (h : n = m):
|
||||
(PureU1 n).Charges ≃ₗ[ℚ] (PureU1 m).Charges where
|
||||
(PureU1 n).Charges ≃ₗ[ℚ] (PureU1 m).Charges where
|
||||
toFun f := f ∘ Fin.cast h.symm
|
||||
invFun f := f ∘ Fin.cast h
|
||||
map_add' _ _ := rfl
|
||||
|
|
|
@ -84,7 +84,7 @@ noncomputable
|
|||
def coordinateMap : ((PureU1 n.succ).LinSols) ≃ₗ[ℚ] Fin n →₀ ℚ where
|
||||
toFun S := Finsupp.equivFunOnFinite.invFun (S.1 ∘ Fin.castSucc)
|
||||
map_add' S T := Finsupp.ext (congrFun rfl)
|
||||
map_smul' a S := Finsupp.ext (congrFun rfl)
|
||||
map_smul' a S := Finsupp.ext (congrFun rfl)
|
||||
invFun f := ∑ i : Fin n, f i • asLinSols i
|
||||
left_inv S := by
|
||||
simp only [PureU1_numberCharges, Equiv.invFun_as_coe, Finsupp.equivFunOnFinite_symm_apply_toFun,
|
||||
|
@ -92,7 +92,7 @@ def coordinateMap : ((PureU1 n.succ).LinSols) ≃ₗ[ℚ] Fin n →₀ ℚ where
|
|||
apply pureU1_anomalyFree_ext
|
||||
intro j
|
||||
rw [sum_of_vectors]
|
||||
simp only [HSMul.hSMul, SMul.smul, PureU1_numberCharges,
|
||||
simp only [HSMul.hSMul, SMul.smul, PureU1_numberCharges,
|
||||
asLinSols_val, Equiv.toFun_as_coe,
|
||||
Fin.coe_eq_castSucc, mul_ite, mul_one, mul_neg, mul_zero, Equiv.invFun_as_coe]
|
||||
rw [Finset.sum_eq_single j]
|
||||
|
@ -127,7 +127,7 @@ instance : Module.Finite ℚ ((PureU1 n.succ).LinSols) :=
|
|||
|
||||
lemma finrank_AnomalyFreeLinear :
|
||||
FiniteDimensional.finrank ℚ (((PureU1 n.succ).LinSols)) = n := by
|
||||
have h := Module.mk_finrank_eq_card_basis (@asBasis n)
|
||||
have h := Module.mk_finrank_eq_card_basis (@asBasis n)
|
||||
simp only [Nat.succ_eq_add_one, finrank_eq_rank, Cardinal.mk_fintype, Fintype.card_fin] at h
|
||||
exact FiniteDimensional.finrank_eq_of_rank_eq h
|
||||
|
||||
|
|
|
@ -14,7 +14,7 @@ We look at charge assignments in which all charges have the same absolute value.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
open BigOperators
|
||||
|
||||
namespace PureU1
|
||||
|
@ -53,7 +53,7 @@ section charges
|
|||
variable {S : (PureU1 n.succ).Charges} {A : (PureU1 n.succ).LinSols}
|
||||
variable (hS : ConstAbsSorted S) (hA : ConstAbsSorted A.val)
|
||||
|
||||
lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := by
|
||||
lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := by
|
||||
have hSS := hS.2 i k hik
|
||||
have ht := hS.1 i k
|
||||
rw [sq_eq_sq_iff_eq_or_eq_neg] at ht
|
||||
|
@ -61,7 +61,7 @@ lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k :=
|
|||
exact h
|
||||
linarith
|
||||
|
||||
lemma val_le_zero {i : Fin n.succ} (hi : S i ≤ 0) : S i = S (0 : Fin n.succ) := by
|
||||
lemma val_le_zero {i : Fin n.succ} (hi : S i ≤ 0) : S i = S (0 : Fin n.succ) := by
|
||||
symm
|
||||
apply lt_eq hS hi
|
||||
simp
|
||||
|
@ -108,7 +108,7 @@ lemma boundary_succ {k : Fin n} (hk : Boundary S k) : S k.succ = - S (0 : Fin n.
|
|||
rw [opposite_signs_eq_neg hS (le_of_lt hk.left) (le_of_lt hk.right)] at hn
|
||||
linear_combination -(1 * hn)
|
||||
|
||||
lemma boundary_split (k : Fin n) : k.succ.val + (n.succ - k.succ.val) = n.succ := by
|
||||
lemma boundary_split (k : Fin n) : k.succ.val + (n.succ - k.succ.val) = n.succ := by
|
||||
omega
|
||||
|
||||
lemma boundary_accGrav' (k : Fin n) : accGrav n.succ S =
|
||||
|
@ -128,7 +128,7 @@ lemma boundary_accGrav'' (k : Fin n) (hk : Boundary S k) :
|
|||
S (Fin.cast (boundary_split k) (Fin.castAdd (n.succ - k.succ.val) i)) = S k.castSucc := by
|
||||
apply lt_eq hS (le_of_lt hk.left) (by rw [Fin.le_def]; simp; omega)
|
||||
have hsnd (i : Fin (n.succ - k.succ.val)) :
|
||||
S (Fin.cast (boundary_split k) (Fin.natAdd (k.succ.val) i)) = S k.succ := by
|
||||
S (Fin.cast (boundary_split k) (Fin.natAdd (k.succ.val) i)) = S k.succ := by
|
||||
apply gt_eq hS (le_of_lt hk.right) (by rw [Fin.le_def]; simp)
|
||||
simp only [hfst, hsnd]
|
||||
simp only [Fin.val_succ, sum_const, card_fin, nsmul_eq_mul, cast_add, cast_one,
|
||||
|
@ -160,7 +160,7 @@ lemma not_hasBoundry_zero (hnot : ¬ (HasBoundary S)) (i : Fin n.succ) :
|
|||
simp at hi
|
||||
exact zero_gt hS hi i
|
||||
|
||||
lemma not_hasBoundary_grav (hnot : ¬ (HasBoundary S)) :
|
||||
lemma not_hasBoundary_grav (hnot : ¬ (HasBoundary S)) :
|
||||
accGrav n.succ S = n.succ * S (0 : Fin n.succ) := by
|
||||
simp [accGrav, ← not_hasBoundry_zero hS hnot]
|
||||
|
||||
|
@ -196,7 +196,7 @@ lemma AFL_odd_zero {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted A.val)
|
|||
theorem AFL_odd (A : (PureU1 (2 * n + 1)).LinSols) (h : ConstAbsSorted A.val) :
|
||||
A = 0 := by
|
||||
apply ACCSystemLinear.LinSols.ext
|
||||
exact is_zero h (AFL_odd_zero h)
|
||||
exact is_zero h (AFL_odd_zero h)
|
||||
|
||||
lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val)
|
||||
(hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) {k : Fin (2 * n + 1)} (hk : Boundary A.val k) :
|
||||
|
@ -209,8 +209,8 @@ lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted
|
|||
linear_combination h0 / 2
|
||||
|
||||
lemma AFL_even_below' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val)
|
||||
(hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) (i : Fin n.succ) :
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) = A.val (0 : Fin (2*n.succ)) := by
|
||||
(hA : A.val (0 : Fin (2 * n.succ)) ≠ 0) (i : Fin n.succ) :
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i)) = A.val (0 : Fin (2*n.succ)) := by
|
||||
obtain ⟨k, hk⟩ := AFL_hasBoundary h hA
|
||||
rw [← boundary_castSucc h hk]
|
||||
apply lt_eq h (le_of_lt hk.left)
|
||||
|
@ -221,7 +221,7 @@ lemma AFL_even_below' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.
|
|||
|
||||
lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.val)
|
||||
(i : Fin n.succ) :
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i))
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i))
|
||||
= A.val (0 : Fin (2*n.succ)) := by
|
||||
by_cases hA : A.val (0 : Fin (2*n.succ)) = 0
|
||||
rw [is_zero h hA]
|
||||
|
@ -230,8 +230,8 @@ lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.v
|
|||
exact AFL_even_below' h hA i
|
||||
|
||||
lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val)
|
||||
(hA : A.val (0 : Fin (2*n.succ)) ≠ 0) (i : Fin n.succ) :
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) =
|
||||
(hA : A.val (0 : Fin (2*n.succ)) ≠ 0) (i : Fin n.succ) :
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) =
|
||||
- A.val (0 : Fin (2*n.succ)) := by
|
||||
obtain ⟨k, hk⟩ := AFL_hasBoundary h hA
|
||||
rw [← boundary_succ h hk]
|
||||
|
@ -243,7 +243,7 @@ lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.
|
|||
|
||||
lemma AFL_even_above (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.val)
|
||||
(i : Fin n.succ) :
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) =
|
||||
A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) =
|
||||
- A.val (0 : Fin (2*n.succ)) := by
|
||||
by_cases hA : A.val (0 : Fin (2*n.succ)) = 0
|
||||
rw [is_zero h hA]
|
||||
|
|
|
@ -16,7 +16,7 @@ that splits into two planes on which every point is a solution to the ACCs.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
open BigOperators
|
||||
|
||||
namespace PureU1
|
||||
|
@ -51,7 +51,7 @@ def δ!₃ : Fin (2 * n.succ) := (Fin.cast (n_cond₂ n) (Fin.castAdd ((n + n) +
|
|||
def δ!₄ : Fin (2 * n.succ) := (Fin.cast (n_cond₂ n) (Fin.natAdd 1 (Fin.natAdd (n + n) 0)))
|
||||
|
||||
lemma ext_δ (S T : Fin (2 * n.succ) → ℚ) (h1 : ∀ i, S (δ₁ i) = T (δ₁ i))
|
||||
(h2 : ∀ i, S (δ₂ i) = T (δ₂ i)) : S = T := by
|
||||
(h2 : ∀ i, S (δ₂ i) = T (δ₂ i)) : S = T := by
|
||||
funext i
|
||||
by_cases hi : i.val < n.succ
|
||||
let j : Fin n.succ := ⟨i, hi⟩
|
||||
|
@ -68,7 +68,7 @@ lemma ext_δ (S T : Fin (2 * n.succ) → ℚ) (h1 : ∀ i, S (δ₁ i) = T (δ
|
|||
rw [h3] at h2
|
||||
exact h2
|
||||
|
||||
lemma sum_δ₁_δ₂ (S : Fin (2 * n.succ) → ℚ) :
|
||||
lemma sum_δ₁_δ₂ (S : Fin (2 * n.succ) → ℚ) :
|
||||
∑ i, S i = ∑ i : Fin n.succ, ((S ∘ δ₁) i + (S ∘ δ₂) i) := by
|
||||
have h1 : ∑ i, S i = ∑ i : Fin (n.succ + n.succ), S (Fin.cast (split_equal n.succ) i) := by
|
||||
rw [Finset.sum_equiv (Fin.castOrderIso (split_equal n.succ)).symm.toEquiv]
|
||||
|
@ -80,7 +80,7 @@ lemma sum_δ₁_δ₂ (S : Fin (2 * n.succ) → ℚ) :
|
|||
rw [Fin.sum_univ_add, Finset.sum_add_distrib]
|
||||
rfl
|
||||
|
||||
lemma sum_δ₁_δ₂' (S : Fin (2 * n.succ) → ℚ) :
|
||||
lemma sum_δ₁_δ₂' (S : Fin (2 * n.succ) → ℚ) :
|
||||
∑ i, S i = ∑ i : Fin n.succ, ((S ∘ δ₁) i + (S ∘ δ₂) i) := by
|
||||
have h1 : ∑ i, S i = ∑ i : Fin (n.succ + n.succ), S (Fin.cast (split_equal n.succ) i) := by
|
||||
rw [Finset.sum_equiv (Fin.castOrderIso (split_equal n.succ)).symm.toEquiv]
|
||||
|
@ -92,8 +92,8 @@ lemma sum_δ₁_δ₂' (S : Fin (2 * n.succ) → ℚ) :
|
|||
rw [Fin.sum_univ_add, Finset.sum_add_distrib]
|
||||
rfl
|
||||
|
||||
lemma sum_δ!₁_δ!₂ (S : Fin (2 * n.succ) → ℚ) :
|
||||
∑ i, S i = S δ!₃ + S δ!₄ + ∑ i : Fin n, ((S ∘ δ!₁) i + (S ∘ δ!₂) i) := by
|
||||
lemma sum_δ!₁_δ!₂ (S : Fin (2 * n.succ) → ℚ) :
|
||||
∑ i, S i = S δ!₃ + S δ!₄ + ∑ i : Fin n, ((S ∘ δ!₁) i + (S ∘ δ!₂) i) := by
|
||||
have h1 : ∑ i, S i = ∑ i : Fin (1 + ((n + n) + 1)), S (Fin.cast (n_cond₂ n) i) := by
|
||||
rw [Finset.sum_equiv (Fin.castOrderIso (n_cond₂ n)).symm.toEquiv]
|
||||
intro i
|
||||
|
@ -180,12 +180,12 @@ lemma basis_on_δ₁_other {k j : Fin n.succ} (h : k ≠ j) :
|
|||
omega
|
||||
rfl
|
||||
|
||||
lemma basis_on_other {k : Fin n.succ} {j : Fin (2 * n.succ)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) :
|
||||
lemma basis_on_other {k : Fin n.succ} {j : Fin (2 * n.succ)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) :
|
||||
basisAsCharges k j = 0 := by
|
||||
simp [basisAsCharges]
|
||||
simp_all only [ne_eq, ↓reduceIte]
|
||||
|
||||
lemma basis!_on_other {k : Fin n} {j : Fin (2 * n.succ)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) :
|
||||
lemma basis!_on_other {k : Fin n} {j : Fin (2 * n.succ)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) :
|
||||
basis!AsCharges k j = 0 := by
|
||||
simp [basis!AsCharges]
|
||||
simp_all only [ne_eq, ↓reduceIte]
|
||||
|
@ -338,11 +338,11 @@ def basisa : (Fin n.succ) ⊕ (Fin n) → (PureU1 (2 * n.succ)).LinSols := fun i
|
|||
/-- Swapping the elements δ!₁ j and δ!₂ j is equivalent to adding a vector basis!AsCharges j. -/
|
||||
lemma swap!_as_add {S S' : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
|
||||
(hS : ((FamilyPermutations (2 * n.succ)).linSolRep
|
||||
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') :
|
||||
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') :
|
||||
S'.val = S.val + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j := by
|
||||
funext i
|
||||
rw [← hS, FamilyPermutations_anomalyFreeLinear_apply]
|
||||
by_cases hi : i = δ!₁ j
|
||||
by_cases hi : i = δ!₁ j
|
||||
subst hi
|
||||
simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst]
|
||||
by_cases hi2 : i = δ!₂ j
|
||||
|
@ -350,7 +350,7 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
|
|||
simp [HSMul.hSMul, basis!_on_δ!₂_self, pairSwap_inv_snd]
|
||||
simp [HSMul.hSMul]
|
||||
rw [basis!_on_other hi hi2]
|
||||
change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_
|
||||
change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_
|
||||
erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)]
|
||||
simp
|
||||
|
||||
|
@ -374,7 +374,7 @@ lemma P_δ₁ (f : Fin n.succ → ℚ) (j : Fin n.succ) : P f (δ₁ j) = f j :=
|
|||
simp only [mul_zero]
|
||||
simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
|
||||
|
||||
lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by
|
||||
lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by
|
||||
rw [P!, sum_of_charges]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
rw [Finset.sum_eq_single j]
|
||||
|
@ -424,7 +424,7 @@ lemma P!_δ!₃ (f : Fin n → ℚ) : P! f (δ!₃) = 0 := by
|
|||
rw [P!, sum_of_charges]
|
||||
simp [HSMul.hSMul, SMul.smul, basis!_on_δ!₃]
|
||||
|
||||
lemma Pa_δ!₃ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₃) = f 0 := by
|
||||
lemma Pa_δ!₃ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₃) = f 0 := by
|
||||
rw [Pa]
|
||||
simp only [ACCSystemCharges.chargesAddCommMonoid_add]
|
||||
rw [P!_δ!₃, δ!₃_δ₁0, P_δ₁]
|
||||
|
@ -434,13 +434,13 @@ lemma P!_δ!₄ (f : Fin n → ℚ) : P! f (δ!₄) = 0 := by
|
|||
rw [P!, sum_of_charges]
|
||||
simp [HSMul.hSMul, SMul.smul, basis!_on_δ!₄]
|
||||
|
||||
lemma Pa_δ!₄ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₄) = - f (Fin.last n) := by
|
||||
lemma Pa_δ!₄ (f : Fin n.succ → ℚ) (g : Fin n → ℚ) : Pa f g (δ!₄) = - f (Fin.last n) := by
|
||||
rw [Pa]
|
||||
simp only [ACCSystemCharges.chargesAddCommMonoid_add]
|
||||
rw [P!_δ!₄, δ!₄_δ₂Last, P_δ₂]
|
||||
simp
|
||||
|
||||
lemma P_δ₁_δ₂ (f : Fin n.succ → ℚ) : P f ∘ δ₂ = - P f ∘ δ₁ := by
|
||||
lemma P_δ₁_δ₂ (f : Fin n.succ → ℚ) : P f ∘ δ₂ = - P f ∘ δ₁ := by
|
||||
funext j
|
||||
simp only [PureU1_numberCharges, Function.comp_apply, Pi.neg_apply]
|
||||
rw [P_δ₁, P_δ₂]
|
||||
|
@ -484,7 +484,7 @@ lemma P_P_P!_accCube (g : Fin n.succ → ℚ) (j : Fin n) :
|
|||
|
||||
lemma P_P!_P!_accCube (g : Fin n → ℚ) (j : Fin n.succ) :
|
||||
accCubeTriLinSymm (P! g) (P! g) (basisAsCharges j)
|
||||
= (P! g (δ₁ j))^2 - (P! g (δ₂ j))^2 := by
|
||||
= (P! g (δ₁ j))^2 - (P! g (δ₂ j))^2 := by
|
||||
simp [accCubeTriLinSymm]
|
||||
rw [sum_δ₁_δ₂]
|
||||
simp only [Function.comp_apply]
|
||||
|
@ -604,7 +604,7 @@ theorem basisa_linear_independent : LinearIndependent ℚ (@basisa n) := by
|
|||
simp_all
|
||||
simp_all
|
||||
|
||||
lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by
|
||||
lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by
|
||||
apply Iff.intro
|
||||
intro h
|
||||
funext i
|
||||
|
@ -625,7 +625,7 @@ lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ℚ) : Pa' f = Pa' f' ↔ f =
|
|||
|
||||
/-! TODO: Replace the definition of `join` with a Mathlib definition, most likely `Sum.elim`. -/
|
||||
/-- A helper function for what follows. -/
|
||||
def join (g : Fin n.succ → ℚ) (f : Fin n → ℚ) : (Fin n.succ) ⊕ (Fin n) → ℚ := fun i =>
|
||||
def join (g : Fin n.succ → ℚ) (f : Fin n → ℚ) : (Fin n.succ) ⊕ (Fin n) → ℚ := fun i =>
|
||||
match i with
|
||||
| .inl i => g i
|
||||
| .inr i => f i
|
||||
|
@ -661,7 +661,7 @@ lemma Pa_eq (g g' : Fin n.succ → ℚ) (f f' : Fin n → ℚ) :
|
|||
rw [← join_ext]
|
||||
exact Pa'_eq _ _
|
||||
|
||||
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) =
|
||||
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) =
|
||||
FiniteDimensional.finrank ℚ (PureU1 (2 * n.succ)).LinSols := by
|
||||
erw [BasisLinear.finrank_AnomalyFreeLinear]
|
||||
simp only [Fintype.card_sum, Fintype.card_fin, mul_eq]
|
||||
|
@ -673,7 +673,7 @@ noncomputable def basisaAsBasis :
|
|||
basisOfLinearIndependentOfCardEqFinrank (@basisa_linear_independent n) basisa_card
|
||||
|
||||
lemma span_basis (S : (PureU1 (2 * n.succ)).LinSols) :
|
||||
∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f := by
|
||||
∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f := by
|
||||
have h := (mem_span_range_iff_exists_fun ℚ).mp (Basis.mem_span basisaAsBasis S)
|
||||
obtain ⟨f, hf⟩ := h
|
||||
simp [basisaAsBasis] at hf
|
||||
|
@ -706,7 +706,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
|
|||
S'.val = P g' + P! f' ∧ P! f' = P! f +
|
||||
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∧ g' = g := by
|
||||
let X := P! f + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j
|
||||
have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by
|
||||
have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by
|
||||
apply Submodule.add_mem
|
||||
exact (P!_in_span f)
|
||||
exact (smul_basis!AsCharges_in_span S j)
|
||||
|
|
|
@ -34,7 +34,7 @@ open BigOperators
|
|||
variable {n : ℕ}
|
||||
open VectorLikeEvenPlane
|
||||
|
||||
/-- A property on `LinSols`, satisfied if every point on the line between the two planes
|
||||
/-- A property on `LinSols`, satisfied if every point on the line between the two planes
|
||||
in the basis through that point is in the cubic. -/
|
||||
def LineInCubic (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) ,
|
||||
|
@ -61,18 +61,18 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S)
|
|||
then `accCubeTriLinSymm.toFun (P g, P g, P! f) = 0`.
|
||||
-/
|
||||
lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S) :
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f),
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f),
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
|
||||
intro g f hS
|
||||
linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1) -
|
||||
(lineInCubic_expand h g f hS 1 2) / 6
|
||||
|
||||
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
|
||||
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
|
||||
def LineInCubicPerm (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=
|
||||
∀ (M : (FamilyPermutations (2 * n.succ)).group ),
|
||||
LineInCubic ((FamilyPermutations (2 * n.succ)).linSolRep M S)
|
||||
|
||||
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
||||
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
||||
lemma lineInCubicPerm_self {S : (PureU1 (2 * n.succ)).LinSols}
|
||||
(hS : LineInCubicPerm S) : LineInCubic S := hS 1
|
||||
|
||||
|
@ -94,7 +94,7 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols}
|
|||
(S.val (δ!₂ j) - S.val (δ!₁ j))
|
||||
* accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by
|
||||
intro j g f h
|
||||
let S' := (FamilyPermutations (2 * n.succ)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j)) S
|
||||
let S' := (FamilyPermutations (2 * n.succ)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j)) S
|
||||
have hSS' : ((FamilyPermutations (2 * n.succ)).linSolRep (pairSwap (δ!₁ j) (δ!₂ j))) S = S' := rfl
|
||||
obtain ⟨g', f', hall⟩ := span_basis_swap! j hSS' g f h
|
||||
have h1 := line_in_cubic_P_P_P! (lineInCubicPerm_self LIC) g f h
|
||||
|
@ -106,21 +106,21 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ)).LinSols}
|
|||
|
||||
lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ )).LinSols}
|
||||
(f : Fin n.succ.succ → ℚ) (g : Fin n.succ → ℚ) (hS : S.val = Pa f g) :
|
||||
accCubeTriLinSymm (P f) (P f) (basis!AsCharges (Fin.last n)) =
|
||||
accCubeTriLinSymm (P f) (P f) (basis!AsCharges (Fin.last n)) =
|
||||
- (S.val (δ!₂ (Fin.last n)) + S.val (δ!₁ (Fin.last n))) * (2 * S.val δ!₄ +
|
||||
S.val (δ!₂ (Fin.last n)) + S.val (δ!₁ (Fin.last n))) := by
|
||||
rw [P_P_P!_accCube f (Fin.last n)]
|
||||
S.val (δ!₂ (Fin.last n)) + S.val (δ!₁ (Fin.last n))) := by
|
||||
rw [P_P_P!_accCube f (Fin.last n)]
|
||||
have h1 := Pa_δ!₄ f g
|
||||
have h2 := Pa_δ!₁ f g (Fin.last n)
|
||||
have h3 := Pa_δ!₂ f g (Fin.last n)
|
||||
simp at h1 h2 h3
|
||||
have hl : f (Fin.succ (Fin.last (n ))) = - Pa f g δ!₄ := by
|
||||
have hl : f (Fin.succ (Fin.last (n ))) = - Pa f g δ!₄ := by
|
||||
simp_all only [Fin.succ_last, neg_neg]
|
||||
erw [hl] at h2
|
||||
have hg : g (Fin.last n) = Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄ := by
|
||||
have hg : g (Fin.last n) = Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄ := by
|
||||
linear_combination -(1 * h2)
|
||||
have hll : f (Fin.castSucc (Fin.last (n ))) =
|
||||
- (Pa f g (δ!₂ (Fin.last n)) + Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄) := by
|
||||
have hll : f (Fin.castSucc (Fin.last (n ))) =
|
||||
- (Pa f g (δ!₂ (Fin.last n)) + Pa f g (δ!₁ (Fin.last n)) + Pa f g δ!₄) := by
|
||||
linear_combination h3 - 1 * hg
|
||||
rw [← hS] at hl hll
|
||||
rw [hl, hll]
|
||||
|
@ -145,7 +145,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols}
|
|||
apply Or.inr
|
||||
exact h1
|
||||
|
||||
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
|
||||
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
|
||||
(LIC : LineInCubicPerm S) : LineInPlaneCond S := by
|
||||
refine @Prop_three (2 * n.succ.succ) LineInPlaneProp S (δ!₂ (Fin.last n)) (δ!₁ (Fin.last n))
|
||||
δ!₄ ?_ ?_ ?_ ?_
|
||||
|
@ -156,15 +156,15 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
|
|||
intro M
|
||||
exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
|
||||
|
||||
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ)).Sols}
|
||||
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ)).Sols}
|
||||
(LIC : LineInCubicPerm S.1.1) : ConstAbs S.val :=
|
||||
linesInPlane_constAbs_AF S (lineInCubicPerm_last_perm LIC)
|
||||
|
||||
theorem lineInCubicPerm_vectorLike {S : (PureU1 (2 * n.succ.succ)).Sols}
|
||||
theorem lineInCubicPerm_vectorLike {S : (PureU1 (2 * n.succ.succ)).Sols}
|
||||
(LIC : LineInCubicPerm S.1.1) : VectorLikeEven S.val :=
|
||||
ConstAbs.boundary_value_even S.1.1 (lineInCubicPerm_constAbs LIC)
|
||||
|
||||
theorem lineInCubicPerm_in_plane (S : (PureU1 (2 * n.succ.succ)).Sols)
|
||||
theorem lineInCubicPerm_in_plane (S : (PureU1 (2 * n.succ.succ)).Sols)
|
||||
(LIC : LineInCubicPerm S.1.1) : ∃ (M : (FamilyPermutations (2 * n.succ.succ)).group),
|
||||
(FamilyPermutations (2 * n.succ.succ)).linSolRep M S.1.1
|
||||
∈ Submodule.span ℚ (Set.range basis) :=
|
||||
|
|
|
@ -14,7 +14,7 @@ import Mathlib.Tactic.Polyrith
|
|||
/-!
|
||||
# Parameterization in even case
|
||||
|
||||
Given maps `g : Fin n.succ → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly
|
||||
Given maps `g : Fin n.succ → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly
|
||||
equations. We show that every solution can be got in this way, up to permutation, unless it, up to
|
||||
permutation, lives in the plane spanned by the first part of the basis vector.
|
||||
|
||||
|
@ -60,7 +60,7 @@ lemma parameterizationCharge_cube (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (
|
|||
accCubeTriLinSymm.map_smul₃]
|
||||
ring
|
||||
|
||||
/-- The construction of a `Sol` from a `Fin n.succ → ℚ`, a `Fin n → ℚ` and a `ℚ`. -/
|
||||
/-- The construction of a `Sol` from a `Fin n.succ → ℚ`, a `Fin n → ℚ` and a `ℚ`. -/
|
||||
def parameterization (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (a : ℚ) :
|
||||
(PureU1 (2 * n.succ)).Sols :=
|
||||
⟨⟨parameterizationAsLinear g f a, by intro i; simp at i; exact Fin.elim0 i⟩,
|
||||
|
@ -77,15 +77,15 @@ lemma anomalyFree_param {S : (PureU1 (2 * n.succ)).Sols}
|
|||
erw [P!_accCube] at hC
|
||||
linear_combination hC / 3
|
||||
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
In this case our parameterization above will be able to recover this point. -/
|
||||
def GenericCase (S : (PureU1 (2 * n.succ)).Sols) : Prop :=
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) ,
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
|
||||
|
||||
lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
||||
(hs : ∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f ∧
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by
|
||||
intro g f hS hC
|
||||
obtain ⟨g', f', hS', hC'⟩ := hs
|
||||
rw [hS] at hS'
|
||||
|
@ -94,13 +94,13 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
|||
exact hC' hC
|
||||
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) = 0`.-/
|
||||
def SpecialCase (S : (PureU1 (2 * n.succ)).Sols) : Prop :=
|
||||
def SpecialCase (S : (PureU1 (2 * n.succ)).Sols) : Prop :=
|
||||
∀ (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f) ,
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0
|
||||
|
||||
lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
||||
(hs : ∃ (g : Fin n.succ → ℚ) (f : Fin n → ℚ), S.val = P g + P! f ∧
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by
|
||||
intro g f hS
|
||||
obtain ⟨g', f', hS', hC'⟩ := hs
|
||||
rw [hS] at hS'
|
||||
|
@ -111,7 +111,7 @@ lemma specialCase_exists (S : (PureU1 (2 * n.succ)).Sols)
|
|||
lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) :
|
||||
GenericCase S ∨ SpecialCase S := by
|
||||
obtain ⟨g, f, h⟩ := span_basis S.1.1
|
||||
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨
|
||||
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
|
||||
exact ne_or_eq _ _
|
||||
cases h1 <;> rename_i h1
|
||||
|
@ -119,7 +119,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) :
|
|||
exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩)
|
||||
|
||||
theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : GenericCase S) :
|
||||
∃ g f a, S = parameterization g f a := by
|
||||
∃ g f a, S = parameterization g f a := by
|
||||
obtain ⟨g, f, hS⟩ := span_basis S.1.1
|
||||
use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹
|
||||
rw [parameterization]
|
||||
|
|
|
@ -14,7 +14,7 @@ import Mathlib.RepresentationTheory.Basic
|
|||
|
||||
We say a `LinSol` satisfies the `line in plane` condition if for all distinct `i1`, `i2`, `i3` in
|
||||
`Fin n`, we have
|
||||
`S i1 = S i2` or `S i1 = - S i2` or `2 S i3 + S i1 + S i2 = 0`.
|
||||
`S i1 = S i2` or `S i1 = - S i2` or `2 S i3 + S i1 + S i2 = 0`.
|
||||
|
||||
We look at various consequences of this.
|
||||
The main reference for this material is
|
||||
|
@ -60,7 +60,7 @@ lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : LineIn
|
|||
simp only [LineInPlaneProp] at hS
|
||||
simp [not_or] at h
|
||||
have h1 (i : Fin n) : S.val i.castSucc.castSucc =
|
||||
- (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by
|
||||
- (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by
|
||||
have h1S := hS (Fin.last n).castSucc ((Fin.last n).succ) i.castSucc.castSucc
|
||||
(by simp; rw [Fin.ext_iff]; simp)
|
||||
(by simp; rw [Fin.ext_iff]; simp; omega)
|
||||
|
@ -82,14 +82,14 @@ lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
|
|||
have h := lineInPlaneCond_eq_last' hS hn
|
||||
rw [sq_eq_sq_iff_eq_or_eq_neg] at hn
|
||||
simp [or_not] at hn
|
||||
have hx : ((2 : ℚ) - ↑(n + 3)) ≠ 0 := by
|
||||
have hx : ((2 : ℚ) - ↑(n + 3)) ≠ 0 := by
|
||||
rw [Nat.cast_add]
|
||||
simp only [Nat.cast_ofNat, ne_eq]
|
||||
apply Not.intro
|
||||
intro a
|
||||
linarith
|
||||
have ht : S.val ((Fin.last n.succ.succ.succ).succ) =
|
||||
- S.val ((Fin.last n.succ.succ.succ).castSucc) := by
|
||||
have ht : S.val ((Fin.last n.succ.succ.succ).succ) =
|
||||
- S.val ((Fin.last n.succ.succ.succ).castSucc) := by
|
||||
rw [← mul_right_inj' hx]
|
||||
simp [Nat.succ_eq_add_one]
|
||||
simp at h
|
||||
|
@ -115,7 +115,7 @@ theorem linesInPlane_constAbs {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
|
|||
rw [hij]
|
||||
|
||||
lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
|
||||
ConstAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by
|
||||
ConstAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by
|
||||
simp [ConstAbsProp]
|
||||
by_contra hn
|
||||
have hLin := pureU1_linear S.1.1
|
||||
|
@ -128,7 +128,7 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
|
|||
have l013 := hS 0 1 3 (by simp) (by simp) (by simp)
|
||||
have l023 := hS 0 2 3 (by simp) (by simp) (by simp)
|
||||
simp_all [LineInPlaneProp]
|
||||
have h1 : S.val (2 : Fin 4) = S.val (3 : Fin 4) := by
|
||||
have h1 : S.val (2 : Fin 4) = S.val (3 : Fin 4) := by
|
||||
linear_combination l012 / 2 + -1 * l013 / 2
|
||||
by_cases h2 : S.val (0 : Fin 4) = S.val (2 : Fin 4)
|
||||
simp_all
|
||||
|
@ -161,7 +161,7 @@ lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols}
|
|||
refine Prop_two ConstAbsProp (by simp : (0 : Fin 4) ≠ 1) ?_
|
||||
intro M
|
||||
let S' := (FamilyPermutations 4).solAction.toFun S M
|
||||
have hS' : LineInPlaneCond S'.1.1 :=
|
||||
have hS' : LineInPlaneCond S'.1.1 :=
|
||||
(lineInPlaneCond_perm hS M)
|
||||
exact linesInPlane_four S' hS'
|
||||
|
||||
|
|
|
@ -13,7 +13,7 @@ We show that in this case the charge must be zero.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace PureU1
|
||||
|
||||
|
|
|
@ -14,7 +14,7 @@ We define a surjective map from `LinSols` with a charge equal to zero to `Sols`.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace PureU1
|
||||
|
||||
|
|
|
@ -13,7 +13,7 @@ We define an equivalence between `LinSols` and `Sols`.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace PureU1
|
||||
|
||||
|
|
|
@ -16,7 +16,7 @@ that splits into two planes on which every point is a solution to the ACCs.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
open BigOperators
|
||||
|
||||
namespace PureU1
|
||||
|
@ -111,7 +111,7 @@ lemma δ₂_δ!₂ (j : Fin n) : δ₂ j = δ!₂ j := by
|
|||
simp [δ₂, δ!₂]
|
||||
omega
|
||||
|
||||
lemma sum_δ (S : Fin (2 * n + 1) → ℚ) :
|
||||
lemma sum_δ (S : Fin (2 * n + 1) → ℚ) :
|
||||
∑ i, S i = S δ₃ + ∑ i : Fin n, ((S ∘ δ₁) i + (S ∘ δ₂) i) := by
|
||||
have h1 : ∑ i, S i = ∑ i : Fin (n + 1 + n), S (Fin.cast (split_odd n) i) := by
|
||||
rw [Finset.sum_equiv (Fin.castOrderIso (split_odd n)).symm.toEquiv]
|
||||
|
@ -127,7 +127,7 @@ lemma sum_δ (S : Fin (2 * n + 1) → ℚ) :
|
|||
rw [Finset.sum_add_distrib]
|
||||
rfl
|
||||
|
||||
lemma sum_δ! (S : Fin (2 * n + 1) → ℚ) :
|
||||
lemma sum_δ! (S : Fin (2 * n + 1) → ℚ) :
|
||||
∑ i, S i = S δ!₃ + ∑ i : Fin n, ((S ∘ δ!₁) i + (S ∘ δ!₂) i) := by
|
||||
have h1 : ∑ i, S i = ∑ i : Fin ((1+n)+n), S (Fin.cast (split_odd! n) i) := by
|
||||
rw [Finset.sum_equiv (Fin.castOrderIso (split_odd! n)).symm.toEquiv]
|
||||
|
@ -210,12 +210,12 @@ lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) :
|
|||
omega
|
||||
rfl
|
||||
|
||||
lemma basis_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) :
|
||||
lemma basis_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) :
|
||||
basisAsCharges k j = 0 := by
|
||||
simp [basisAsCharges]
|
||||
simp_all only [ne_eq, ↓reduceIte]
|
||||
|
||||
lemma basis!_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) :
|
||||
lemma basis!_on_other {k : Fin n} {j : Fin (2 * n + 1)} (h1 : j ≠ δ!₁ k) (h2 : j ≠ δ!₂ k) :
|
||||
basis!AsCharges k j = 0 := by
|
||||
simp [basis!AsCharges]
|
||||
simp_all only [ne_eq, ↓reduceIte]
|
||||
|
@ -333,11 +333,11 @@ end theBasisVectors
|
|||
/-- Swapping the elements δ!₁ j and δ!₂ j is equivalent to adding a vector basis!AsCharges j. -/
|
||||
lemma swap!_as_add {S S' : (PureU1 (2 * n + 1)).LinSols} (j : Fin n)
|
||||
(hS : ((FamilyPermutations (2 * n + 1)).linSolRep
|
||||
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') :
|
||||
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') :
|
||||
S'.val = S.val + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j := by
|
||||
funext i
|
||||
rw [← hS, FamilyPermutations_anomalyFreeLinear_apply]
|
||||
by_cases hi : i = δ!₁ j
|
||||
by_cases hi : i = δ!₁ j
|
||||
subst hi
|
||||
simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst]
|
||||
by_cases hi2 : i = δ!₂ j
|
||||
|
@ -345,7 +345,7 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n + 1)).LinSols} (j : Fin n)
|
|||
simp [HSMul.hSMul,basis!_on_δ!₂_self, pairSwap_inv_snd]
|
||||
simp [HSMul.hSMul]
|
||||
rw [basis!_on_other hi hi2]
|
||||
change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_
|
||||
change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_
|
||||
erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)]
|
||||
simp
|
||||
|
||||
|
@ -369,7 +369,7 @@ lemma P_δ₁ (f : Fin n → ℚ) (j : Fin n) : P f (δ₁ j) = f j := by
|
|||
simp only [mul_zero]
|
||||
simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
|
||||
|
||||
lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by
|
||||
lemma P!_δ!₁ (f : Fin n → ℚ) (j : Fin n) : P! f (δ!₁ j) = f j := by
|
||||
rw [P!, sum_of_charges]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
rw [Finset.sum_eq_single j]
|
||||
|
@ -418,14 +418,14 @@ lemma Pa_δa₁ (f g : Fin n.succ → ℚ) : Pa f g δa₁ = f 0 := by
|
|||
rw [P_δ₁, P!_δ!₃]
|
||||
simp
|
||||
|
||||
lemma Pa_δa₂ (f g : Fin n.succ → ℚ) (j : Fin n) : Pa f g (δa₂ j) = f j.succ + g j.castSucc := by
|
||||
lemma Pa_δa₂ (f g : Fin n.succ → ℚ) (j : Fin n) : Pa f g (δa₂ j) = f j.succ + g j.castSucc := by
|
||||
rw [Pa]
|
||||
simp only [ACCSystemCharges.chargesAddCommMonoid_add]
|
||||
nth_rewrite 1 [δa₂_δ₁]
|
||||
rw [δa₂_δ!₁]
|
||||
rw [P_δ₁, P!_δ!₁]
|
||||
|
||||
lemma Pa_δa₃ (f g : Fin n.succ → ℚ) : Pa f g (δa₃) = g (Fin.last n) := by
|
||||
lemma Pa_δa₃ (f g : Fin n.succ → ℚ) : Pa f g (δa₃) = g (Fin.last n) := by
|
||||
rw [Pa]
|
||||
simp only [ACCSystemCharges.chargesAddCommMonoid_add]
|
||||
nth_rewrite 1 [δa₃_δ₃]
|
||||
|
@ -495,7 +495,7 @@ lemma P!_zero (f : Fin n → ℚ) (h : P! f = 0) : ∀ i, f i = 0 := by
|
|||
rw [h]
|
||||
rfl
|
||||
|
||||
lemma Pa_zero (f g : Fin n.succ → ℚ) (h : Pa f g = 0) :
|
||||
lemma Pa_zero (f g : Fin n.succ → ℚ) (h : Pa f g = 0) :
|
||||
∀ i, f i = 0 := by
|
||||
have h₃ := Pa_δa₁ f g
|
||||
rw [h] at h₃
|
||||
|
@ -589,7 +589,7 @@ theorem basisa_linear_independent : LinearIndependent ℚ (@basisa n.succ) := by
|
|||
simp_all
|
||||
simp_all
|
||||
|
||||
lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by
|
||||
lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ℚ) : Pa' f = Pa' f' ↔ f = f' := by
|
||||
apply Iff.intro
|
||||
intro h
|
||||
funext i
|
||||
|
@ -610,7 +610,7 @@ lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ℚ) : Pa' f = Pa' f'
|
|||
|
||||
/-! TODO: Replace the definition of `join` with a Mathlib definition, most likely `Sum.elim`. -/
|
||||
/-- A helper function for what follows. -/
|
||||
def join (g f : Fin n → ℚ) : Fin n ⊕ Fin n → ℚ := fun i =>
|
||||
def join (g f : Fin n → ℚ) : Fin n ⊕ Fin n → ℚ := fun i =>
|
||||
match i with
|
||||
| .inl i => g i
|
||||
| .inr i => f i
|
||||
|
@ -646,7 +646,7 @@ lemma Pa_eq (g g' : Fin n.succ → ℚ) (f f' : Fin n.succ → ℚ) :
|
|||
rw [← join_ext]
|
||||
exact Pa'_eq _ _
|
||||
|
||||
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n.succ)) =
|
||||
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n.succ)) =
|
||||
FiniteDimensional.finrank ℚ (PureU1 (2 * n.succ + 1)).LinSols := by
|
||||
erw [BasisLinear.finrank_AnomalyFreeLinear]
|
||||
simp only [Fintype.card_sum, Fintype.card_fin]
|
||||
|
@ -658,7 +658,7 @@ noncomputable def basisaAsBasis :
|
|||
basisOfLinearIndependentOfCardEqFinrank (@basisa_linear_independent n) basisa_card
|
||||
|
||||
lemma span_basis (S : (PureU1 (2 * n.succ + 1)).LinSols) :
|
||||
∃ (g f : Fin n.succ → ℚ) , S.val = P g + P! f := by
|
||||
∃ (g f : Fin n.succ → ℚ) , S.val = P g + P! f := by
|
||||
have h := (mem_span_range_iff_exists_fun ℚ).mp (Basis.mem_span basisaAsBasis S)
|
||||
obtain ⟨f, hf⟩ := h
|
||||
simp [basisaAsBasis] at hf
|
||||
|
@ -686,7 +686,7 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ + 1)).LinSols} (j : Fin n.succ)
|
|||
apply SetLike.mem_of_subset
|
||||
apply Submodule.subset_span
|
||||
simp_all only [Set.mem_range, exists_apply_eq_apply]
|
||||
have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by
|
||||
have hX : X ∈ Submodule.span ℚ (Set.range (basis!AsCharges)) := by
|
||||
apply Submodule.add_mem
|
||||
exact hf
|
||||
exact hP
|
||||
|
|
|
@ -34,10 +34,10 @@ open BigOperators
|
|||
variable {n : ℕ}
|
||||
open VectorLikeOddPlane
|
||||
|
||||
/-- A property on `LinSols`, satisfied if every point on the line between the two planes
|
||||
/-- A property on `LinSols`, satisfied if every point on the line between the two planes
|
||||
in the basis through that point is in the cubic. -/
|
||||
def LineInCubic (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
|
||||
∀ (g f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) ,
|
||||
∀ (g f : Fin n → ℚ) (_ : S.val = Pa g f) (a b : ℚ) ,
|
||||
accCube (2 * n + 1) (a • P g + b • P! f) = 0
|
||||
|
||||
lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) :
|
||||
|
@ -55,18 +55,18 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S)
|
|||
ring
|
||||
|
||||
lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S) :
|
||||
∀ (g : Fin n → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f),
|
||||
∀ (g : Fin n → ℚ) (f : Fin n → ℚ) (_ : S.val = P g + P! f),
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
|
||||
intro g f hS
|
||||
linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1 ) -
|
||||
(lineInCubic_expand h g f hS 1 2 ) / 6
|
||||
|
||||
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
|
||||
/-- We say a `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
|
||||
def LineInCubicPerm (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
|
||||
∀ (M : (FamilyPermutations (2 * n + 1)).group ),
|
||||
LineInCubic ((FamilyPermutations (2 * n + 1)).linSolRep M S)
|
||||
|
||||
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
||||
/-- If `lineInCubicPerm S` then `lineInCubic S`. -/
|
||||
lemma lineInCubicPerm_self {S : (PureU1 (2 * n + 1)).LinSols} (hS : LineInCubicPerm S) :
|
||||
LineInCubic S := hS 1
|
||||
|
||||
|
@ -89,7 +89,7 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ + 1)).LinSols}
|
|||
(S.val (δ!₂ j) - S.val (δ!₁ j))
|
||||
* accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by
|
||||
intro j g f h
|
||||
let S' := (FamilyPermutations (2 * n.succ + 1)).linSolRep
|
||||
let S' := (FamilyPermutations (2 * n.succ + 1)).linSolRep
|
||||
(pairSwap (δ!₁ j) (δ!₂ j)) S
|
||||
have hSS' : ((FamilyPermutations (2 * n.succ + 1)).linSolRep
|
||||
(pairSwap (δ!₁ j) (δ!₂ j))) S = S' := rfl
|
||||
|
@ -119,7 +119,7 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
|||
have h2 := Pa_δa₂ f g 0
|
||||
rw [← hS] at h1 h2 h4
|
||||
simp at h2
|
||||
have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0):= by
|
||||
have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0):= by
|
||||
linear_combination -(1 * h1) - 1 * h4 - 1 * h2
|
||||
rw [h5]
|
||||
rw [δa₄_δ!₂]
|
||||
|
@ -133,7 +133,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols}
|
|||
(LIC : LineInCubicPerm S) :
|
||||
LineInPlaneProp ((S.val (δ!₂ 0)), ((S.val (δ!₁ 0)), (S.val δ!₃))) := by
|
||||
obtain ⟨g, f, hfg⟩ := span_basis S
|
||||
have h1 := lineInCubicPerm_swap LIC 0 g f hfg
|
||||
have h1 := lineInCubicPerm_swap LIC 0 g f hfg
|
||||
rw [P_P_P!_accCube' g f hfg] at h1
|
||||
simp at h1
|
||||
cases h1 <;> rename_i h1
|
||||
|
@ -147,7 +147,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols}
|
|||
apply Or.inr
|
||||
linear_combination h1
|
||||
|
||||
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
||||
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
||||
(LIC : LineInCubicPerm S) : LineInPlaneCond S := by
|
||||
refine @Prop_three (2 * n.succ.succ + 1) LineInPlaneProp S (δ!₂ 0) (δ!₁ 0)
|
||||
δ!₃ ?_ ?_ ?_ ?_
|
||||
|
@ -157,11 +157,11 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
|||
intro M
|
||||
exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
|
||||
|
||||
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
||||
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
||||
(LIC : LineInCubicPerm S) : ConstAbs S.val :=
|
||||
linesInPlane_constAbs (lineInCubicPerm_last_perm LIC)
|
||||
|
||||
theorem lineInCubicPerm_zero {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
||||
theorem lineInCubicPerm_zero {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
||||
(LIC : LineInCubicPerm S) : S = 0 :=
|
||||
ConstAbs.boundary_value_odd S (lineInCubicPerm_constAbs LIC)
|
||||
|
||||
|
|
|
@ -14,7 +14,7 @@ import Mathlib.RepresentationTheory.Basic
|
|||
/-!
|
||||
# Parameterization in odd case
|
||||
|
||||
Given maps `g : Fin n → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly
|
||||
Given maps `g : Fin n → ℚ`, `f : Fin n → ℚ` and `a : ℚ` we form a solution to the anomaly
|
||||
equations. We show that every solution can be got in this way, up to permutation, unless it is zero.
|
||||
|
||||
The main reference is:
|
||||
|
@ -31,7 +31,7 @@ open VectorLikeOddPlane
|
|||
|
||||
/-- Given a `g f : Fin n → ℚ` and a `a : ℚ` we form a linear solution. We will later
|
||||
show that this can be extended to a complete solution. -/
|
||||
def parameterizationAsLinear (g f : Fin n → ℚ) (a : ℚ) :
|
||||
def parameterizationAsLinear (g f : Fin n → ℚ) (a : ℚ) :
|
||||
(PureU1 (2 * n + 1)).LinSols :=
|
||||
a • ((accCubeTriLinSymm (P! f) (P! f) (P g)) • P' g +
|
||||
(- accCubeTriLinSymm (P g) (P g) (P! f)) • P!' f)
|
||||
|
@ -44,7 +44,7 @@ lemma parameterizationAsLinear_val (g f : Fin n → ℚ) (a : ℚ) :
|
|||
change a • (_ • (P' g).val + _ • (P!' f).val) = _
|
||||
rw [P'_val, P!'_val]
|
||||
|
||||
lemma parameterizationCharge_cube (g f : Fin n → ℚ) (a : ℚ):
|
||||
lemma parameterizationCharge_cube (g f : Fin n → ℚ) (a : ℚ):
|
||||
(accCube (2 * n + 1)) (parameterizationAsLinear g f a).val = 0 := by
|
||||
change accCubeTriLinSymm.toCubic _ = 0
|
||||
rw [parameterizationAsLinear_val]
|
||||
|
@ -64,7 +64,7 @@ def parameterization (g f : Fin n → ℚ) (a : ℚ) :
|
|||
parameterizationCharge_cube g f a⟩
|
||||
|
||||
lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
|
||||
(g f : Fin n → ℚ) (hS : S.val = P g + P! f) :
|
||||
(g f : Fin n → ℚ) (hS : S.val = P g + P! f) :
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) =
|
||||
- accCubeTriLinSymm (P! f) (P! f) (P g) := by
|
||||
have hC := S.cubicSol
|
||||
|
@ -75,15 +75,15 @@ lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
|
|||
erw [P!_accCube] at hC
|
||||
linear_combination hC / 3
|
||||
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
In this case our parameterization above will be able to recover this point. -/
|
||||
def GenericCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
|
||||
∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) ,
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
|
||||
∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) ,
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
|
||||
|
||||
lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
|
||||
(hs : ∃ (g f : Fin n.succ → ℚ), S.val = P g + P! f ∧
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : GenericCase S := by
|
||||
intro g f hS hC
|
||||
obtain ⟨g', f', hS', hC'⟩ := hs
|
||||
rw [hS] at hS'
|
||||
|
@ -91,15 +91,15 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
|
|||
rw [hS'.1, hS'.2] at hC
|
||||
exact hC' hC
|
||||
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
|
||||
In this case we will show that S is zero if it is true for all permutations. -/
|
||||
def SpecialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
|
||||
def SpecialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
|
||||
∀ (g f : Fin n.succ → ℚ) (_ : S.val = P g + P! f) ,
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0
|
||||
|
||||
lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
|
||||
(hs : ∃ (g f : Fin n.succ → ℚ), S.val = P g + P! f ∧
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0) : SpecialCase S := by
|
||||
intro g f hS
|
||||
obtain ⟨g', f', hS', hC'⟩ := hs
|
||||
rw [hS] at hS'
|
||||
|
@ -110,7 +110,7 @@ lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
|
|||
lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) :
|
||||
GenericCase S ∨ SpecialCase S := by
|
||||
obtain ⟨g, f, h⟩ := span_basis S.1.1
|
||||
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨
|
||||
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0 ∨
|
||||
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
|
||||
exact ne_or_eq _ _
|
||||
cases h1 <;> rename_i h1
|
||||
|
@ -118,7 +118,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) :
|
|||
exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩)
|
||||
|
||||
theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
|
||||
∃ g f a, S = parameterization g f a := by
|
||||
∃ g f a, S = parameterization g f a := by
|
||||
obtain ⟨g, f, hS⟩ := span_basis S.1.1
|
||||
use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹
|
||||
rw [parameterization]
|
||||
|
@ -164,7 +164,7 @@ theorem special_case {S : (PureU1 (2 * n.succ.succ + 1)).Sols}
|
|||
(h : ∀ (M : (FamilyPermutations (2 * n.succ.succ + 1)).group),
|
||||
SpecialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun S M)) :
|
||||
S.1.1 = 0 := by
|
||||
have ht := special_case_lineInCubic_perm h
|
||||
have ht := special_case_lineInCubic_perm h
|
||||
exact lineInCubicPerm_zero ht
|
||||
end Odd
|
||||
end PureU1
|
||||
|
|
|
@ -18,13 +18,13 @@ We further define the action on the ACC System.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace PureU1
|
||||
|
||||
/-- The permutation group of the n-fermions. -/
|
||||
@[simp]
|
||||
def PermGroup (n : ℕ) := Equiv.Perm (Fin n)
|
||||
def PermGroup (n : ℕ) := Equiv.Perm (Fin n)
|
||||
|
||||
instance {n : ℕ} : Group (PermGroup n) := by
|
||||
simp [PermGroup]
|
||||
|
@ -35,7 +35,7 @@ section Charges
|
|||
/-- The image of an element of `permGroup` under the representation on charges. -/
|
||||
@[simps!]
|
||||
def chargeMap {n : ℕ} (f : PermGroup n) :
|
||||
(PureU1 n).Charges →ₗ[ℚ] (PureU1 n).Charges where
|
||||
(PureU1 n).Charges →ₗ[ℚ] (PureU1 n).Charges where
|
||||
toFun S := S ∘ f.toFun
|
||||
map_add' S T := by
|
||||
funext i
|
||||
|
@ -73,7 +73,7 @@ open BigOperators
|
|||
lemma accCube_invariant {n : ℕ} (f : (PermGroup n)) (S : (PureU1 n).Charges) :
|
||||
accCube n (permCharges f S) = accCube n S := by
|
||||
rw [accCube_explicit, accCube_explicit]
|
||||
change ∑ i : Fin n, ((((fun a => a^3) ∘ S) (f.symm i))) = _
|
||||
change ∑ i : Fin n, ((((fun a => a^3) ∘ S) (f.symm i))) = _
|
||||
apply (Equiv.Perm.sum_comp _ _ _ ?_)
|
||||
simp
|
||||
|
||||
|
@ -130,22 +130,22 @@ def pairSwap {n : ℕ} (i j : Fin n) : (FamilyPermutations n).group where
|
|||
right_inv s := by
|
||||
aesop
|
||||
|
||||
lemma pairSwap_self_inv {n : ℕ} (i j : Fin n) : (pairSwap i j)⁻¹ = (pairSwap i j) := by
|
||||
lemma pairSwap_self_inv {n : ℕ} (i j : Fin n) : (pairSwap i j)⁻¹ = (pairSwap i j) := by
|
||||
rfl
|
||||
|
||||
lemma pairSwap_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun i = j := by
|
||||
lemma pairSwap_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun i = j := by
|
||||
simp [pairSwap]
|
||||
|
||||
lemma pairSwap_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun j = i := by
|
||||
lemma pairSwap_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).toFun j = i := by
|
||||
simp [pairSwap]
|
||||
|
||||
lemma pairSwap_inv_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun i = j := by
|
||||
lemma pairSwap_inv_fst {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun i = j := by
|
||||
simp [pairSwap]
|
||||
|
||||
lemma pairSwap_inv_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun j = i := by
|
||||
lemma pairSwap_inv_snd {n : ℕ} (i j : Fin n) : (pairSwap i j).invFun j = i := by
|
||||
simp [pairSwap]
|
||||
|
||||
lemma pairSwap_other {n : ℕ} (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) :
|
||||
lemma pairSwap_other {n : ℕ} (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) :
|
||||
(pairSwap i j).toFun k = k := by
|
||||
simp [pairSwap]
|
||||
split
|
||||
|
@ -154,7 +154,7 @@ lemma pairSwap_other {n : ℕ} (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k)
|
|||
simp_all
|
||||
rfl
|
||||
|
||||
lemma pairSwap_inv_other {n : ℕ} {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) :
|
||||
lemma pairSwap_inv_other {n : ℕ} {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) :
|
||||
(pairSwap i j).invFun k = k := by
|
||||
simp [pairSwap]
|
||||
split
|
||||
|
@ -164,12 +164,12 @@ lemma pairSwap_inv_other {n : ℕ} {i j k : Fin n} (hik : i ≠ k) (hjk : j
|
|||
rfl
|
||||
|
||||
/-- A permutation of fermions which takes one ordered subset into another. -/
|
||||
noncomputable def permOfInjection (f g : Fin m ↪ Fin n) : (FamilyPermutations n).group :=
|
||||
noncomputable def permOfInjection (f g : Fin m ↪ Fin n) : (FamilyPermutations n).group :=
|
||||
Equiv.extendSubtype (g.toEquivRange.symm.trans f.toEquivRange)
|
||||
|
||||
section permTwo
|
||||
|
||||
variable {i j i' j' : Fin n} (hij : i ≠ j) (hij' : i' ≠ j')
|
||||
variable {i j i' j' : Fin n} (hij : i ≠ j) (hij' : i' ≠ j')
|
||||
|
||||
/-- Given two distinct elements, an embedding of `Fin 2` into `Fin n`. -/
|
||||
def permTwoInj : Fin 2 ↪ Fin n where
|
||||
|
@ -185,7 +185,7 @@ lemma permTwoInj_fst : i ∈ Set.range ⇑(permTwoInj hij) := by
|
|||
rfl
|
||||
|
||||
lemma permTwoInj_fst_apply :
|
||||
(Function.Embedding.toEquivRange (permTwoInj hij)).symm ⟨i, permTwoInj_fst hij⟩ = 0 := by
|
||||
(Function.Embedding.toEquivRange (permTwoInj hij)).symm ⟨i, permTwoInj_fst hij⟩ = 0 := by
|
||||
exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permTwoInj hij))).mpr rfl
|
||||
|
||||
lemma permTwoInj_snd : j ∈ Set.range ⇑(permTwoInj hij) := by
|
||||
|
@ -195,12 +195,12 @@ lemma permTwoInj_snd : j ∈ Set.range ⇑(permTwoInj hij) := by
|
|||
|
||||
lemma permTwoInj_snd_apply :
|
||||
(Function.Embedding.toEquivRange (permTwoInj hij)).symm
|
||||
⟨j, permTwoInj_snd hij⟩ = 1 := by
|
||||
⟨j, permTwoInj_snd hij⟩ = 1 := by
|
||||
exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permTwoInj hij))).mpr rfl
|
||||
|
||||
/-- A permutation which swaps `i` with `i'` and `j` with `j'`. -/
|
||||
noncomputable def permTwo : (FamilyPermutations n).group :=
|
||||
permOfInjection (permTwoInj hij) (permTwoInj hij')
|
||||
permOfInjection (permTwoInj hij) (permTwoInj hij')
|
||||
|
||||
lemma permTwo_fst : (permTwo hij hij').toFun i' = i := by
|
||||
rw [permTwo, permOfInjection]
|
||||
|
@ -243,7 +243,7 @@ lemma permThreeInj_fst : i ∈ Set.range ⇑(permThreeInj hij hjk hik) := by
|
|||
|
||||
lemma permThreeInj_fst_apply :
|
||||
(Function.Embedding.toEquivRange (permThreeInj hij hjk hik)).symm
|
||||
⟨i, permThreeInj_fst hij hjk hik⟩ = 0 := by
|
||||
⟨i, permThreeInj_fst hij hjk hik⟩ = 0 := by
|
||||
exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permThreeInj hij hjk hik))).mpr rfl
|
||||
|
||||
lemma permThreeInj_snd : j ∈ Set.range ⇑(permThreeInj hij hjk hik) := by
|
||||
|
@ -253,7 +253,7 @@ lemma permThreeInj_snd : j ∈ Set.range ⇑(permThreeInj hij hjk hik) := by
|
|||
|
||||
lemma permThreeInj_snd_apply :
|
||||
(Function.Embedding.toEquivRange (permThreeInj hij hjk hik)).symm
|
||||
⟨j, permThreeInj_snd hij hjk hik⟩ = 1 := by
|
||||
⟨j, permThreeInj_snd hij hjk hik⟩ = 1 := by
|
||||
exact (Equiv.symm_apply_eq (Function.Embedding.toEquivRange (permThreeInj hij hjk hik))).mpr rfl
|
||||
|
||||
lemma permThreeInj_thd : k ∈ Set.range ⇑(permThreeInj hij hjk hik) := by
|
||||
|
@ -268,7 +268,7 @@ lemma permThreeInj_thd_apply :
|
|||
|
||||
/-- A permutation which swaps three distinct elements with another three. -/
|
||||
noncomputable def permThree : (FamilyPermutations n).group :=
|
||||
permOfInjection (permThreeInj hij hjk hik) (permThreeInj hij' hjk' hik')
|
||||
permOfInjection (permThreeInj hij hjk hik) (permThreeInj hij' hjk' hik')
|
||||
|
||||
lemma permThree_fst : (permThree hij hjk hik hij' hjk' hik').toFun i' = i := by
|
||||
rw [permThree, permOfInjection]
|
||||
|
@ -310,7 +310,7 @@ lemma Prop_two (P : ℚ × ℚ → Prop) {S : (PureU1 n).LinSols}
|
|||
have h1 := h (permTwo hij hab ).symm
|
||||
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply] at h1
|
||||
simp at h1
|
||||
change P
|
||||
change P
|
||||
(S.val ((permTwo hij hab ).toFun a),
|
||||
S.val ((permTwo hij hab ).toFun b)) at h1
|
||||
erw [permTwo_fst,permTwo_snd] at h1
|
||||
|
@ -322,14 +322,14 @@ lemma Prop_three (P : ℚ × ℚ × ℚ → Prop) {S : (PureU1 n).LinSols}
|
|||
P ((((FamilyPermutations n).linSolRep f S).val a),(
|
||||
(((FamilyPermutations n).linSolRep f S).val b),
|
||||
(((FamilyPermutations n).linSolRep f S).val c)
|
||||
))) : ∀ (i j k : Fin n) (_ : i ≠ j) (_ : j ≠ k) (_ : i ≠ k),
|
||||
))) : ∀ (i j k : Fin n) (_ : i ≠ j) (_ : j ≠ k) (_ : i ≠ k),
|
||||
P (S.val i, (S.val j, S.val k)) := by
|
||||
intro i j k hij hjk hik
|
||||
have h1 := h (permThree hij hjk hik hab hbc hac).symm
|
||||
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply,
|
||||
FamilyPermutations_anomalyFreeLinear_apply] at h1
|
||||
simp at h1
|
||||
change P
|
||||
change P
|
||||
(S.val ((permThree hij hjk hik hab hbc hac).toFun a),
|
||||
S.val ((permThree hij hjk hik hab hbc hac).toFun b),
|
||||
S.val ((permThree hij hjk hik hab hbc hac).toFun c)) at h1
|
||||
|
|
|
@ -15,7 +15,7 @@ We define the sort function for Pure U(1) charges, and prove some basic properti
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace PureU1
|
||||
|
||||
|
@ -23,15 +23,15 @@ variable {n : ℕ}
|
|||
|
||||
/-- We say a charge is shorted if for all `i ≤ j`, then `S i ≤ S j`. -/
|
||||
@[simp]
|
||||
def Sorted {n : ℕ} (S : (PureU1 n).Charges) : Prop :=
|
||||
∀ i j (_ : i ≤ j), S i ≤ S j
|
||||
def Sorted {n : ℕ} (S : (PureU1 n).Charges) : Prop :=
|
||||
∀ i j (_ : i ≤ j), S i ≤ S j
|
||||
|
||||
/-- Given a charge assignment `S`, the corresponding sorted charge assignment. -/
|
||||
@[simp]
|
||||
def sort {n : ℕ} (S : (PureU1 n).Charges) : (PureU1 n).Charges :=
|
||||
((FamilyPermutations n).rep (Tuple.sort S).symm S)
|
||||
|
||||
lemma sort_sorted {n : ℕ} (S : (PureU1 n).Charges) : Sorted (sort S) := by
|
||||
lemma sort_sorted {n : ℕ} (S : (PureU1 n).Charges) : Sorted (sort S) := by
|
||||
simp only [Sorted, PureU1_numberCharges, sort, FamilyPermutations, PermGroup, permCharges,
|
||||
MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply]
|
||||
intro i j hij
|
||||
|
@ -45,7 +45,7 @@ lemma sort_apply {n : ℕ} (S : (PureU1 n).Charges) (j : Fin n) :
|
|||
sort S j = S ((Tuple.sort S) j) := by
|
||||
rfl
|
||||
|
||||
lemma sort_zero {n : ℕ} (S : (PureU1 n).Charges) (hS : sort S = 0) : S = 0 := by
|
||||
lemma sort_zero {n : ℕ} (S : (PureU1 n).Charges) (hS : sort S = 0) : S = 0 := by
|
||||
funext i
|
||||
have hj : ∀ j, sort S j = 0 := by
|
||||
rw [hS]
|
||||
|
@ -61,13 +61,13 @@ lemma sort_projection {n : ℕ} (S : (PureU1 n).Charges) : sort (sort S) = sort
|
|||
sort_perm S (Tuple.sort S).symm
|
||||
|
||||
/-- The sort function acting on `LinSols`. -/
|
||||
def sortAFL {n : ℕ} (S : (PureU1 n).LinSols) : (PureU1 n).LinSols :=
|
||||
def sortAFL {n : ℕ} (S : (PureU1 n).LinSols) : (PureU1 n).LinSols :=
|
||||
((FamilyPermutations n).linSolRep (Tuple.sort S.val).symm S)
|
||||
|
||||
lemma sortAFL_val {n : ℕ} (S : (PureU1 n).LinSols) : (sortAFL S).val = sort S.val := by
|
||||
lemma sortAFL_val {n : ℕ} (S : (PureU1 n).LinSols) : (sortAFL S).val = sort S.val := by
|
||||
rfl
|
||||
|
||||
lemma sortAFL_zero {n : ℕ} (S : (PureU1 n).LinSols) (hS : sortAFL S = 0) : S = 0 := by
|
||||
lemma sortAFL_zero {n : ℕ} (S : (PureU1 n).LinSols) (hS : sortAFL S = 0) : S = 0 := by
|
||||
apply ACCSystemLinear.LinSols.ext
|
||||
have h1 : sort S.val = 0 := by
|
||||
rw [← sortAFL_val]
|
||||
|
|
|
@ -15,7 +15,7 @@ For the `n`-even case we define the property of a charge assignment being vector
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
open BigOperators
|
||||
|
||||
namespace PureU1
|
||||
|
@ -30,10 +30,10 @@ lemma split_equal (n : ℕ) : n + n = 2 * n := (Nat.two_mul n).symm
|
|||
lemma split_odd (n : ℕ) : n + 1 + n = 2 * n + 1 := by omega
|
||||
|
||||
/-- A charge configuration for n even is vector like if when sorted the `i`th element
|
||||
is equal to the negative of the `n + i`th element. -/
|
||||
is equal to the negative of the `n + i`th element. -/
|
||||
@[simp]
|
||||
def VectorLikeEven (S : (PureU1 (2 * n)).Charges) : Prop :=
|
||||
∀ (i : Fin n), (sort S) (Fin.cast (split_equal n) (Fin.castAdd n i))
|
||||
= - (sort S) (Fin.cast (split_equal n) (Fin.natAdd n i))
|
||||
def VectorLikeEven (S : (PureU1 (2 * n)).Charges) : Prop :=
|
||||
∀ (i : Fin n), (sort S) (Fin.cast (split_equal n) (Fin.castAdd n i))
|
||||
= - (sort S) (Fin.cast (split_equal n) (Fin.natAdd n i))
|
||||
|
||||
end PureU1
|
||||
|
|
|
@ -31,7 +31,7 @@ def SMSpecies (n : ℕ) : ACCSystemCharges := ACCSystemChargesMk n
|
|||
|
||||
namespace SMCharges
|
||||
|
||||
variable {n : ℕ}
|
||||
variable {n : ℕ}
|
||||
|
||||
/-- An equivalence between the set `(SMCharges n).charges` and the set
|
||||
`(Fin 5 → Fin n → ℚ)`. -/
|
||||
|
@ -54,7 +54,7 @@ lemma charges_eq_toSpecies_eq (S T : (SMCharges n).Charges) :
|
|||
apply toSpeciesEquiv.injective
|
||||
exact (Set.eqOn_univ (toSpeciesEquiv S) (toSpeciesEquiv T)).mp fun ⦃x⦄ _ => h x
|
||||
|
||||
lemma toSMSpecies_toSpecies_inv (i : Fin 5) (f : (Fin 5 → Fin n → ℚ) ) :
|
||||
lemma toSMSpecies_toSpecies_inv (i : Fin 5) (f : (Fin 5 → Fin n → ℚ) ) :
|
||||
(toSpecies i) (toSpeciesEquiv.symm f) = f i := by
|
||||
change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i= f i
|
||||
simp
|
||||
|
@ -80,7 +80,7 @@ namespace SMACCs
|
|||
|
||||
open SMCharges
|
||||
|
||||
variable {n : ℕ}
|
||||
variable {n : ℕ}
|
||||
|
||||
/-- The gravitational anomaly equation. -/
|
||||
@[simp]
|
||||
|
@ -89,7 +89,7 @@ def accGrav : (SMCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
|
@ -103,7 +103,7 @@ def accGrav : (SMCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accGrav`. -/
|
||||
lemma accGrav_ext {S T : (SMCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accGrav S = accGrav T := by
|
||||
simp only [accGrav, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
|
||||
AddHom.coe_mk]
|
||||
|
@ -133,7 +133,7 @@ def accSU2 : (SMCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accSU2`. -/
|
||||
lemma accSU2_ext {S T : (SMCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accSU2 S = accSU2 T := by
|
||||
simp only [accSU2, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
|
||||
AddHom.coe_mk]
|
||||
|
@ -162,7 +162,7 @@ def accSU3 : (SMCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accSU3`. -/
|
||||
lemma accSU3_ext {S T : (SMCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accSU3 S = accSU3 T := by
|
||||
simp only [accSU3, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
|
||||
AddHom.coe_mk]
|
||||
|
@ -193,7 +193,7 @@ def accYY : (SMCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
|
||||
/-- Extensionality lemma for `accYY`. -/
|
||||
lemma accYY_ext {S T : (SMCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 5), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accYY S = accYY T := by
|
||||
simp only [accYY, SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
|
||||
AddHom.coe_mk]
|
||||
|
@ -238,7 +238,7 @@ def quadBiLin : BiLinearSymm (SMCharges n).Charges := BiLinearSymm.mk₂
|
|||
|
||||
/-- The quadratic anomaly cancellation condition. -/
|
||||
@[simp]
|
||||
def accQuad : HomogeneousQuadratic (SMCharges n).Charges :=
|
||||
def accQuad : HomogeneousQuadratic (SMCharges n).Charges :=
|
||||
(@quadBiLin n).toHomogeneousQuad
|
||||
|
||||
/-- Extensionality lemma for `accQuad`. -/
|
||||
|
@ -263,7 +263,7 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
|
|||
+ 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i))
|
||||
+ 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i))
|
||||
+ 2 * ((L S.1 i) * (L S.2.1 i) * (L S.2.2 i))
|
||||
+ ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i))))
|
||||
+ ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i))))
|
||||
(by
|
||||
intro a S T R
|
||||
simp only
|
||||
|
|
|
@ -37,7 +37,7 @@ def chargesMapOfSpeciesMap {n m : ℕ} (f : (SMSpecies n).Charges →ₗ[ℚ] (S
|
|||
rw [map_smul]
|
||||
rfl
|
||||
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/
|
||||
@[simps!]
|
||||
def speciesFamilyProj {m n : ℕ} (h : n ≤ m) :
|
||||
(SMSpecies m).Charges →ₗ[ℚ] (SMSpecies n).Charges where
|
||||
|
@ -50,12 +50,12 @@ def speciesFamilyProj {m n : ℕ} (h : n ≤ m) :
|
|||
simp [HSMul.hSMul]
|
||||
--rw [show Rat.cast a = a from rfl]
|
||||
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges. -/
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges. -/
|
||||
def familyProjection {m n : ℕ} (h : n ≤ m) : (SMCharges m).Charges →ₗ[ℚ] (SMCharges n).Charges :=
|
||||
chargesMapOfSpeciesMap (speciesFamilyProj h)
|
||||
|
||||
/-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all
|
||||
other charges zero. -/
|
||||
other charges zero. -/
|
||||
@[simps!]
|
||||
def speciesEmbed (m n : ℕ) :
|
||||
(SMSpecies m).Charges →ₗ[ℚ] (SMSpecies n).Charges where
|
||||
|
@ -80,7 +80,7 @@ def speciesEmbed (m n : ℕ) :
|
|||
simp
|
||||
|
||||
/-- The embedding of the `m`-family charges onto the `n`-family charges, with all
|
||||
other charges zero. -/
|
||||
other charges zero. -/
|
||||
def familyEmbedding (m n : ℕ) : (SMCharges m).Charges →ₗ[ℚ] (SMCharges n).Charges :=
|
||||
chargesMapOfSpeciesMap (speciesEmbed m n)
|
||||
|
||||
|
|
|
@ -36,17 +36,17 @@ namespace SMNoGrav
|
|||
|
||||
variable {n : ℕ}
|
||||
|
||||
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
|
||||
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
|
||||
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 1
|
||||
|
||||
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
|
||||
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
|
||||
exact S.cubicSol
|
||||
|
||||
/-- An element of `charges` which satisfies the linear ACCs
|
||||
|
@ -84,7 +84,7 @@ def chargeToAF (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S
|
|||
(hc : accCube S = 0) : (SMNoGrav n).Sols :=
|
||||
quadToAF (chargeToQuad S hSU2 hSU3) hc
|
||||
|
||||
/-- An element of `AnomalyFreeLinear` which satisfies the quadratic and cubic ACCs
|
||||
/-- An element of `AnomalyFreeLinear` which satisfies the quadratic and cubic ACCs
|
||||
gives us a element of `AnomalyFree`. -/
|
||||
def linearToAF (S : (SMNoGrav n).LinSols)
|
||||
(hc : accCube S.val = 0) : (SMNoGrav n).Sols :=
|
||||
|
|
|
@ -25,11 +25,11 @@ open SMACCs
|
|||
open BigOperators
|
||||
|
||||
lemma E_zero_iff_Q_zero {S : (SMNoGrav 1).Sols} : Q S.val (0 : Fin 1) = 0 ↔
|
||||
E S.val (0 : Fin 1) = 0 := by
|
||||
E S.val (0 : Fin 1) = 0 := by
|
||||
let S' := linearParameters.bijection.symm S.1.1
|
||||
have hC := cubeSol S
|
||||
have hS' := congrArg (fun S => S.val) (linearParameters.bijection.right_inv S.1.1)
|
||||
change S'.asCharges = S.val at hS'
|
||||
change S'.asCharges = S.val at hS'
|
||||
rw [← hS'] at hC
|
||||
apply Iff.intro
|
||||
intro hQ
|
||||
|
@ -37,7 +37,7 @@ lemma E_zero_iff_Q_zero {S : (SMNoGrav 1).Sols} : Q S.val (0 : Fin 1) = 0 ↔
|
|||
intro hE
|
||||
exact S'.cubic_zero_E'_zero hC hE
|
||||
|
||||
lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) :
|
||||
lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) :
|
||||
accGrav S.val = 0 := by
|
||||
rw [accGrav]
|
||||
simp only [SMSpecies_numberCharges, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
|
@ -59,12 +59,12 @@ lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0
|
|||
have hC := cubeSol S
|
||||
have hS' := congrArg (fun S => S.val.val)
|
||||
(linearParametersQENeqZero.bijection.right_inv ⟨S.1.1, And.intro hQ hE⟩)
|
||||
change _ = S.val at hS'
|
||||
change _ = S.val at hS'
|
||||
rw [← hS'] at hC
|
||||
rw [← hS']
|
||||
exact S'.grav_of_cubic hC FLTThree
|
||||
|
||||
/-- Any solution to the ACCs without gravity satisfies the gravitational ACC. -/
|
||||
/-- Any solution to the ACCs without gravity satisfies the gravitational ACC. -/
|
||||
theorem accGravSatisfied {S : (SMNoGrav 1).Sols} (FLTThree : FermatLastTheoremWith ℚ 3) :
|
||||
accGrav S.val = 0 := by
|
||||
by_cases hQ : Q S.val (0 : Fin 1)= 0
|
||||
|
|
|
@ -198,15 +198,15 @@ namespace linearParametersQENeqZero
|
|||
|
||||
@[ext]
|
||||
lemma ext {S T : linearParametersQENeqZero} (hx : S.x = T.x) (hv : S.v = T.v)
|
||||
(hw : S.w = T.w) : S = T := by
|
||||
(hw : S.w = T.w) : S = T := by
|
||||
cases' S
|
||||
simp_all only
|
||||
|
||||
/-- A map from `linearParametersQENeqZero` to `linearParameters`. -/
|
||||
@[simps!]
|
||||
def toLinearParameters (S : linearParametersQENeqZero) :
|
||||
{S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} :=
|
||||
⟨⟨S.x, 3 * S.x * (S.v - S.w) / (S.v + S.w), - 6 * S.x / (S.v + S.w)⟩,
|
||||
{S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} :=
|
||||
⟨⟨S.x, 3 * S.x * (S.v - S.w) / (S.v + S.w), - 6 * S.x / (S.v + S.w)⟩,
|
||||
by
|
||||
apply And.intro S.hx
|
||||
simp only [neg_mul, ne_eq, div_eq_zero_iff, neg_eq_zero, mul_eq_zero, OfNat.ofNat_ne_zero,
|
||||
|
@ -217,9 +217,9 @@ def toLinearParameters (S : linearParametersQENeqZero) :
|
|||
/-- A map from `linearParameters` to `linearParametersQENeqZero` in the special case when
|
||||
`Q'` and `E'` of the linear parameters are non-zero. -/
|
||||
@[simps!]
|
||||
def tolinearParametersQNeqZero (S : {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0}) :
|
||||
def tolinearParametersQNeqZero (S : {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0}) :
|
||||
linearParametersQENeqZero :=
|
||||
⟨S.1.Q', - (3 * S.1.Q' + S.1.Y) / S.1.E', - (3 * S.1.Q' - S.1.Y)/ S.1.E', S.2.1,
|
||||
⟨S.1.Q', - (3 * S.1.Q' + S.1.Y) / S.1.E', - (3 * S.1.Q' - S.1.Y)/ S.1.E', S.2.1,
|
||||
by
|
||||
simp only [ne_eq, neg_add_rev, neg_sub]
|
||||
field_simp
|
||||
|
@ -231,7 +231,7 @@ def tolinearParametersQNeqZero (S : {S : linearParameters // S.Q' ≠ 0 ∧ S.
|
|||
with `Q'` and `E'` non-zero. -/
|
||||
@[simps!]
|
||||
def bijectionLinearParameters :
|
||||
linearParametersQENeqZero ≃ {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} where
|
||||
linearParametersQENeqZero ≃ {S : linearParameters // S.Q' ≠ 0 ∧ S.E' ≠ 0} where
|
||||
toFun := toLinearParameters
|
||||
invFun := tolinearParametersQNeqZero
|
||||
left_inv S := by
|
||||
|
@ -260,7 +260,7 @@ def bijectionLinearParameters :
|
|||
|
||||
/-- The bijection between `linearParametersQENeqZero` and `LinSols` with `Q` and `E` non-zero. -/
|
||||
def bijection : linearParametersQENeqZero ≃
|
||||
{S : (SMNoGrav 1).LinSols // Q S.val (0 : Fin 1) ≠ 0 ∧ E S.val (0 : Fin 1) ≠ 0} :=
|
||||
{S : (SMNoGrav 1).LinSols // Q S.val (0 : Fin 1) ≠ 0 ∧ E S.val (0 : Fin 1) ≠ 0} :=
|
||||
bijectionLinearParameters.trans (linearParameters.bijectionQEZero)
|
||||
|
||||
lemma cubic (S : linearParametersQENeqZero) :
|
||||
|
@ -297,7 +297,7 @@ lemma cubic_v_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.
|
|||
have h' : (S.w + 1) * (1 * S.w * S.w + (-1) * S.w + 1) = 0 := by
|
||||
ring_nf
|
||||
exact add_eq_zero_iff_neg_eq.mpr (id (Eq.symm h))
|
||||
have h'' : (1 * S.w * S.w + (-1) * S.w + 1) ≠ 0 := by
|
||||
have h'' : (1 * S.w * S.w + (-1) * S.w + 1) ≠ 0 := by
|
||||
refine quadratic_ne_zero_of_discrim_ne_sq ?_ S.w
|
||||
intro s
|
||||
by_contra hn
|
||||
|
@ -315,7 +315,7 @@ lemma cube_w_zero (S : linearParametersQENeqZero) (h : accCube (bijection S).1.v
|
|||
have h' : (S.v + 1) * (1 * S.v * S.v + (-1) * S.v + 1) = 0 := by
|
||||
ring_nf
|
||||
exact add_eq_zero_iff_neg_eq.mpr (id (Eq.symm h))
|
||||
have h'' : (1 * S.v * S.v + (-1) * S.v + 1) ≠ 0 := by
|
||||
have h'' : (1 * S.v * S.v + (-1) * S.v + 1) ≠ 0 := by
|
||||
refine quadratic_ne_zero_of_discrim_ne_sq ?_ S.v
|
||||
intro s
|
||||
by_contra hn
|
||||
|
|
|
@ -16,7 +16,7 @@ We define the group of permutations for the SM charges with no RHN.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace SM
|
||||
|
||||
|
@ -26,7 +26,7 @@ open BigOperators
|
|||
|
||||
/-- The group of `Sₙ` permutations for each species. -/
|
||||
@[simp]
|
||||
def PermGroup (n : ℕ) := ∀ (_ : Fin 5), Equiv.Perm (Fin n)
|
||||
def PermGroup (n : ℕ) := ∀ (_ : Fin 5), Equiv.Perm (Fin n)
|
||||
|
||||
variable {n : ℕ}
|
||||
|
||||
|
@ -72,32 +72,32 @@ lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup n) (S : (SMCharges n).Cha
|
|||
exact Fintype.sum_equiv (f⁻¹ j) (fun x => ((fun a => a ^ m) ∘ (toSpecies j) S ∘ ⇑(f⁻¹ j)) x)
|
||||
(fun x => ((fun a => a ^ m) ∘ (toSpecies j) S) x) (congrFun rfl)
|
||||
|
||||
lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
accGrav (repCharges f S) = accGrav S :=
|
||||
accGrav_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accSU2_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
lemma accSU2_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
accSU2 (repCharges f S) = accSU2 S :=
|
||||
accSU2_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accSU3_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
lemma accSU3_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
accSU3 (repCharges f S) = accSU3 S :=
|
||||
accSU3_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accYY_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
lemma accYY_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
accYY (repCharges f S) = accYY S :=
|
||||
accYY_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accQuad_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
lemma accQuad_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
accQuad (repCharges f S) = accQuad S :=
|
||||
accQuad_ext
|
||||
(toSpecies_sum_invariant 2 f S)
|
||||
|
||||
lemma accCube_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
lemma accCube_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
|
||||
accCube (repCharges f S) = accCube S :=
|
||||
accCube_ext (fun j => toSpecies_sum_invariant 3 f S j)
|
||||
|
||||
|
|
|
@ -28,7 +28,7 @@ def SMνSpecies (n : ℕ) : ACCSystemCharges := ACCSystemChargesMk n
|
|||
|
||||
namespace SMνCharges
|
||||
|
||||
variable {n : ℕ}
|
||||
variable {n : ℕ}
|
||||
|
||||
/-- An equivalence between `(SMνCharges n).charges` and `(Fin 6 → Fin n → ℚ)`
|
||||
splitting the charges into species.-/
|
||||
|
@ -54,12 +54,12 @@ lemma charges_eq_toSpecies_eq (S T : (SMνCharges n).Charges) :
|
|||
funext i
|
||||
exact h i
|
||||
|
||||
lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin n → ℚ) ) :
|
||||
lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin n → ℚ) ) :
|
||||
(toSpecies i) (toSpeciesEquiv.symm f) = f i := by
|
||||
change (toSpeciesEquiv ∘ toSpeciesEquiv.symm ) _ i = f i
|
||||
simp
|
||||
|
||||
lemma toSpecies_one (S : (SMνCharges 1).Charges) (j : Fin 6) :
|
||||
lemma toSpecies_one (S : (SMνCharges 1).Charges) (j : Fin 6) :
|
||||
toSpecies j S ⟨0, by simp⟩ = S j := by
|
||||
match j with
|
||||
| 0 => rfl
|
||||
|
@ -88,7 +88,7 @@ namespace SMνACCs
|
|||
|
||||
open SMνCharges
|
||||
|
||||
variable {n : ℕ}
|
||||
variable {n : ℕ}
|
||||
|
||||
/-- The gravitational anomaly equation. -/
|
||||
@[simp]
|
||||
|
@ -97,7 +97,7 @@ def accGrav : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
|
@ -119,7 +119,7 @@ lemma accGrav_decomp (S : (SMνCharges n).Charges) :
|
|||
|
||||
/-- Extensionality lemma for `accGrav`. -/
|
||||
lemma accGrav_ext {S T : (SMνCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accGrav S = accGrav T := by
|
||||
rw [accGrav_decomp, accGrav_decomp]
|
||||
repeat erw [hj]
|
||||
|
@ -131,7 +131,7 @@ def accSU2 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
|
@ -152,7 +152,7 @@ lemma accSU2_decomp (S : (SMνCharges n).Charges) :
|
|||
|
||||
/-- Extensionality lemma for `accSU2`. -/
|
||||
lemma accSU2_ext {S T : (SMνCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accSU2 S = accSU2 T := by
|
||||
rw [accSU2_decomp, accSU2_decomp]
|
||||
repeat erw [hj]
|
||||
|
@ -164,7 +164,7 @@ def accSU3 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [ Pi.add_apply, mul_add]
|
||||
simp [ Pi.add_apply, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
|
@ -185,7 +185,7 @@ lemma accSU3_decomp (S : (SMνCharges n).Charges) :
|
|||
|
||||
/-- Extensionality lemma for `accSU3`. -/
|
||||
lemma accSU3_ext {S T : (SMνCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accSU3 S = accSU3 T := by
|
||||
rw [accSU3_decomp, accSU3_decomp]
|
||||
repeat rw [hj]
|
||||
|
@ -198,7 +198,7 @@ def accYY : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
|
@ -219,7 +219,7 @@ lemma accYY_decomp (S : (SMνCharges n).Charges) :
|
|||
|
||||
/-- Extensionality lemma for `accYY`. -/
|
||||
lemma accYY_ext {S T : (SMνCharges n).Charges}
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
|
||||
accYY S = accYY T := by
|
||||
rw [accYY_decomp, accYY_decomp]
|
||||
repeat rw [hj]
|
||||
|
@ -259,8 +259,8 @@ def quadBiLin : BiLinearSymm (SMνCharges n).Charges := BiLinearSymm.mk₂
|
|||
ring)
|
||||
|
||||
lemma quadBiLin_decomp (S T : (SMνCharges n).Charges) :
|
||||
quadBiLin S T = ∑ i, Q S i * Q T i - 2 * ∑ i, U S i * U T i +
|
||||
∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by
|
||||
quadBiLin S T = ∑ i, Q S i * Q T i - 2 * ∑ i, U S i * U T i +
|
||||
∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by
|
||||
erw [← quadBiLin.toFun_eq_coe]
|
||||
rw [quadBiLin]
|
||||
simp only [quadBiLin, BiLinearSymm.mk₂, AddHom.toFun_eq_coe, AddHom.coe_mk, LinearMap.coe_mk]
|
||||
|
@ -271,7 +271,7 @@ lemma quadBiLin_decomp (S T : (SMνCharges n).Charges) :
|
|||
|
||||
/-- The quadratic anomaly cancellation condition. -/
|
||||
@[simp]
|
||||
def accQuad : HomogeneousQuadratic (SMνCharges n).Charges :=
|
||||
def accQuad : HomogeneousQuadratic (SMνCharges n).Charges :=
|
||||
(@quadBiLin n).toHomogeneousQuad
|
||||
|
||||
lemma accQuad_decomp (S : (SMνCharges n).Charges) :
|
||||
|
@ -296,8 +296,8 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃
|
|||
+ 3 * ((U S.1 i) * (U S.2.1 i) * (U S.2.2 i))
|
||||
+ 3 * ((D S.1 i) * (D S.2.1 i) * (D S.2.2 i))
|
||||
+ 2 * ((L S.1 i) * (L S.2.1 i) * (L S.2.2 i))
|
||||
+ ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i))
|
||||
+ ((N S.1 i) * (N S.2.1 i) * (N S.2.2 i))))
|
||||
+ ((E S.1 i) * (E S.2.1 i) * (E S.2.2 i))
|
||||
+ ((N S.1 i) * (N S.2.1 i) * (N S.2.2 i))))
|
||||
(by
|
||||
intro a S T R
|
||||
simp only
|
||||
|
@ -330,8 +330,8 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃
|
|||
ring)
|
||||
|
||||
lemma cubeTriLin_decomp (S T R : (SMνCharges n).Charges) :
|
||||
cubeTriLin S T R = 6 * ∑ i, (Q S i * Q T i * Q R i) + 3 * ∑ i, (U S i * U T i * U R i) +
|
||||
3 * ∑ i, (D S i * D T i * D R i) + 2 * ∑ i, (L S i * L T i * L R i) +
|
||||
cubeTriLin S T R = 6 * ∑ i, (Q S i * Q T i * Q R i) + 3 * ∑ i, (U S i * U T i * U R i) +
|
||||
3 * ∑ i, (D S i * D T i * D R i) + 2 * ∑ i, (L S i * L T i * L R i) +
|
||||
∑ i, (E S i * E T i * E R i) + ∑ i, (N S i * N T i * N R i) := by
|
||||
erw [← cubeTriLin.toFun_eq_coe]
|
||||
rw [cubeTriLin]
|
||||
|
|
|
@ -5,7 +5,7 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import HepLean.AnomalyCancellation.SMNu.Basic
|
||||
/-!
|
||||
# Family maps for the Standard Model for RHN ACCs
|
||||
# Family maps for the Standard Model for RHN ACCs
|
||||
|
||||
We define the a series of maps between the charges for different numbers of families.
|
||||
|
||||
|
@ -39,10 +39,10 @@ def chargesMapOfSpeciesMap {n m : ℕ} (f : (SMνSpecies n).Charges →ₗ[ℚ]
|
|||
lemma chargesMapOfSpeciesMap_toSpecies {n m : ℕ}
|
||||
(f : (SMνSpecies n).Charges →ₗ[ℚ] (SMνSpecies m).Charges)
|
||||
(S : (SMνCharges n).Charges) (j : Fin 6) :
|
||||
toSpecies j (chargesMapOfSpeciesMap f S) = (LinearMap.comp f (toSpecies j)) S := by
|
||||
toSpecies j (chargesMapOfSpeciesMap f S) = (LinearMap.comp f (toSpecies j)) S := by
|
||||
erw [toSMSpecies_toSpecies_inv]
|
||||
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/
|
||||
@[simps!]
|
||||
def speciesFamilyProj {m n : ℕ} (h : n ≤ m) :
|
||||
(SMνSpecies m).Charges →ₗ[ℚ] (SMνSpecies n).Charges where
|
||||
|
@ -50,12 +50,12 @@ def speciesFamilyProj {m n : ℕ} (h : n ≤ m) :
|
|||
map_add' _ _ := rfl
|
||||
map_smul' _ _ := rfl
|
||||
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges. -/
|
||||
/-- The projection of the `m`-family charges onto the first `n`-family charges. -/
|
||||
def familyProjection {m n : ℕ} (h : n ≤ m) : (SMνCharges m).Charges →ₗ[ℚ] (SMνCharges n).Charges :=
|
||||
chargesMapOfSpeciesMap (speciesFamilyProj h)
|
||||
|
||||
/-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all
|
||||
other charges zero. -/
|
||||
other charges zero. -/
|
||||
@[simps!]
|
||||
def speciesEmbed (m n : ℕ) :
|
||||
(SMνSpecies m).Charges →ₗ[ℚ] (SMνSpecies n).Charges where
|
||||
|
@ -80,7 +80,7 @@ def speciesEmbed (m n : ℕ) :
|
|||
exact Eq.symm (Rat.mul_zero a)
|
||||
|
||||
/-- The embedding of the `m`-family charges onto the `n`-family charges, with all
|
||||
other charges zero. -/
|
||||
other charges zero. -/
|
||||
def familyEmbedding (m n : ℕ) : (SMνCharges m).Charges →ₗ[ℚ] (SMνCharges n).Charges :=
|
||||
chargesMapOfSpeciesMap (speciesEmbed m n)
|
||||
|
||||
|
@ -118,7 +118,7 @@ lemma sum_familyUniversal {n : ℕ} (m : ℕ) (S : (SMνCharges 1).Charges) (j :
|
|||
intro i _
|
||||
erw [toSpecies_familyUniversal]
|
||||
|
||||
lemma sum_familyUniversal_one {n : ℕ} (S : (SMνCharges 1).Charges) (j : Fin 6) :
|
||||
lemma sum_familyUniversal_one {n : ℕ} (S : (SMνCharges 1).Charges) (j : Fin 6) :
|
||||
∑ i, toSpecies j (familyUniversal n S) i = n * (toSpecies j S ⟨0, by simp⟩) := by
|
||||
simpa using @sum_familyUniversal n 1 S j
|
||||
|
||||
|
@ -134,7 +134,7 @@ lemma sum_familyUniversal_two {n : ℕ} (S : (SMνCharges 1).Charges)
|
|||
erw [toSpecies_familyUniversal]
|
||||
rfl
|
||||
|
||||
lemma sum_familyUniversal_three {n : ℕ} (S : (SMνCharges 1).Charges)
|
||||
lemma sum_familyUniversal_three {n : ℕ} (S : (SMνCharges 1).Charges)
|
||||
(T L : (SMνCharges n).Charges) (j : Fin 6) :
|
||||
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i * toSpecies j L i) =
|
||||
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i * toSpecies j L i := by
|
||||
|
@ -148,7 +148,7 @@ lemma sum_familyUniversal_three {n : ℕ} (S : (SMνCharges 1).Charges)
|
|||
ring
|
||||
|
||||
lemma familyUniversal_accGrav (S : (SMνCharges 1).Charges) :
|
||||
accGrav (familyUniversal n S) = n * (accGrav S) := by
|
||||
accGrav (familyUniversal n S) = n * (accGrav S) := by
|
||||
rw [accGrav_decomp, accGrav_decomp]
|
||||
repeat rw [sum_familyUniversal_one]
|
||||
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
|
||||
|
@ -156,7 +156,7 @@ lemma familyUniversal_accGrav (S : (SMνCharges 1).Charges) :
|
|||
ring
|
||||
|
||||
lemma familyUniversal_accSU2 (S : (SMνCharges 1).Charges) :
|
||||
accSU2 (familyUniversal n S) = n * (accSU2 S) := by
|
||||
accSU2 (familyUniversal n S) = n * (accSU2 S) := by
|
||||
rw [accSU2_decomp, accSU2_decomp]
|
||||
repeat rw [sum_familyUniversal_one]
|
||||
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
|
||||
|
@ -164,7 +164,7 @@ lemma familyUniversal_accSU2 (S : (SMνCharges 1).Charges) :
|
|||
ring
|
||||
|
||||
lemma familyUniversal_accSU3 (S : (SMνCharges 1).Charges) :
|
||||
accSU3 (familyUniversal n S) = n * (accSU3 S) := by
|
||||
accSU3 (familyUniversal n S) = n * (accSU3 S) := by
|
||||
rw [accSU3_decomp, accSU3_decomp]
|
||||
repeat rw [sum_familyUniversal_one]
|
||||
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
|
||||
|
@ -172,7 +172,7 @@ lemma familyUniversal_accSU3 (S : (SMνCharges 1).Charges) :
|
|||
ring
|
||||
|
||||
lemma familyUniversal_accYY (S : (SMνCharges 1).Charges) :
|
||||
accYY (familyUniversal n S) = n * (accYY S) := by
|
||||
accYY (familyUniversal n S) = n * (accYY S) := by
|
||||
rw [accYY_decomp, accYY_decomp]
|
||||
repeat rw [sum_familyUniversal_one]
|
||||
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
|
||||
|
@ -182,7 +182,7 @@ lemma familyUniversal_accYY (S : (SMνCharges 1).Charges) :
|
|||
lemma familyUniversal_quadBiLin (S : (SMνCharges 1).Charges) (T : (SMνCharges n).Charges) :
|
||||
quadBiLin (familyUniversal n S) T =
|
||||
S (0 : Fin 6) * ∑ i, Q T i - 2 * S (1 : Fin 6) * ∑ i, U T i + S (2 : Fin 6) *∑ i, D T i -
|
||||
S (3 : Fin 6) * ∑ i, L T i + S (4 : Fin 6) * ∑ i, E T i := by
|
||||
S (3 : Fin 6) * ∑ i, L T i + S (4 : Fin 6) * ∑ i, E T i := by
|
||||
rw [quadBiLin_decomp]
|
||||
repeat rw [sum_familyUniversal_two]
|
||||
repeat rw [toSpecies_one]
|
||||
|
@ -191,7 +191,7 @@ lemma familyUniversal_quadBiLin (S : (SMνCharges 1).Charges) (T : (SMνCharges
|
|||
ring
|
||||
|
||||
lemma familyUniversal_accQuad (S : (SMνCharges 1).Charges) :
|
||||
accQuad (familyUniversal n S) = n * (accQuad S) := by
|
||||
accQuad (familyUniversal n S) = n * (accQuad S) := by
|
||||
rw [accQuad_decomp]
|
||||
repeat erw [sum_familyUniversal]
|
||||
rw [accQuad_decomp]
|
||||
|
@ -201,7 +201,7 @@ lemma familyUniversal_accQuad (S : (SMνCharges 1).Charges) :
|
|||
|
||||
lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).Charges) (T R : (SMνCharges n).Charges) :
|
||||
cubeTriLin (familyUniversal n S) T R = 6 * S (0 : Fin 6) * ∑ i, (Q T i * Q R i) +
|
||||
3 * S (1 : Fin 6) * ∑ i, (U T i * U R i) + 3 * S (2 : Fin 6) * ∑ i, (D T i * D R i)
|
||||
3 * S (1 : Fin 6) * ∑ i, (U T i * U R i) + 3 * S (2 : Fin 6) * ∑ i, (D T i * D R i)
|
||||
+ 2 * S (3 : Fin 6) * ∑ i, (L T i * L R i) +
|
||||
S (4 : Fin 6) * ∑ i, (E T i * E R i) + S (5 : Fin 6) * ∑ i, (N T i * N R i) := by
|
||||
rw [cubeTriLin_decomp]
|
||||
|
|
|
@ -38,17 +38,17 @@ namespace SMNoGrav
|
|||
|
||||
variable {n : ℕ}
|
||||
|
||||
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
|
||||
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
|
||||
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 1
|
||||
|
||||
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
|
||||
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
|
||||
exact S.cubicSol
|
||||
|
||||
/-- An element of `charges` which satisfies the linear ACCs
|
||||
|
@ -86,7 +86,7 @@ def chargeToAF (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S
|
|||
(hc : accCube S = 0) : (SMNoGrav n).Sols :=
|
||||
quadToAF (chargeToQuad S hSU2 hSU3) hc
|
||||
|
||||
/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs
|
||||
/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs
|
||||
gives us a element of `Sols`. -/
|
||||
def linearToAF (S : (SMNoGrav n).LinSols)
|
||||
(hc : accCube S.val = 0) : (SMNoGrav n).Sols :=
|
||||
|
|
|
@ -39,22 +39,22 @@ namespace SM
|
|||
|
||||
variable {n : ℕ}
|
||||
|
||||
lemma gravSol (S : (SM n).LinSols) : accGrav S.val = 0 := by
|
||||
lemma gravSol (S : (SM n).LinSols) : accGrav S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU2Sol (S : (SM n).LinSols) : accSU2 S.val = 0 := by
|
||||
lemma SU2Sol (S : (SM n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 1
|
||||
|
||||
lemma SU3Sol (S : (SM n).LinSols) : accSU3 S.val = 0 := by
|
||||
lemma SU3Sol (S : (SM n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 2
|
||||
|
||||
lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := by
|
||||
lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := by
|
||||
exact S.cubicSol
|
||||
|
||||
/-- An element of `charges` which satisfies the linear ACCs
|
||||
|
@ -94,13 +94,13 @@ def chargeToAF (S : (SM n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0
|
|||
(hSU3 : accSU3 S = 0) (hc : accCube S = 0) : (SM n).Sols :=
|
||||
quadToAF (chargeToQuad S hGrav hSU2 hSU3) hc
|
||||
|
||||
/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs
|
||||
/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs
|
||||
gives us a element of `Sols`. -/
|
||||
def linearToAF (S : (SM n).LinSols)
|
||||
(hc : accCube S.val = 0) : (SM n).Sols :=
|
||||
quadToAF (linearToQuad S) hc
|
||||
|
||||
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
|
||||
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
|
||||
def perm (n : ℕ) : ACCSystemGroupAction (SM n) where
|
||||
group := PermGroup n
|
||||
groupInst := inferInstance
|
||||
|
|
|
@ -32,7 +32,7 @@ def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
|||
)
|
||||
|
||||
lemma B₀_cubic (S T : (SM 3).Charges) : cubeTriLin B₀ S T =
|
||||
6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by
|
||||
6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₀, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
ring
|
||||
|
||||
|
@ -45,7 +45,7 @@ def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
|||
)
|
||||
|
||||
lemma B₁_cubic (S T : (SM 3).Charges) : cubeTriLin B₁ S T =
|
||||
3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by
|
||||
3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₁, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
ring
|
||||
|
||||
|
@ -58,7 +58,7 @@ def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
|||
)
|
||||
|
||||
lemma B₂_cubic (S T : (SM 3).Charges) : cubeTriLin B₂ S T =
|
||||
3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by
|
||||
3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₂, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
ring
|
||||
|
||||
|
@ -71,7 +71,7 @@ def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
|||
)
|
||||
|
||||
lemma B₃_cubic (S T : (SM 3).Charges) : cubeTriLin B₃ S T =
|
||||
2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by
|
||||
2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₃, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
ring_nf
|
||||
rfl
|
||||
|
@ -85,7 +85,7 @@ def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
|||
)
|
||||
|
||||
lemma B₄_cubic (S T : (SM 3).Charges) : cubeTriLin B₄ S T =
|
||||
(S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by
|
||||
(S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₄, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
ring_nf
|
||||
rfl
|
||||
|
@ -99,7 +99,7 @@ def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
|||
)
|
||||
|
||||
lemma B₅_cubic (S T : (SM 3).Charges) : cubeTriLin B₅ S T =
|
||||
(S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by
|
||||
(S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₅, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
ring_nf
|
||||
rfl
|
||||
|
@ -113,7 +113,7 @@ def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
|
|||
)
|
||||
|
||||
lemma B₆_cubic (S T : (SM 3).Charges) : cubeTriLin B₆ S T =
|
||||
3* (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by
|
||||
3 * (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
ring_nf
|
||||
|
||||
|
@ -336,7 +336,7 @@ theorem basis_linear_independent : LinearIndependent ℚ B := by
|
|||
end PlaneSeven
|
||||
|
||||
theorem seven_dim_plane_exists : ∃ (B : Fin 7 → (SM 3).Charges),
|
||||
LinearIndependent ℚ B ∧ ∀ (f : Fin 7 → ℚ), (SM 3).IsSolution (∑ i, f i • B i) := by
|
||||
LinearIndependent ℚ B ∧ ∀ (f : Fin 7 → ℚ), (SM 3).IsSolution (∑ i, f i • B i) := by
|
||||
use PlaneSeven.B
|
||||
apply And.intro
|
||||
exact PlaneSeven.basis_linear_independent
|
||||
|
|
|
@ -16,7 +16,7 @@ We define the group of permutations for the SM charges with RHN.
|
|||
universe v u
|
||||
|
||||
open Nat
|
||||
open Finset
|
||||
open Finset
|
||||
|
||||
namespace SMRHN
|
||||
|
||||
|
@ -69,36 +69,36 @@ lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup n) (S : (SMνCharges n).C
|
|||
∑ i, ((fun a => a ^ m) ∘ toSpecies j (repCharges f S)) i =
|
||||
∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by
|
||||
erw [repCharges_toSpecies]
|
||||
change ∑ i : Fin n, ((fun a => a ^ m) ∘ _) (⇑(f⁻¹ _) i) = ∑ i : Fin n, ((fun a => a ^ m) ∘ _) i
|
||||
change ∑ i : Fin n, ((fun a => a ^ m) ∘ _) (⇑(f⁻¹ _) i) = ∑ i : Fin n, ((fun a => a ^ m) ∘ _) i
|
||||
refine Equiv.Perm.sum_comp _ _ _ ?_
|
||||
simp only [PermGroup, Fin.isValue, Pi.inv_apply, ne_eq, coe_univ, Set.subset_univ]
|
||||
|
||||
lemma accGrav_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
lemma accGrav_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
accGrav (repCharges f S) = accGrav S :=
|
||||
accGrav_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accSU2_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
lemma accSU2_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
accSU2 (repCharges f S) = accSU2 S :=
|
||||
accSU2_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accSU3_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
lemma accSU3_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
accSU3 (repCharges f S) = accSU3 S :=
|
||||
accSU3_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accYY_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
lemma accYY_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
accYY (repCharges f S) = accYY S :=
|
||||
accYY_ext
|
||||
(by simpa using toSpecies_sum_invariant 1 f S)
|
||||
|
||||
lemma accQuad_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
lemma accQuad_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
accQuad (repCharges f S) = accQuad S :=
|
||||
accQuad_ext
|
||||
(toSpecies_sum_invariant 2 f S)
|
||||
|
||||
lemma accCube_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
lemma accCube_invariant (f : PermGroup n) (S : (SMνCharges n).Charges) :
|
||||
accCube (repCharges f S) = accCube S :=
|
||||
accCube_ext
|
||||
(by simpa using toSpecies_sum_invariant 3 f S)
|
||||
|
|
|
@ -39,32 +39,32 @@ namespace PlusU1
|
|||
|
||||
variable {n : ℕ}
|
||||
|
||||
lemma gravSol (S : (PlusU1 n).LinSols) : accGrav S.val = 0 := by
|
||||
lemma gravSol (S : (PlusU1 n).LinSols) : accGrav S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU2Sol (S : (PlusU1 n).LinSols) : accSU2 S.val = 0 := by
|
||||
lemma SU2Sol (S : (PlusU1 n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 1
|
||||
|
||||
lemma SU3Sol (S : (PlusU1 n).LinSols) : accSU3 S.val = 0 := by
|
||||
lemma SU3Sol (S : (PlusU1 n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 2
|
||||
|
||||
lemma YYsol (S : (PlusU1 n).LinSols) : accYY S.val = 0 := by
|
||||
lemma YYsol (S : (PlusU1 n).LinSols) : accYY S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
exact hS 3
|
||||
|
||||
lemma quadSol (S : (PlusU1 n).QuadSols) : accQuad S.val = 0 := by
|
||||
lemma quadSol (S : (PlusU1 n).QuadSols) : accQuad S.val = 0 := by
|
||||
have hS := S.quadSol
|
||||
simp at hS
|
||||
exact hS 0
|
||||
|
||||
lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by
|
||||
lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by
|
||||
exact S.cubicSol
|
||||
|
||||
/-- An element of `charges` which satisfies the linear ACCs
|
||||
|
@ -110,13 +110,13 @@ def chargeToAF (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S
|
|||
(PlusU1 n).Sols :=
|
||||
quadToAF (chargeToQuad S hGrav hSU2 hSU3 hYY hQ) hc
|
||||
|
||||
/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs
|
||||
/-- An element of `LinSols` which satisfies the quadratic and cubic ACCs
|
||||
gives us a element of `Sols`. -/
|
||||
def linearToAF (S : (PlusU1 n).LinSols) (hQ : accQuad S.val = 0)
|
||||
(hc : accCube S.val = 0) : (PlusU1 n).Sols :=
|
||||
quadToAF (linearToQuad S hQ) hc
|
||||
|
||||
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
|
||||
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
|
||||
def perm (n : ℕ) : ACCSystemGroupAction (PlusU1 n) where
|
||||
group := PermGroup n
|
||||
groupInst := inferInstance
|
||||
|
|
|
@ -26,11 +26,11 @@ def Y₁ : (PlusU1 1).Sols where
|
|||
val := fun i =>
|
||||
match i with
|
||||
| (0 : Fin 6) => 1
|
||||
| (1 : Fin 6) => -4
|
||||
| (2 : Fin 6) => 2
|
||||
| (3 : Fin 6) => -3
|
||||
| (4 : Fin 6) => 6
|
||||
| (5 : Fin 6) => 0
|
||||
| (1 : Fin 6) => -4
|
||||
| (2 : Fin 6) => 2
|
||||
| (3 : Fin 6) => -3
|
||||
| (4 : Fin 6) => 6
|
||||
| (5 : Fin 6) => 0
|
||||
linearSol := by
|
||||
intro i
|
||||
simp at i
|
||||
|
@ -89,7 +89,7 @@ lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ℚ): addQuad S a 0 = a • S
|
|||
rfl
|
||||
|
||||
lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
|
||||
cubeTriLin (Y n).val (Y n).val S = 6 * accYY S := by
|
||||
cubeTriLin (Y n).val (Y n).val S = 6 * accYY S := by
|
||||
erw [familyUniversal_cubeTriLin']
|
||||
rw [accYY_decomp]
|
||||
simp only [Fin.isValue, Y₁_val, mul_one, SMνSpecies_numberCharges, toSpecies_apply, mul_neg,
|
||||
|
@ -103,7 +103,7 @@ lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
|
|||
simp
|
||||
|
||||
lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
|
||||
cubeTriLin (Y n).val S S = 6 * accQuad S := by
|
||||
cubeTriLin (Y n).val S S = 6 * accQuad S := by
|
||||
erw [familyUniversal_cubeTriLin]
|
||||
rw [accQuad_decomp]
|
||||
simp only [Fin.isValue, Y₁_val, mul_one, SMνSpecies_numberCharges, toSpecies_apply, mul_neg,
|
||||
|
@ -127,7 +127,7 @@ lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ℚ) :
|
|||
ring
|
||||
|
||||
lemma add_AFQ_cube (S : (PlusU1 n).QuadSols) (a b : ℚ) :
|
||||
accCube (a • S.val + b • (Y n).val) = a ^ 3 * accCube S.val := by
|
||||
accCube (a • S.val + b • (Y n).val) = a ^ 3 * accCube S.val := by
|
||||
rw [add_AFL_cube]
|
||||
rw [cubeTriLin.swap₃]
|
||||
rw [on_cubeTriLin'_ALQ]
|
||||
|
|
|
@ -50,7 +50,7 @@ def B₆ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0,
|
|||
def B₇ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0]
|
||||
|
||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||
def B₈ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0]
|
||||
def B₈ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0]
|
||||
|
||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||
def B₉ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 2, 0]
|
||||
|
@ -79,7 +79,7 @@ lemma Bi_Bj_quad {i j : Fin 11} (hi : i ≠ j) : quadBiLin (B i) (B j) = 0 := by
|
|||
all_goals simp at hi
|
||||
|
||||
lemma Bi_sum_quad (i : Fin 11) (f : Fin 11 → ℚ) :
|
||||
quadBiLin (B i) (∑ k, f k • B k) = f i * quadBiLin (B i) (B i) := by
|
||||
quadBiLin (B i) (∑ k, f k • B k) = f i * quadBiLin (B i) (B i) := by
|
||||
rw [quadBiLin.map_sum₂]
|
||||
rw [Fintype.sum_eq_single i]
|
||||
rw [quadBiLin.map_smul₂]
|
||||
|
@ -96,7 +96,7 @@ lemma quadCoeff_eq_bilinear (i : Fin 11) : quadCoeff i = quadBiLin (B i) (B i) :
|
|||
all_goals rfl
|
||||
|
||||
lemma on_accQuad (f : Fin 11 → ℚ) :
|
||||
accQuad (∑ i, f i • B i) = ∑ i, quadCoeff i * (f i)^2 := by
|
||||
accQuad (∑ i, f i • B i) = ∑ i, quadCoeff i * (f i)^2 := by
|
||||
change quadBiLin _ _ = _
|
||||
rw [quadBiLin.map_sum₁]
|
||||
apply Fintype.sum_congr
|
||||
|
@ -105,7 +105,7 @@ lemma on_accQuad (f : Fin 11 → ℚ) :
|
|||
ring
|
||||
|
||||
lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i))
|
||||
(k : Fin 11) : quadCoeff k * (f k)^2 = 0 := by
|
||||
(k : Fin 11) : quadCoeff k * (f k)^2 = 0 := by
|
||||
obtain ⟨S, hS⟩ := hS
|
||||
have hQ := quadSol S.1
|
||||
rw [hS, on_accQuad] at hQ
|
||||
|
@ -169,7 +169,7 @@ lemma isSolution_sum_part (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑
|
|||
simp
|
||||
rfl
|
||||
|
||||
lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
||||
lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
||||
f 10 = - 3 * f 9 := by
|
||||
have hx := isSolution_sum_part f hS
|
||||
obtain ⟨S, hS'⟩ := hS
|
||||
|
|
|
@ -61,7 +61,7 @@ lemma accQuad_α₁_α₂_zero (S : (PlusU1 n).LinSols) (h1 : α₁ C S = 0)
|
|||
/-- The construction of a `QuadSol` from a `LinSols` in the generic case. -/
|
||||
def genericToQuad (S : (PlusU1 n).LinSols) :
|
||||
(PlusU1 n).QuadSols :=
|
||||
linearToQuad ((α₁ C S) • S + α₂ S • C.1) (accQuad_α₁_α₂ C S)
|
||||
linearToQuad ((α₁ C S) • S + α₂ S • C.1) (accQuad_α₁_α₂ C S)
|
||||
|
||||
lemma genericToQuad_on_quad (S : (PlusU1 n).QuadSols) :
|
||||
genericToQuad C S.1 = (α₁ C S.1) • S := by
|
||||
|
@ -80,7 +80,7 @@ def specialToQuad (S : (PlusU1 n).LinSols) (a b : ℚ) (h1 : α₁ C S = 0)
|
|||
(h2 : α₂ S = 0) : (PlusU1 n).QuadSols :=
|
||||
linearToQuad (a • S + b • C.1) (accQuad_α₁_α₂_zero C S h1 h2 a b)
|
||||
|
||||
lemma special_on_quad (S : (PlusU1 n).QuadSols) (h1 : α₁ C S.1 = 0) :
|
||||
lemma special_on_quad (S : (PlusU1 n).QuadSols) (h1 : α₁ C S.1 = 0) :
|
||||
specialToQuad C S.1 1 0 h1 (α₂_AFQ S) = S := by
|
||||
apply ACCSystemQuad.QuadSols.ext
|
||||
change (1 • S.val + 0 • C.val) = S.val
|
||||
|
@ -123,9 +123,9 @@ lemma toQuadInv_special (S : (PlusU1 n).QuadSols) (h : α₁ C S.1 = 0) :
|
|||
rw [special_on_quad]
|
||||
|
||||
lemma toQuadInv_generic (S : (PlusU1 n).QuadSols) (h : α₁ C S.1 ≠ 0) :
|
||||
(toQuadInv C S).2.1 • genericToQuad C (toQuadInv C S).1 = S := by
|
||||
(toQuadInv C S).2.1 • genericToQuad C (toQuadInv C S).1 = S := by
|
||||
simp only [toQuadInv_fst]
|
||||
rw [show (toQuadInv C S).2.1 = (α₁ C S.1)⁻¹ by rw [toQuadInv, if_neg h]]
|
||||
rw [show (toQuadInv C S).2.1 = (α₁ C S.1)⁻¹ by rw [toQuadInv, if_neg h]]
|
||||
rw [genericToQuad_neq_zero C S h]
|
||||
|
||||
lemma toQuad_rightInverse : Function.RightInverse (@toQuadInv n C) (toQuad C) := by
|
||||
|
|
|
@ -71,7 +71,7 @@ def special (S : (PlusU1 n).QuadSols) (a b : ℚ) (h1 : α₁ S = 0) (h2 : α₂
|
|||
(PlusU1 n).Sols :=
|
||||
quadToAF (BL.addQuad S a b) (cube_α₁_α₂_zero S a b h1 h2)
|
||||
|
||||
lemma special_on_AF (S : (PlusU1 n).Sols) (h1 : α₁ S.1 = 0) :
|
||||
lemma special_on_AF (S : (PlusU1 n).Sols) (h1 : α₁ S.1 = 0) :
|
||||
special S.1 1 0 h1 (α₂_AF S) = S := by
|
||||
apply ACCSystem.Sols.ext
|
||||
change (BL.addQuad S.1 1 0).val = _
|
||||
|
@ -84,7 +84,7 @@ open QuadSolToSol
|
|||
/-- A map from `QuadSols × ℚ × ℚ` to `Sols` taking account of the special and generic cases.
|
||||
We will show that this map is a surjection. -/
|
||||
def quadSolToSol {n : ℕ} : (PlusU1 n).QuadSols × ℚ × ℚ → (PlusU1 n).Sols := fun S =>
|
||||
if h1 : α₁ S.1 = 0 ∧ α₂ S.1 = 0 then
|
||||
if h1 : α₁ S.1 = 0 ∧ α₂ S.1 = 0 then
|
||||
special S.1 S.2.1 S.2.2 h1.1 h1.2
|
||||
else
|
||||
S.2.1 • generic S.1
|
||||
|
@ -98,7 +98,7 @@ def quadSolToSolInv {n : ℕ} : (PlusU1 n).Sols → (PlusU1 n).QuadSols × ℚ
|
|||
else
|
||||
(S.1, (α₁ S.1)⁻¹, 0)
|
||||
|
||||
lemma quadSolToSolInv_1 (S : (PlusU1 n).Sols) :
|
||||
lemma quadSolToSolInv_1 (S : (PlusU1 n).Sols) :
|
||||
(quadSolToSolInv S).1 = S.1 := by
|
||||
simp [quadSolToSolInv]
|
||||
split
|
||||
|
@ -110,24 +110,24 @@ lemma quadSolToSolInv_α₁_α₂_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 = 0)
|
|||
rw [quadSolToSolInv_1, α₂_AF S, h]
|
||||
simp
|
||||
|
||||
lemma quadSolToSolInv_α₁_α₂_neq_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) :
|
||||
¬ (α₁ (quadSolToSolInv S).1 = 0 ∧ α₂ (quadSolToSolInv S).1 = 0) := by
|
||||
lemma quadSolToSolInv_α₁_α₂_neq_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) :
|
||||
¬ (α₁ (quadSolToSolInv S).1 = 0 ∧ α₂ (quadSolToSolInv S).1 = 0) := by
|
||||
rw [not_and, quadSolToSolInv_1, α₂_AF S]
|
||||
intro hn
|
||||
simp_all
|
||||
|
||||
lemma quadSolToSolInv_special (S : (PlusU1 n).Sols) (h : α₁ S.1 = 0) :
|
||||
special (quadSolToSolInv S).1 (quadSolToSolInv S).2.1 (quadSolToSolInv S).2.2
|
||||
(quadSolToSolInv_α₁_α₂_zero S h).1 (quadSolToSolInv_α₁_α₂_zero S h).2 = S := by
|
||||
(quadSolToSolInv_α₁_α₂_zero S h).1 (quadSolToSolInv_α₁_α₂_zero S h).2 = S := by
|
||||
simp [quadSolToSolInv_1]
|
||||
rw [show (quadSolToSolInv S).2.1 = 1 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [show (quadSolToSolInv S).2.2 = 0 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [show (quadSolToSolInv S).2.1 = 1 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [show (quadSolToSolInv S).2.2 = 0 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [special_on_AF]
|
||||
|
||||
lemma quadSolToSolInv_generic (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) :
|
||||
(quadSolToSolInv S).2.1 • generic (quadSolToSolInv S).1 = S := by
|
||||
(quadSolToSolInv S).2.1 • generic (quadSolToSolInv S).1 = S := by
|
||||
simp [quadSolToSolInv_1]
|
||||
rw [show (quadSolToSolInv S).2.1 = (α₁ S.1)⁻¹ by rw [quadSolToSolInv, if_neg h]]
|
||||
rw [show (quadSolToSolInv S).2.1 = (α₁ S.1)⁻¹ by rw [quadSolToSolInv, if_neg h]]
|
||||
rw [generic_on_AF_α₁_ne_zero S h]
|
||||
|
||||
lemma quadSolToSolInv_rightInverse : Function.RightInverse (@quadSolToSolInv n) quadSolToSol := by
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue