refactor: Change case of type and props

This commit is contained in:
jstoobysmith 2024-06-26 11:54:02 -04:00
parent 18b83f582e
commit f7a638d32e
58 changed files with 695 additions and 696 deletions

View file

@ -34,17 +34,17 @@ variable {n : }
/-- An equivalence between `(SMνCharges n).charges` and `(Fin 6 → Fin n → )`
splitting the charges into species.-/
@[simps!]
def toSpeciesEquiv : (SMνCharges n).charges ≃ (Fin 6 → Fin n → ) :=
def toSpeciesEquiv : (SMνCharges n).Charges ≃ (Fin 6 → Fin n → ) :=
((Equiv.curry _ _ _).symm.trans ((@finProdFinEquiv 6 n).arrowCongr (Equiv.refl ))).symm
/-- Given an `i ∈ Fin 6`, the projection of charges onto a given species. -/
@[simps!]
def toSpecies (i : Fin 6) : (SMνCharges n).charges →ₗ[] (SMνSpecies n).charges where
def toSpecies (i : Fin 6) : (SMνCharges n).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S := toSpeciesEquiv S i
map_add' _ _ := by aesop
map_smul' _ _ := by aesop
lemma charges_eq_toSpecies_eq (S T : (SMνCharges n).charges) :
lemma charges_eq_toSpecies_eq (S T : (SMνCharges n).Charges) :
S = T ↔ ∀ i, toSpecies i S = toSpecies i T := by
apply Iff.intro
intro h
@ -60,7 +60,7 @@ lemma toSMSpecies_toSpecies_inv (i : Fin 6) (f : (Fin 6 → Fin n → ) ) :
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
@ -94,7 +94,7 @@ variable {n : }
/-- The gravitational anomaly equation. -/
@[simp]
def accGrav : (SMνCharges n).charges →ₗ[] where
def accGrav : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (6 * Q S i + 3 * U S i + 3 * D S i + 2 * L S i + E S i + N S i)
map_add' S T := by
simp only
@ -112,7 +112,7 @@ def accGrav : (SMνCharges n).charges →ₗ[] where
ring
lemma accGrav_decomp (S : (SMνCharges n).charges) :
lemma accGrav_decomp (S : (SMνCharges n).Charges) :
accGrav S = 6 * ∑ i, Q S i + 3 * ∑ i, U S i + 3 * ∑ i, D S i + 2 * ∑ i, L S i + ∑ i, E S i +
∑ i, N S i := by
simp only [accGrav, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
@ -121,7 +121,7 @@ lemma accGrav_decomp (S : (SMνCharges n).charges) :
repeat erw [← Finset.mul_sum]
/-- Extensionality lemma for `accGrav`. -/
lemma accGrav_ext {S T : (SMνCharges n).charges}
lemma accGrav_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accGrav S = accGrav T := by
rw [accGrav_decomp, accGrav_decomp]
@ -130,7 +130,7 @@ lemma accGrav_ext {S T : (SMνCharges n).charges}
/-- The `SU(2)` anomaly equation. -/
@[simp]
def accSU2 : (SMνCharges n).charges →ₗ[] where
def accSU2 : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (3 * Q S i + L S i)
map_add' S T := by
simp only
@ -147,7 +147,7 @@ def accSU2 : (SMνCharges n).charges →ₗ[] where
-- rw [show Rat.cast a = a from rfl]
ring
lemma accSU2_decomp (S : (SMνCharges n).charges) :
lemma accSU2_decomp (S : (SMνCharges n).Charges) :
accSU2 S = 3 * ∑ i, Q S i + ∑ i, L S i := by
simp only [accSU2, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
AddHom.coe_mk]
@ -155,7 +155,7 @@ lemma accSU2_decomp (S : (SMνCharges n).charges) :
repeat erw [← Finset.mul_sum]
/-- Extensionality lemma for `accSU2`. -/
lemma accSU2_ext {S T : (SMνCharges n).charges}
lemma accSU2_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accSU2 S = accSU2 T := by
rw [accSU2_decomp, accSU2_decomp]
@ -163,7 +163,7 @@ lemma accSU2_ext {S T : (SMνCharges n).charges}
/-- The `SU(3)` anomaly equations. -/
@[simp]
def accSU3 : (SMνCharges n).charges →ₗ[] where
def accSU3 : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (2 * Q S i + U S i + D S i)
map_add' S T := by
simp only
@ -180,7 +180,7 @@ def accSU3 : (SMνCharges n).charges →ₗ[] where
-- rw [show Rat.cast a = a from rfl]
ring
lemma accSU3_decomp (S : (SMνCharges n).charges) :
lemma accSU3_decomp (S : (SMνCharges n).Charges) :
accSU3 S = 2 * ∑ i, Q S i + ∑ i, U S i + ∑ i, D S i := by
simp only [accSU3, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
AddHom.coe_mk]
@ -188,7 +188,7 @@ lemma accSU3_decomp (S : (SMνCharges n).charges) :
repeat rw [← Finset.mul_sum]
/-- Extensionality lemma for `accSU3`. -/
lemma accSU3_ext {S T : (SMνCharges n).charges}
lemma accSU3_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accSU3 S = accSU3 T := by
rw [accSU3_decomp, accSU3_decomp]
@ -196,7 +196,7 @@ lemma accSU3_ext {S T : (SMνCharges n).charges}
/-- The `Y²` anomaly equation. -/
@[simp]
def accYY : (SMνCharges n).charges →ₗ[] where
def accYY : (SMνCharges n).Charges →ₗ[] where
toFun S := ∑ i, (Q S i + 8 * U S i + 2 * D S i + 3 * L S i
+ 6 * E S i)
map_add' S T := by
@ -214,7 +214,7 @@ def accYY : (SMνCharges n).charges →ₗ[] where
-- rw [show Rat.cast a = a from rfl]
ring
lemma accYY_decomp (S : (SMνCharges n).charges) :
lemma accYY_decomp (S : (SMνCharges n).Charges) :
accYY S = ∑ i, Q S i + 8 * ∑ i, U S i + 2 * ∑ i, D S i + 3 * ∑ i, L S i + 6 * ∑ i, E S i := by
simp only [accYY, SMνSpecies_numberCharges, toSpecies_apply, Fin.isValue, LinearMap.coe_mk,
AddHom.coe_mk]
@ -222,7 +222,7 @@ lemma accYY_decomp (S : (SMνCharges n).charges) :
repeat rw [← Finset.mul_sum]
/-- Extensionality lemma for `accYY`. -/
lemma accYY_ext {S T : (SMνCharges n).charges}
lemma accYY_ext {S T : (SMνCharges n).Charges}
(hj : ∀ (j : Fin 6), ∑ i, (toSpecies j) S i = ∑ i, (toSpecies j) T i) :
accYY S = accYY T := by
rw [accYY_decomp, accYY_decomp]
@ -230,7 +230,7 @@ lemma accYY_ext {S T : (SMνCharges n).charges}
/-- The quadratic bilinear map. -/
@[simps!]
def quadBiLin : BiLinearSymm (SMνCharges n).charges := BiLinearSymm.mk₂
def quadBiLin : BiLinearSymm (SMνCharges n).Charges := BiLinearSymm.mk₂
(fun S => ∑ i, (Q S.1 i * Q S.2 i +
- 2 * (U S.1 i * U S.2 i) +
D S.1 i * D S.2 i +
@ -262,7 +262,7 @@ def quadBiLin : BiLinearSymm (SMνCharges n).charges := BiLinearSymm.mk₂
intro i
ring)
lemma quadBiLin_decomp (S T : (SMνCharges n).charges) :
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
erw [← quadBiLin.toFun_eq_coe]
@ -275,17 +275,17 @@ 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) :
lemma accQuad_decomp (S : (SMνCharges n).Charges) :
accQuad S = ∑ i, (Q S i)^2 - 2 * ∑ i, (U S i)^2 + ∑ i, (D S i)^2 - ∑ i, (L S i)^2
+ ∑ i, (E S i)^2 := by
erw [quadBiLin_decomp]
ring_nf
/-- Extensionality lemma for `accQuad`. -/
lemma accQuad_ext {S T : (SMνCharges n).charges}
lemma accQuad_ext {S T : (SMνCharges n).Charges}
(h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSpecies j S) i =
∑ i, ((fun a => a^2) ∘ toSpecies j T) i) :
accQuad S = accQuad T := by
@ -295,7 +295,7 @@ lemma accQuad_ext {S T : (SMνCharges n).charges}
/-- The symmetric trilinear form used to define the cubic acc. -/
@[simps!]
def cubeTriLin : TriLinearSymm (SMνCharges n).charges := TriLinearSymm.mk₃
def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃
(fun S => ∑ i, (6 * ((Q S.1 i) * (Q S.2.1 i) * (Q S.2.2 i))
+ 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))
@ -333,7 +333,7 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).charges := TriLinearSymm.mk₃
intro i
ring)
lemma cubeTriLin_decomp (S T R : (SMνCharges n).charges) :
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) +
∑ i, (E S i * E T i * E R i) + ∑ i, (N S i * N T i * N R i) := by
@ -347,16 +347,16 @@ lemma cubeTriLin_decomp (S T R : (SMνCharges n).charges) :
/-- The cubic ACC. -/
@[simp]
def accCube : HomogeneousCubic (SMνCharges n).charges := cubeTriLin.toCubic
def accCube : HomogeneousCubic (SMνCharges n).Charges := cubeTriLin.toCubic
lemma accCube_decomp (S : (SMνCharges n).charges) :
lemma accCube_decomp (S : (SMνCharges n).Charges) :
accCube S = 6 * ∑ i, (Q S i)^3 + 3 * ∑ i, (U S i)^3 + 3 * ∑ i, (D S i)^3 + 2 * ∑ i, (L S i)^3 +
∑ i, (E S i)^3 + ∑ i, (N S i)^3 := by
erw [cubeTriLin_decomp]
ring_nf
/-- Extensionality lemma for `accCube`. -/
lemma accCube_ext {S T : (SMνCharges n).charges}
lemma accCube_ext {S T : (SMνCharges n).Charges}
(h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSpecies j S) i =
∑ i, ((fun a => a^3) ∘ toSpecies j T) i) :
accCube S = accCube T := by

View file

@ -19,8 +19,8 @@ open BigOperators
/-- Given a map of for a generic species, the corresponding map for charges. -/
@[simps!]
def chargesMapOfSpeciesMap {n m : } (f : (SMνSpecies n).charges →ₗ[] (SMνSpecies m).charges) :
(SMνCharges n).charges →ₗ[] (SMνCharges m).charges where
def chargesMapOfSpeciesMap {n m : } (f : (SMνSpecies n).Charges →ₗ[] (SMνSpecies m).Charges) :
(SMνCharges n).Charges →ₗ[] (SMνCharges m).Charges where
toFun S := toSpeciesEquiv.symm (fun i => (LinearMap.comp f (toSpecies i)) S)
map_add' S T := by
rw [charges_eq_toSpecies_eq]
@ -37,28 +37,28 @@ def chargesMapOfSpeciesMap {n m : } (f : (SMνSpecies n).charges →ₗ[]
rfl
lemma chargesMapOfSpeciesMap_toSpecies {n m : }
(f : (SMνSpecies n).charges →ₗ[] (SMνSpecies m).charges)
(S : (SMνCharges n).charges) (j : Fin 6) :
(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
erw [toSMSpecies_toSpecies_inv]
/-- 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
(SMνSpecies m).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S := S ∘ Fin.castLE h
map_add' _ _ := rfl
map_smul' _ _ := rfl
/-- 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 :=
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. -/
@[simps!]
def speciesEmbed (m n : ) :
(SMνSpecies m).charges →ₗ[] (SMνSpecies n).charges where
(SMνSpecies m).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S := fun i =>
if hi : i.val < m then
S ⟨i, hi⟩
@ -82,29 +82,29 @@ def speciesEmbed (m n : ) :
/-- The embedding of the `m`-family charges onto the `n`-family charges, with all
other charges zero. -/
def familyEmbedding (m n : ) : (SMνCharges m).charges →ₗ[] (SMνCharges n).charges :=
def familyEmbedding (m n : ) : (SMνCharges m).Charges →ₗ[] (SMνCharges n).Charges :=
chargesMapOfSpeciesMap (speciesEmbed m n)
/-- For species, the embedding of the `1`-family charges into the `n`-family charges in
a universal manor. -/
@[simps!]
def speciesFamilyUniversial (n : ) :
(SMνSpecies 1).charges →ₗ[] (SMνSpecies n).charges where
(SMνSpecies 1).Charges →ₗ[] (SMνSpecies n).Charges where
toFun S _ := S ⟨0, by simp⟩
map_add' S T := rfl
map_smul' a S := rfl
/-- The embedding of the `1`-family charges into the `n`-family charges in
a universal manor. -/
def familyUniversal (n : ) : (SMνCharges 1).charges →ₗ[] (SMνCharges n).charges :=
def familyUniversal (n : ) : (SMνCharges 1).Charges →ₗ[] (SMνCharges n).Charges :=
chargesMapOfSpeciesMap (speciesFamilyUniversial n)
lemma toSpecies_familyUniversal {n : } (j : Fin 6) (S : (SMνCharges 1).charges)
lemma toSpecies_familyUniversal {n : } (j : Fin 6) (S : (SMνCharges 1).Charges)
(i : Fin n) : toSpecies j (familyUniversal n S) i = toSpecies j S ⟨0, by simp⟩ := by
erw [chargesMapOfSpeciesMap_toSpecies]
rfl
lemma sum_familyUniversal {n : } (m : ) (S : (SMνCharges 1).charges) (j : Fin 6) :
lemma sum_familyUniversal {n : } (m : ) (S : (SMνCharges 1).Charges) (j : Fin 6) :
∑ i, ((fun a => a ^ m) ∘ toSpecies j (familyUniversal n S)) i =
n * (toSpecies j S ⟨0, by simp⟩) ^ m := by
simp only [SMνSpecies_numberCharges, Function.comp_apply, toSpecies_apply, Fin.zero_eta,
@ -119,12 +119,12 @@ 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
lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).charges)
(T : (SMνCharges n).charges) (j : Fin 6) :
lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).Charges)
(T : (SMνCharges n).Charges) (j : Fin 6) :
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i) =
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i := by
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
@ -135,8 +135,8 @@ lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).charges)
erw [toSpecies_familyUniversal]
rfl
lemma sum_familyUniversal_three {n : } (S : (SMνCharges 1).charges)
(T L : (SMνCharges n).charges) (j : Fin 6) :
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
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
@ -148,7 +148,7 @@ lemma sum_familyUniversal_three {n : } (S : (SMνCharges 1).charges)
simp only [SMνSpecies_numberCharges, Fin.zero_eta, Fin.isValue, toSpecies_apply]
ring
lemma familyUniversal_accGrav (S : (SMνCharges 1).charges) :
lemma familyUniversal_accGrav (S : (SMνCharges 1).Charges) :
accGrav (familyUniversal n S) = n * (accGrav S) := by
rw [accGrav_decomp, accGrav_decomp]
repeat rw [sum_familyUniversal_one]
@ -156,7 +156,7 @@ lemma familyUniversal_accGrav (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_accSU2 (S : (SMνCharges 1).charges) :
lemma familyUniversal_accSU2 (S : (SMνCharges 1).Charges) :
accSU2 (familyUniversal n S) = n * (accSU2 S) := by
rw [accSU2_decomp, accSU2_decomp]
repeat rw [sum_familyUniversal_one]
@ -164,7 +164,7 @@ lemma familyUniversal_accSU2 (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_accSU3 (S : (SMνCharges 1).charges) :
lemma familyUniversal_accSU3 (S : (SMνCharges 1).Charges) :
accSU3 (familyUniversal n S) = n * (accSU3 S) := by
rw [accSU3_decomp, accSU3_decomp]
repeat rw [sum_familyUniversal_one]
@ -172,7 +172,7 @@ lemma familyUniversal_accSU3 (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_accYY (S : (SMνCharges 1).charges) :
lemma familyUniversal_accYY (S : (SMνCharges 1).Charges) :
accYY (familyUniversal n S) = n * (accYY S) := by
rw [accYY_decomp, accYY_decomp]
repeat rw [sum_familyUniversal_one]
@ -180,7 +180,7 @@ lemma familyUniversal_accYY (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_quadBiLin (S : (SMνCharges 1).charges) (T : (SMνCharges n).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
@ -191,7 +191,7 @@ lemma familyUniversal_quadBiLin (S : (SMνCharges 1).charges) (T : (SMνCharges
sub_right_inj]
ring
lemma familyUniversal_accQuad (S : (SMνCharges 1).charges) :
lemma familyUniversal_accQuad (S : (SMνCharges 1).Charges) :
accQuad (familyUniversal n S) = n * (accQuad S) := by
rw [accQuad_decomp]
repeat erw [sum_familyUniversal]
@ -200,7 +200,7 @@ lemma familyUniversal_accQuad (S : (SMνCharges 1).charges) :
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
ring
lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).charges) (T R : (SMνCharges n).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)
+ 2 * S (3 : Fin 6) * ∑ i, (L T i * L R i) +
@ -211,7 +211,7 @@ lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).charges) (T R : (SMνCharg
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply, add_left_inj]
ring
lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).charges) (R : (SMνCharges n).charges) :
lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).Charges) (R : (SMνCharges n).Charges) :
cubeTriLin (familyUniversal n S) (familyUniversal n T) R =
6 * S (0 : Fin 6) * T (0 : Fin 6) * ∑ i, Q R i +
3 * S (1 : Fin 6) * T (1 : Fin 6) * ∑ i, U R i
@ -224,7 +224,7 @@ lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).charges) (R : (SMνChar
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply]
ring
lemma familyUniversal_accCube (S : (SMνCharges 1).charges) :
lemma familyUniversal_accCube (S : (SMνCharges 1).Charges) :
accCube (familyUniversal n S) = n * (accCube S) := by
rw [accCube_decomp]
repeat erw [sum_familyUniversal]

View file

@ -53,7 +53,7 @@ lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs
gives us a element of `LinSols`. -/
def chargeToLinear (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
def chargeToLinear (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SMNoGrav n).LinSols :=
⟨S, by
intro i
@ -76,13 +76,13 @@ def quadToAF (S : (SMNoGrav n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `QuadSols`. -/
def chargeToQuad (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
def chargeToQuad (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SMNoGrav n).QuadSols :=
linearToQuad $ chargeToLinear S hSU2 hSU3
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `Sols`. -/
def chargeToAF (S : (SMNoGrav n).charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0)
def chargeToAF (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0)
(hc : accCube S = 0) : (SMNoGrav n).Sols :=
quadToAF (chargeToQuad S hSU2 hSU3) hc
@ -95,7 +95,7 @@ def linearToAF (S : (SMNoGrav n).LinSols)
/-- The permutations acting on the ACC system corresponding to the SM with RHN,
and no gravitational anomaly. -/
def perm (n : ) : ACCSystemGroupAction (SMNoGrav n) where
group := permGroup n
group := PermGroup n
groupInst := inferInstance
rep := repCharges
linearInvariant := by

View file

@ -60,7 +60,7 @@ lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs
gives us a element of `LinSols`. -/
def chargeToLinear (S : (SM n).charges) (hGrav : accGrav S = 0)
def chargeToLinear (S : (SM n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (SM n).LinSols :=
⟨S, by
intro i
@ -84,14 +84,14 @@ def quadToAF (S : (SM n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `QuadSols`. -/
def chargeToQuad (S : (SM n).charges) (hGrav : accGrav S = 0)
def chargeToQuad (S : (SM n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) :
(SM n).QuadSols :=
linearToQuad $ chargeToLinear S hGrav hSU2 hSU3
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `Sols`. -/
def chargeToAF (S : (SM n).charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0)
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
@ -103,7 +103,7 @@ def linearToAF (S : (SM n).LinSols)
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
def perm (n : ) : ACCSystemGroupAction (SM n) where
group := permGroup n
group := PermGroup n
groupInst := inferInstance
rep := repCharges
linearInvariant := by

View file

@ -24,102 +24,102 @@ open BigOperators
namespace PlaneSeven
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₀ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i =>
def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 0, 0 => 1
| 0, 1 => - 1
| _, _ => 0
)
lemma B₀_cubic (S T : (SM 3).charges) : cubeTriLin B₀ S T =
lemma B₀_cubic (S T : (SM 3).Charges) : cubeTriLin B₀ S T =
6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by
simp [Fin.sum_univ_three, B₀, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₁ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i =>
def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 1, 0 => 1
| 1, 1 => - 1
| _, _ => 0
)
lemma B₁_cubic (S T : (SM 3).charges) : cubeTriLin B₁ S T =
lemma B₁_cubic (S T : (SM 3).Charges) : cubeTriLin B₁ S T =
3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by
simp [Fin.sum_univ_three, B₁, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₂ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i =>
def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 2, 0 => 1
| 2, 1 => - 1
| _, _ => 0
)
lemma B₂_cubic (S T : (SM 3).charges) : cubeTriLin B₂ S T =
lemma B₂_cubic (S T : (SM 3).Charges) : cubeTriLin B₂ S T =
3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by
simp [Fin.sum_univ_three, B₂, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₃ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i =>
def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 3, 0 => 1
| 3, 1 => - 1
| _, _ => 0
)
lemma B₃_cubic (S T : (SM 3).charges) : cubeTriLin B₃ S T =
lemma B₃_cubic (S T : (SM 3).Charges) : cubeTriLin B₃ S T =
2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by
simp [Fin.sum_univ_three, B₃, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf
rfl
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₄ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i =>
def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 4, 0 => 1
| 4, 1 => - 1
| _, _ => 0
)
lemma B₄_cubic (S T : (SM 3).charges) : cubeTriLin B₄ S T =
lemma B₄_cubic (S T : (SM 3).Charges) : cubeTriLin B₄ S T =
(S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by
simp [Fin.sum_univ_three, B₄, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf
rfl
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₅ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i =>
def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 5, 0 => 1
| 5, 1 => - 1
| _, _ => 0
)
lemma B₅_cubic (S T : (SM 3).charges) : cubeTriLin B₅ S T =
lemma B₅_cubic (S T : (SM 3).Charges) : cubeTriLin B₅ S T =
(S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by
simp [Fin.sum_univ_three, B₅, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf
rfl
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₆ : (SM 3).charges := toSpeciesEquiv.invFun ( fun s => fun i =>
def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun ( fun s => fun i =>
match s, i with
| 1, 2 => 1
| 2, 2 => -1
| _, _ => 0
)
lemma B₆_cubic (S T : (SM 3).charges) : cubeTriLin B₆ S T =
lemma B₆_cubic (S T : (SM 3).Charges) : cubeTriLin B₆ S T =
3* (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by
simp [Fin.sum_univ_three, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
ring_nf
/-- The charge assignments forming a basis of the plane. -/
@[simp]
def B : Fin 7 → (SM 3).charges := fun i =>
def B : Fin 7 → (SM 3).Charges := fun i =>
match i with
| 0 => B₀
| 1 => B₁
@ -129,7 +129,7 @@ def B : Fin 7 → (SM 3).charges := fun i =>
| 5 => B₅
| 6 => B₆
lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).charges) :
lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 0) (B i) S = 0 := by
change cubeTriLin B₀ (B i) S = 0
rw [B₀_cubic]
@ -138,7 +138,7 @@ lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).charges) :
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).charges) :
lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 1) (B i) S = 0 := by
change cubeTriLin B₁ (B i) S = 0
rw [B₁_cubic]
@ -146,7 +146,7 @@ lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).charges) :
simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).charges) :
lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 2) (B i) S = 0 := by
change cubeTriLin B₂ (B i) S = 0
rw [B₂_cubic]
@ -154,7 +154,7 @@ lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).charges) :
simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).charges) :
lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 3) (B i) S = 0 := by
change cubeTriLin (B₃) (B i) S = 0
rw [B₃_cubic]
@ -162,7 +162,7 @@ lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).charges) :
simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).charges) :
lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 4) (B i) S = 0 := by
change cubeTriLin (B₄) (B i) S = 0
rw [B₄_cubic]
@ -170,7 +170,7 @@ lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).charges) :
simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).charges) :
lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 5) (B i) S = 0 := by
change cubeTriLin (B₅) (B i) S = 0
rw [B₅_cubic]
@ -178,7 +178,7 @@ lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).charges) :
simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).charges) :
lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).Charges) :
cubeTriLin (B 6) (B i) S = 0 := by
change cubeTriLin (B₆) (B i) S = 0
rw [B₆_cubic]
@ -186,7 +186,7 @@ lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).charges) :
simp at hi <;>
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
lemma Bi_Bj_ne_cubic {i j : Fin 7} (h : i ≠ j) (S : (SM 3).charges) :
lemma Bi_Bj_ne_cubic {i j : Fin 7} (h : i ≠ j) (S : (SM 3).Charges) :
cubeTriLin (B i) (B j) S = 0 := by
fin_cases i
exact B₀_Bi_cubic h S
@ -279,7 +279,7 @@ theorem B_in_accCube (f : Fin 7 → ) : accCube (∑ i, f i • B i) = 0 := b
rw [Bi_Bj_Bk_cubic]
simp
lemma B_sum_is_sol (f : Fin 7 → ) : (SM 3).isSolution (∑ i, f i • B i) := by
lemma B_sum_is_sol (f : Fin 7 → ) : (SM 3).IsSolution (∑ i, f i • B i) := by
let X := chargeToAF (∑ i, f i • B i) (by
rw [map_sum]
apply Fintype.sum_eq_zero
@ -337,8 +337,8 @@ 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
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
use PlaneSeven.B
apply And.intro
exact PlaneSeven.basis_linear_independent

View file

@ -26,26 +26,26 @@ open BigOperators
/-- The group of `Sₙ` permutations for each species. -/
@[simp]
def permGroup (n : ) := Fin 6 → Equiv.Perm (Fin n)
def PermGroup (n : ) := Fin 6 → Equiv.Perm (Fin n)
variable {n : }
@[simp]
instance : Group (permGroup n) := Pi.group
instance : Group (PermGroup n) := Pi.group
/-- The image of an element of `permGroup n` under the representation on charges. -/
@[simps!]
def chargeMap (f : permGroup n) : (SMνCharges n).charges →ₗ[] (SMνCharges n).charges where
def chargeMap (f : PermGroup n) : (SMνCharges n).Charges →ₗ[] (SMνCharges n).Charges where
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i )
map_add' _ _ := rfl
map_smul' _ _ := rfl
/-- The representation of `(permGroup n)` acting on the vector space of charges. -/
@[simp]
def repCharges {n : } : Representation (permGroup n) (SMνCharges n).charges where
def repCharges {n : } : Representation (PermGroup n) (SMνCharges n).Charges where
toFun f := chargeMap f⁻¹
map_mul' f g := by
simp only [permGroup, mul_inv_rev]
simp only [PermGroup, mul_inv_rev]
apply LinearMap.ext
intro S
rw [charges_eq_toSpecies_eq]
@ -61,49 +61,49 @@ def repCharges {n : } : Representation (permGroup n) (SMνCharges n).char
erw [toSMSpecies_toSpecies_inv]
rfl
lemma repCharges_toSpecies (f : permGroup n) (S : (SMνCharges n).charges) (j : Fin 6) :
lemma repCharges_toSpecies (f : PermGroup n) (S : (SMνCharges n).Charges) (j : Fin 6) :
toSpecies j (repCharges f S) = toSpecies j S ∘ f⁻¹ j := by
erw [toSMSpecies_toSpecies_inv]
lemma toSpecies_sum_invariant (m : ) (f : permGroup n) (S : (SMνCharges n).charges) (j : Fin 6) :
lemma toSpecies_sum_invariant (m : ) (f : PermGroup n) (S : (SMνCharges n).Charges) (j : Fin 6) :
∑ 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
refine Equiv.Perm.sum_comp _ _ _ ?_
simp only [permGroup, Fin.isValue, Pi.inv_apply, ne_eq, coe_univ, Set.subset_univ]
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)

View file

@ -57,7 +57,7 @@ namespace BL
variable {n : }
lemma on_quadBiLin (S : (PlusU1 n).charges) :
lemma on_quadBiLin (S : (PlusU1 n).Charges) :
quadBiLin (BL n).val S = 1/2 * accYY S + 3/2 * accSU2 S - 2 * accSU3 S := by
erw [familyUniversal_quadBiLin]
rw [accYY_decomp, accSU2_decomp, accSU3_decomp]
@ -90,7 +90,7 @@ lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ): addQuad S a 0 = a • S
simp [addQuad, linearToQuad]
rfl
lemma on_cubeTriLin (S : (PlusU1 n).charges) :
lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
cubeTriLin (BL n).val (BL n).val S = 9 * accGrav S - 24 * accSU3 S := by
erw [familyUniversal_cubeTriLin']
rw [accGrav_decomp, accSU3_decomp]

View file

@ -70,7 +70,7 @@ lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by
/-- An element of `charges` which satisfies the linear ACCs
gives us a element of `LinSols`. -/
def chargeToLinear (S : (PlusU1 n).charges) (hGrav : accGrav S = 0)
def chargeToLinear (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) (hYY : accYY S = 0) :
(PlusU1 n).LinSols :=
⟨S, by
@ -99,14 +99,14 @@ def quadToAF (S : (PlusU1 n).QuadSols) (hc : accCube S.val = 0) :
/-- An element of `charges` which satisfies the linear and quadratic ACCs
gives us a element of `QuadSols`. -/
def chargeToQuad (S : (PlusU1 n).charges) (hGrav : accGrav S = 0)
def chargeToQuad (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0)
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) (hYY : accYY S = 0) (hQ : accQuad S = 0) :
(PlusU1 n).QuadSols :=
linearToQuad (chargeToLinear S hGrav hSU2 hSU3 hYY) hQ
/-- An element of `charges` which satisfies the linear, quadratic and cubic ACCs
gives us a element of `Sols`. -/
def chargeToAF (S : (PlusU1 n).charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0)
def chargeToAF (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0) (hSU2 : accSU2 S = 0)
(hSU3 : accSU3 S = 0) (hYY : accYY S = 0) (hQ : accQuad S = 0) (hc : accCube S = 0) :
(PlusU1 n).Sols :=
quadToAF (chargeToQuad S hGrav hSU2 hSU3 hYY hQ) hc
@ -119,7 +119,7 @@ def linearToAF (S : (PlusU1 n).LinSols) (hQ : accQuad S.val = 0)
/-- The permutations acting on the ACC system corresponding to the SM with RHN. -/
def perm (n : ) : ACCSystemGroupAction (PlusU1 n) where
group := permGroup n
group := PermGroup n
groupInst := inferInstance
rep := repCharges
linearInvariant := by

View file

@ -23,11 +23,11 @@ open BigOperators
/-- A proposition which is true if for a given `n` a plane of charges of dimension `n` exists
in which each point is a solution. -/
def existsPlane (n : ) : Prop := ∃ (B : Fin n → (PlusU1 3).charges),
LinearIndependent B ∧ ∀ (f : Fin n → ), (PlusU1 3).isSolution (∑ i, f i • B i)
def ExistsPlane (n : ) : Prop := ∃ (B : Fin n → (PlusU1 3).Charges),
LinearIndependent B ∧ ∀ (f : Fin n → ), (PlusU1 3).IsSolution (∑ i, f i • B i)
lemma exists_plane_exists_basis {n : } (hE : existsPlane n) :
∃ (B : Fin 11 ⊕ Fin n → (PlusU1 3).charges), LinearIndependent B := by
lemma exists_plane_exists_basis {n : } (hE : ExistsPlane n) :
∃ (B : Fin 11 ⊕ Fin n → (PlusU1 3).Charges), LinearIndependent B := by
obtain ⟨E, hE1, hE2⟩ := hE
obtain ⟨B, hB1, hB2⟩ := eleven_dim_plane_of_no_sols_exists
let Y := Sum.elim B E
@ -59,11 +59,11 @@ lemma exists_plane_exists_basis {n : } (hE : existsPlane n) :
| Sum.inr i => exact h4 i
theorem plane_exists_dim_le_7 {n : } (hn : existsPlane n) : n ≤ 7 := by
theorem plane_exists_dim_le_7 {n : } (hn : ExistsPlane n) : n ≤ 7 := by
obtain ⟨B, hB⟩ := exists_plane_exists_basis hn
have h1 := LinearIndependent.fintype_card_le_finrank hB
simp at h1
rw [show FiniteDimensional.finrank (PlusU1 3).charges = 18 from
rw [show FiniteDimensional.finrank (PlusU1 3).Charges = 18 from
FiniteDimensional.finrank_fin_fun ] at h1
exact Nat.le_of_add_le_add_left h1

View file

@ -55,7 +55,7 @@ namespace Y
variable {n : }
lemma on_quadBiLin (S : (PlusU1 n).charges) :
lemma on_quadBiLin (S : (PlusU1 n).Charges) :
quadBiLin (Y n).val S = accYY S := by
erw [familyUniversal_quadBiLin]
rw [accYY_decomp]
@ -89,7 +89,7 @@ lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ): addQuad S a 0 = a • S
simp [addQuad, linearToQuad]
rfl
lemma on_cubeTriLin (S : (PlusU1 n).charges) :
lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
cubeTriLin (Y n).val (Y n).val S = 6 * accYY S := by
erw [familyUniversal_cubeTriLin']
rw [accYY_decomp]
@ -103,7 +103,7 @@ lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
rw [YYsol S]
simp
lemma on_cubeTriLin' (S : (PlusU1 n).charges) :
lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
cubeTriLin (Y n).val S S = 6 * accQuad S := by
erw [familyUniversal_cubeTriLin]
rw [accQuad_decomp]

View file

@ -26,40 +26,40 @@ open BigOperators
namespace ElevenPlane
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₀ : (PlusU1 3).charges := ![1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
def B₀ : (PlusU1 3).Charges := ![1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₁ : (PlusU1 3).charges := ![0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
def B₁ : (PlusU1 3).Charges := ![0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
/-- A charge assignment forming one of the basis elements of the plane. -/
def B₂ : (PlusU1 3).charges := ![0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
def B₂ : (PlusU1 3).Charges := ![0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 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, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
def B₃ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 0, 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, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0]
def B₄ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 0, 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, 1, 0, 0, 0, 0, 0, 0, 0, 0, 0]
def B₅ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 0, 0, 0, 0, 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, 1, 0, 0, 0, 0, 0]
def B₆ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 0, 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, 1, 0, 0, 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]
def B₉ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1, 2, 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, 0, 0, 1]
def B₁₀ : (PlusU1 3).Charges := ![0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 0, 1]
/-- The charge assignment forming a basis of the plane. -/
def B : Fin 11 → (PlusU1 3).charges := fun i =>
def B : Fin 11 → (PlusU1 3).Charges := fun i =>
match i with
| 0 => B₀
| 1 => B₁
@ -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))
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
obtain ⟨S, hS⟩ := hS
have hQ := quadSol S.1
@ -120,46 +120,46 @@ lemma isSolution_quadCoeff_f_sq_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSol
fin_cases i <;> rfl
exact sq_nonneg (f i)
lemma isSolution_f0 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 0 = 0 := by
lemma isSolution_f0 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 0 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 0)
lemma isSolution_f1 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 1 = 0 := by
lemma isSolution_f1 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 1 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 1)
lemma isSolution_f2 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 2 = 0 := by
lemma isSolution_f2 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 2 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 2)
lemma isSolution_f3 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 3 = 0 := by
lemma isSolution_f3 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 3 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 3)
lemma isSolution_f4 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 4 = 0 := by
lemma isSolution_f4 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 4 = 0 := by
simpa using (isSolution_quadCoeff_f_sq_zero f hS 4)
lemma isSolution_f5 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 5 = 0 := by
lemma isSolution_f5 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 5 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 5
rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h
simpa using h
lemma isSolution_f6 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 6 = 0 := by
lemma isSolution_f6 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 6 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 6
rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h
simpa using h
lemma isSolution_f7 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 7 = 0 := by
lemma isSolution_f7 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 7 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 7
rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h
simpa using h
lemma isSolution_f8 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) : f 8 = 0 := by
lemma isSolution_f8 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) : f 8 = 0 := by
have h := isSolution_quadCoeff_f_sq_zero f hS 8
rw [mul_eq_zero] at h
change 1 = 0 _ = _ at h
simpa using h
lemma isSolution_sum_part (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) :
lemma isSolution_sum_part (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
∑ i, f i • B i = f 9 • B₉ + f 10 • B₁₀ := by
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc, Fin.sum_univ_castSucc, Fin.sum_univ_eight]
change f 0 • B 0 + f 1 • B 1 + f 2 • B 2 + f 3 • B 3 + f 4 • B 4 + f 5 • B 5 + f 6 • B 6 +
@ -170,7 +170,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
@ -183,12 +183,12 @@ lemma isSolution_grav (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i,
simp at hg
linear_combination hg
lemma isSolution_sum_part' (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) :
lemma isSolution_sum_part' (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
∑ i, f i • B i = f 9 • B₉ + (- 3 * f 9) • B₁₀ := by
rw [isSolution_sum_part f hS]
rw [isSolution_grav f hS]
lemma isSolution_f9 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) :
lemma isSolution_f9 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
f 9 = 0 := by
have hx := isSolution_sum_part' f hS
obtain ⟨S, hS'⟩ := hS
@ -211,12 +211,12 @@ lemma isSolution_f9 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i
rw [h1] at hc
simpa using hc
lemma isSolution_f10 (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) :
lemma isSolution_f10 (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
f 10 = 0 := by
rw [isSolution_grav f hS, isSolution_f9 f hS]
simp
lemma isSolution_f_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i))
lemma isSolution_f_zero (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i))
(k : Fin 11) : f k = 0 := by
fin_cases k
exact isSolution_f0 f hS
@ -232,7 +232,7 @@ lemma isSolution_f_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i,
exact isSolution_f10 f hS
lemma isSolution_only_if_zero (f : Fin 11 → ) (hS : (PlusU1 3).isSolution (∑ i, f i • B i)) :
lemma isSolution_only_if_zero (f : Fin 11 → ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
∑ i, f i • B i = 0 := by
rw [isSolution_sum_part f hS]
rw [isSolution_grav f hS]
@ -244,7 +244,7 @@ theorem basis_linear_independent : LinearIndependent B := by
apply Fintype.linearIndependent_iff.mpr
intro f h
let X : (PlusU1 3).Sols := chargeToAF 0 (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl)
have hS : (PlusU1 3).isSolution (∑ i, f i • B i) := by
have hS : (PlusU1 3).IsSolution (∑ i, f i • B i) := by
use X
rw [h]
rfl
@ -252,9 +252,9 @@ theorem basis_linear_independent : LinearIndependent B := by
end ElevenPlane
theorem eleven_dim_plane_of_no_sols_exists : ∃ (B : Fin 11 → (PlusU1 3).charges),
theorem eleven_dim_plane_of_no_sols_exists : ∃ (B : Fin 11 → (PlusU1 3).Charges),
LinearIndependent B ∧
∀ (f : Fin 11 → ), (PlusU1 3).isSolution (∑ i, f i • B i) → ∑ i, f i • B i = 0 := by
∀ (f : Fin 11 → ), (PlusU1 3).IsSolution (∑ i, f i • B i) → ∑ i, f i • B i = 0 := by
use ElevenPlane.B
apply And.intro
exact ElevenPlane.basis_linear_independent