refactor: Change case of type and props
This commit is contained in:
parent
18b83f582e
commit
f7a638d32e
58 changed files with 695 additions and 696 deletions
|
@ -26,7 +26,7 @@ open BigOperators
|
|||
|
||||
/-- `B₃` is the charge which is $B-L$ in all families, but with the third
|
||||
family of the opposite sign. -/
|
||||
def B₃AsCharge : MSSMACC.charges := toSpecies.symm
|
||||
def B₃AsCharge : MSSMACC.Charges := toSpecies.symm
|
||||
⟨fun s => fun i =>
|
||||
match s, i with
|
||||
| 0, 0 => 1
|
||||
|
|
|
@ -37,7 +37,7 @@ namespace MSSMCharges
|
|||
`(Fin 18 ⊕ Fin 2 → ℚ)`. The first 18 factors corresponds to the SM fermions, while the last two
|
||||
are the higgsions. -/
|
||||
@[simps!]
|
||||
def toSMPlusH : MSSMCharges.charges ≃ (Fin 18 ⊕ Fin 2 → ℚ) :=
|
||||
def toSMPlusH : MSSMCharges.Charges ≃ (Fin 18 ⊕ Fin 2 → ℚ) :=
|
||||
((@finSumFinEquiv 18 2).arrowCongr (Equiv.refl ℚ)).symm
|
||||
|
||||
/-- An equivalence between `Fin 18 ⊕ Fin 2 → ℚ` and `(Fin 18 → ℚ) × (Fin 2 → ℚ)`. -/
|
||||
|
@ -51,7 +51,7 @@ def splitSMPlusH : (Fin 18 ⊕ Fin 2 → ℚ) ≃ (Fin 18 → ℚ) × (Fin 2 →
|
|||
/-- An equivalence between `MSSMCharges.charges` and `(Fin 18 → ℚ) × (Fin 2 → ℚ)`. This
|
||||
splits the charges up into the SM and the additional ones for the MSSM. -/
|
||||
@[simps!]
|
||||
def toSplitSMPlusH : MSSMCharges.charges ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) :=
|
||||
def toSplitSMPlusH : MSSMCharges.Charges ≃ (Fin 18 → ℚ) × (Fin 2 → ℚ) :=
|
||||
toSMPlusH.trans splitSMPlusH
|
||||
|
||||
/-- An equivalence between `(Fin 18 → ℚ)` and `(Fin 6 → Fin 3 → ℚ)`. -/
|
||||
|
@ -64,13 +64,13 @@ def toSpeciesMaps' : (Fin 18 → ℚ) ≃ (Fin 6 → Fin 3 → ℚ) :=
|
|||
This split charges up into the SM and additional fermions, and further splits the SM into
|
||||
species. -/
|
||||
@[simps!]
|
||||
def toSpecies : MSSMCharges.charges ≃ (Fin 6 → Fin 3 → ℚ) × (Fin 2 → ℚ) :=
|
||||
def toSpecies : MSSMCharges.Charges ≃ (Fin 6 → Fin 3 → ℚ) × (Fin 2 → ℚ) :=
|
||||
toSplitSMPlusH.trans (Equiv.prodCongr toSpeciesMaps' (Equiv.refl _))
|
||||
|
||||
/-- For a given `i ∈ Fin 6` the projection of `MSSMCharges.charges` down to the
|
||||
corresponding SM species of charges. -/
|
||||
@[simps!]
|
||||
def toSMSpecies (i : Fin 6) : MSSMCharges.charges →ₗ[ℚ] MSSMSpecies.charges where
|
||||
def toSMSpecies (i : Fin 6) : MSSMCharges.Charges →ₗ[ℚ] MSSMSpecies.Charges where
|
||||
toFun S := (Prod.fst ∘ toSpecies) S i
|
||||
map_add' _ _ := by rfl
|
||||
map_smul' _ _ := by rfl
|
||||
|
@ -95,19 +95,19 @@ abbrev N := toSMSpecies 5
|
|||
|
||||
/-- The charge `Hd`. -/
|
||||
@[simps!]
|
||||
def Hd : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
def Hd : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
||||
toFun S := S ⟨18, Nat.lt_of_sub_eq_succ rfl⟩
|
||||
map_add' _ _ := by rfl
|
||||
map_smul' _ _ := by rfl
|
||||
|
||||
/-- The charge `Hu`. -/
|
||||
@[simps!]
|
||||
def Hu : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
def Hu : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
||||
toFun S := S ⟨19, Nat.lt_of_sub_eq_succ rfl⟩
|
||||
map_add' _ _ := by rfl
|
||||
map_smul' _ _ := by rfl
|
||||
|
||||
lemma charges_eq_toSpecies_eq (S T : MSSMCharges.charges) :
|
||||
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
|
||||
apply Iff.intro
|
||||
intro h
|
||||
|
@ -139,7 +139,7 @@ open MSSMCharges
|
|||
|
||||
/-- The gravitational anomaly equation. -/
|
||||
@[simp]
|
||||
def accGrav : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
def accGrav : MSSMCharges.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) + 2 * (Hd S + Hu S)
|
||||
map_add' S T := by
|
||||
|
@ -159,7 +159,7 @@ def accGrav : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
|||
ring
|
||||
|
||||
/-- Extensionality lemma for `accGrav`. -/
|
||||
lemma accGrav_ext {S T : MSSMCharges.charges}
|
||||
lemma accGrav_ext {S T : MSSMCharges.Charges}
|
||||
(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
|
||||
|
@ -173,7 +173,7 @@ lemma accGrav_ext {S T : MSSMCharges.charges}
|
|||
|
||||
/-- The anomaly cancellation condition for SU(2) anomaly. -/
|
||||
@[simp]
|
||||
def accSU2 : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
def accSU2 : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
||||
toFun S := ∑ i, (3 * Q S i + L S i) + Hd S + Hu S
|
||||
map_add' S T := by
|
||||
simp only
|
||||
|
@ -192,7 +192,7 @@ def accSU2 : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
|||
ring
|
||||
|
||||
/-- Extensionality lemma for `accSU2`. -/
|
||||
lemma accSU2_ext {S T : MSSMCharges.charges}
|
||||
lemma accSU2_ext {S T : MSSMCharges.Charges}
|
||||
(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
|
||||
|
@ -206,7 +206,7 @@ lemma accSU2_ext {S T : MSSMCharges.charges}
|
|||
|
||||
/-- The anomaly cancellation condition for SU(3) anomaly. -/
|
||||
@[simp]
|
||||
def accSU3 : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
def accSU3 : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
||||
toFun S := ∑ i, (2 * (Q S i) + (U S i) + (D S i))
|
||||
map_add' S T := by
|
||||
simp only
|
||||
|
@ -224,7 +224,7 @@ def accSU3 : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
|||
ring
|
||||
|
||||
/-- Extensionality lemma for `accSU3`. -/
|
||||
lemma accSU3_ext {S T : MSSMCharges.charges}
|
||||
lemma accSU3_ext {S T : MSSMCharges.Charges}
|
||||
(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,
|
||||
|
@ -236,7 +236,7 @@ lemma accSU3_ext {S T : MSSMCharges.charges}
|
|||
|
||||
/-- The acc for `Y²`. -/
|
||||
@[simp]
|
||||
def accYY : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
||||
def accYY : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
||||
toFun S := ∑ i, ((Q S) i + 8 * (U S) i + 2 * (D S) i + 3 * (L S) i
|
||||
+ 6 * (E S) i) + 3 * (Hd S + Hu S)
|
||||
map_add' S T := by
|
||||
|
@ -256,7 +256,7 @@ def accYY : MSSMCharges.charges →ₗ[ℚ] ℚ where
|
|||
ring
|
||||
|
||||
/-- Extensionality lemma for `accGrav`. -/
|
||||
lemma accYY_ext {S T : MSSMCharges.charges}
|
||||
lemma accYY_ext {S T : MSSMCharges.Charges}
|
||||
(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
|
||||
|
@ -270,7 +270,7 @@ 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₂ (
|
||||
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))
|
||||
|
@ -317,10 +317,10 @@ def quadBiLin : BiLinearSymm MSSMCharges.charges := BiLinearSymm.mk₂ (
|
|||
|
||||
/-- The quadratic ACC. -/
|
||||
@[simp]
|
||||
def accQuad : HomogeneousQuadratic MSSMCharges.charges := quadBiLin.toHomogeneousQuad
|
||||
def accQuad : HomogeneousQuadratic MSSMCharges.Charges := quadBiLin.toHomogeneousQuad
|
||||
|
||||
/-- Extensionality lemma for `accQuad`. -/
|
||||
lemma accQuad_ext {S T : (MSSMCharges).charges}
|
||||
lemma accQuad_ext {S T : (MSSMCharges).Charges}
|
||||
(h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSMSpecies j S) i =
|
||||
∑ i, ((fun a => a^2) ∘ toSMSpecies j T) i)
|
||||
(hd : Hd S = Hd T) (hu : Hu S = Hu T) :
|
||||
|
@ -339,7 +339,7 @@ lemma accQuad_ext {S T : (MSSMCharges).charges}
|
|||
/-- The function underlying the symmetric trilinear form used to define the cubic ACC. -/
|
||||
@[simp]
|
||||
def cubeTriLinToFun
|
||||
(S : MSSMCharges.charges × MSSMCharges.charges × MSSMCharges.charges) : ℚ :=
|
||||
(S : MSSMCharges.Charges × MSSMCharges.Charges × MSSMCharges.Charges) : ℚ :=
|
||||
∑ 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)
|
||||
|
@ -349,7 +349,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]
|
||||
|
@ -364,7 +364,7 @@ lemma cubeTriLinToFun_map_smul₁ (a : ℚ) (S T R : MSSMCharges.charges) :
|
|||
ring
|
||||
|
||||
|
||||
lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.charges) :
|
||||
lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.Charges) :
|
||||
cubeTriLinToFun (S + T, R, L) = cubeTriLinToFun (S, R, L) + cubeTriLinToFun (T, R, L) := by
|
||||
simp only [cubeTriLinToFun]
|
||||
rw [add_assoc, ← add_assoc (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _ _]
|
||||
|
@ -383,7 +383,7 @@ lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.charges) :
|
|||
ring
|
||||
|
||||
|
||||
lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.charges) :
|
||||
lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.Charges) :
|
||||
cubeTriLinToFun (S, T, R) = cubeTriLinToFun (T, S, R) := by
|
||||
simp only [cubeTriLinToFun, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply,
|
||||
Fin.reduceFinMk, Hu_apply]
|
||||
|
@ -393,7 +393,7 @@ lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.charges) :
|
|||
ring
|
||||
ring
|
||||
|
||||
lemma cubeTriLinToFun_swap2 (S T R : MSSMCharges.charges) :
|
||||
lemma cubeTriLinToFun_swap2 (S T R : MSSMCharges.Charges) :
|
||||
cubeTriLinToFun (S, T, R) = cubeTriLinToFun (S, R, T) := by
|
||||
simp only [cubeTriLinToFun, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply,
|
||||
Fin.reduceFinMk, Hu_apply]
|
||||
|
@ -405,7 +405,7 @@ lemma cubeTriLinToFun_swap2 (S T R : MSSMCharges.charges) :
|
|||
|
||||
/-- The symmetric trilinear form used to define the cubic ACC. -/
|
||||
@[simps!]
|
||||
def cubeTriLin : TriLinearSymm MSSMCharges.charges := TriLinearSymm.mk₃
|
||||
def cubeTriLin : TriLinearSymm MSSMCharges.Charges := TriLinearSymm.mk₃
|
||||
cubeTriLinToFun
|
||||
cubeTriLinToFun_map_smul₁
|
||||
cubeTriLinToFun_map_add₁
|
||||
|
@ -414,10 +414,10 @@ def cubeTriLin : TriLinearSymm MSSMCharges.charges := TriLinearSymm.mk₃
|
|||
|
||||
/-- The cubic ACC. -/
|
||||
@[simp]
|
||||
def accCube : HomogeneousCubic MSSMCharges.charges := cubeTriLin.toCubic
|
||||
def accCube : HomogeneousCubic MSSMCharges.Charges := cubeTriLin.toCubic
|
||||
|
||||
/-- Extensionality lemma for `accCube`. -/
|
||||
lemma accCube_ext {S T : MSSMCharges.charges}
|
||||
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) :
|
||||
|
@ -464,7 +464,7 @@ lemma quadSol (S : MSSMACC.QuadSols) : accQuad S.val = 0 := by
|
|||
|
||||
/-- A solution from a charge satisfying the ACCs. -/
|
||||
@[simp]
|
||||
def AnomalyFreeMk (S : MSSMACC.charges) (hg : accGrav S = 0)
|
||||
def AnomalyFreeMk (S : MSSMACC.Charges) (hg : accGrav S = 0)
|
||||
(hsu2 : accSU2 S = 0) (hsu3 : accSU3 S = 0) (hyy : accYY S = 0)
|
||||
(hquad : accQuad S = 0) (hcube : accCube S = 0) : MSSMACC.Sols :=
|
||||
⟨⟨⟨S, by
|
||||
|
@ -481,7 +481,7 @@ def AnomalyFreeMk (S : MSSMACC.charges) (hg : accGrav S = 0)
|
|||
| 0 => exact hquad
|
||||
⟩ , by exact hcube ⟩
|
||||
|
||||
lemma AnomalyFreeMk_val (S : MSSMACC.charges) (hg : accGrav S = 0)
|
||||
lemma AnomalyFreeMk_val (S : MSSMACC.Charges) (hg : accGrav S = 0)
|
||||
(hsu2 : accSU2 S = 0) (hsu3 : accSU3 S = 0) (hyy : accYY S = 0)
|
||||
(hquad : accQuad S = 0) (hcube : accCube S = 0) :
|
||||
(AnomalyFreeMk S hg hsu2 hsu3 hyy hquad hcube).val = S := by
|
||||
|
@ -521,7 +521,7 @@ lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols)
|
|||
|
||||
/-- The dot product on the vector space of charges. -/
|
||||
@[simps!]
|
||||
def dot : BiLinearSymm MSSMCharges.charges := BiLinearSymm.mk₂
|
||||
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)
|
||||
|
|
|
@ -20,7 +20,7 @@ open MSSMACCs
|
|||
open BigOperators
|
||||
|
||||
/-- The hypercharge as an element of `MSSMACC.charges`. -/
|
||||
def YAsCharge : MSSMACC.charges := toSpecies.invFun
|
||||
def YAsCharge : MSSMACC.Charges := toSpecies.invFun
|
||||
⟨fun s => fun i =>
|
||||
match s, i with
|
||||
| 0, 0 => 1
|
||||
|
|
|
@ -38,14 +38,14 @@ namespace AnomalyFreePerp
|
|||
|
||||
/-- A condition for the quad line in the plane spanned by R, Y₃ and B₃ to sit in the cubic,
|
||||
and for the cube line to sit in the quad. -/
|
||||
def lineEqProp (R : MSSMACC.AnomalyFreePerp) : Prop := α₁ R = 0 ∧ α₂ R = 0 ∧ α₃ R = 0
|
||||
def LineEqProp (R : MSSMACC.AnomalyFreePerp) : Prop := α₁ R = 0 ∧ α₂ R = 0 ∧ α₃ R = 0
|
||||
|
||||
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (lineEqProp R) := by
|
||||
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (LineEqProp R) := by
|
||||
apply And.decidable
|
||||
|
||||
/-- A condition on `Sols` which we will show in `linEqPropSol_iff_proj_linEqProp` that is equivalent
|
||||
to the condition that the `proj` of the solution satisfies `lineEqProp`. -/
|
||||
def lineEqPropSol (R : MSSMACC.Sols) : Prop :=
|
||||
def LineEqPropSol (R : MSSMACC.Sols) : Prop :=
|
||||
cubeTriLin R.val R.val Y₃.val * quadBiLin B₃.val R.val -
|
||||
cubeTriLin R.val R.val B₃.val * quadBiLin Y₃.val R.val = 0
|
||||
|
||||
|
@ -54,8 +54,8 @@ equivalent to satisfying `lineEqPropSol`. -/
|
|||
def lineEqCoeff (T : MSSMACC.Sols) : ℚ := dot Y₃.val B₃.val * α₃ (proj T.1.1)
|
||||
|
||||
lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) :
|
||||
lineEqPropSol T ↔ lineEqCoeff T = 0 := by
|
||||
rw [lineEqPropSol, lineEqCoeff, α₃]
|
||||
LineEqPropSol T ↔ lineEqCoeff T = 0 := by
|
||||
rw [LineEqPropSol, lineEqCoeff, α₃]
|
||||
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero,
|
||||
false_or]
|
||||
rw [cube_proj_proj_B₃, cube_proj_proj_Y₃, quad_B₃_proj, quad_Y₃_proj]
|
||||
|
@ -66,8 +66,8 @@ lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) :
|
|||
simp
|
||||
|
||||
lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) :
|
||||
lineEqPropSol R ↔ lineEqProp (proj R.1.1) := by
|
||||
rw [lineEqPropSol_iff_lineEqCoeff_zero, lineEqCoeff, lineEqProp]
|
||||
LineEqPropSol R ↔ LineEqProp (proj R.1.1) := by
|
||||
rw [lineEqPropSol_iff_lineEqCoeff_zero, lineEqCoeff, LineEqProp]
|
||||
apply Iff.intro
|
||||
intro h
|
||||
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
||||
|
@ -81,15 +81,15 @@ lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) :
|
|||
|
||||
/-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies
|
||||
entirely in the quadratic surface. -/
|
||||
def inQuadProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
|
||||
def InQuadProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
|
||||
quadBiLin R.val R.val = 0 ∧ quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0
|
||||
|
||||
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (inQuadProp R) := by
|
||||
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InQuadProp R) := by
|
||||
apply And.decidable
|
||||
|
||||
/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
|
||||
lies entirely in the quadratic surface. -/
|
||||
def inQuadSolProp (R : MSSMACC.Sols) : Prop :=
|
||||
def InQuadSolProp (R : MSSMACC.Sols) : Prop :=
|
||||
quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0
|
||||
|
||||
/-- A rational which has two properties. It is zero for a solution `T` if and only if
|
||||
|
@ -98,7 +98,7 @@ def quadCoeff (T : MSSMACC.Sols) : ℚ :=
|
|||
2 * dot Y₃.val B₃.val ^ 2 *
|
||||
(quadBiLin Y₃.val T.val ^ 2 + quadBiLin B₃.val T.val ^ 2)
|
||||
|
||||
lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : inQuadSolProp T ↔ quadCoeff T = 0 := by
|
||||
lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : InQuadSolProp T ↔ quadCoeff T = 0 := by
|
||||
apply Iff.intro
|
||||
intro h
|
||||
rw [quadCoeff, h.1, h.2]
|
||||
|
@ -117,8 +117,8 @@ lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : inQuadSolProp T ↔
|
|||
/-- The conditions `inQuadSolProp R` and `inQuadProp (proj R.1.1)` are equivalent. This is to be
|
||||
expected since both `R` and `proj R.1.1` define the same plane with `Y₃` and `B₃`. -/
|
||||
lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
|
||||
inQuadSolProp R ↔ inQuadProp (proj R.1.1) := by
|
||||
rw [inQuadSolProp, inQuadProp]
|
||||
InQuadSolProp R ↔ InQuadProp (proj R.1.1) := by
|
||||
rw [InQuadSolProp, InQuadProp]
|
||||
rw [quad_proj, quad_Y₃_proj, quad_B₃_proj]
|
||||
apply Iff.intro
|
||||
intro h
|
||||
|
@ -133,17 +133,17 @@ lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
|
|||
|
||||
/-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies
|
||||
entirely in the cubic surface. -/
|
||||
def inCubeProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
|
||||
def InCubeProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
|
||||
cubeTriLin R.val R.val R.val = 0 ∧ cubeTriLin R.val R.val B₃.val = 0 ∧
|
||||
cubeTriLin R.val R.val Y₃.val = 0
|
||||
|
||||
|
||||
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (inCubeProp R) := by
|
||||
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InCubeProp R) := by
|
||||
apply And.decidable
|
||||
|
||||
/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
|
||||
lies entirely in the cubic surface. -/
|
||||
def inCubeSolProp (R : MSSMACC.Sols) : Prop :=
|
||||
def InCubeSolProp (R : MSSMACC.Sols) : Prop :=
|
||||
cubeTriLin R.val R.val B₃.val = 0 ∧ cubeTriLin R.val R.val Y₃.val = 0
|
||||
|
||||
/-- A rational which has two properties. It is zero for a solution `T` if and only if
|
||||
|
@ -153,7 +153,7 @@ def cubicCoeff (T : MSSMACC.Sols) : ℚ :=
|
|||
cubeTriLin T.val T.val B₃.val ^ 2 )
|
||||
|
||||
lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
||||
inCubeSolProp T ↔ cubicCoeff T = 0 := by
|
||||
InCubeSolProp T ↔ cubicCoeff T = 0 := by
|
||||
apply Iff.intro
|
||||
intro h
|
||||
rw [cubicCoeff, h.1, h.2]
|
||||
|
@ -170,8 +170,8 @@ lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
|||
exact h.symm
|
||||
|
||||
lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) :
|
||||
inCubeSolProp R ↔ inCubeProp (proj R.1.1) := by
|
||||
rw [inCubeSolProp, inCubeProp]
|
||||
InCubeSolProp R ↔ InCubeProp (proj R.1.1) := by
|
||||
rw [InCubeSolProp, InCubeProp]
|
||||
rw [cube_proj, cube_proj_proj_Y₃, cube_proj_proj_B₃]
|
||||
apply Iff.intro
|
||||
intro h
|
||||
|
@ -186,30 +186,30 @@ lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) :
|
|||
|
||||
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
|
||||
`lineEqProp`. -/
|
||||
def inLineEq : Type := {R : MSSMACC.AnomalyFreePerp // lineEqProp R}
|
||||
def InLineEq : Type := {R : MSSMACC.AnomalyFreePerp // LineEqProp R}
|
||||
|
||||
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
|
||||
`lineEqProp` and `inQuadProp`. -/
|
||||
def inQuad : Type := {R : inLineEq // inQuadProp R.val}
|
||||
def InQuad : Type := {R : InLineEq // InQuadProp R.val}
|
||||
|
||||
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
|
||||
`lineEqProp`, `inQuadProp` and `inCubeProp`. -/
|
||||
def inQuadCube : Type := {R : inQuad // inCubeProp R.val.val}
|
||||
def InQuadCube : Type := {R : InQuad // InCubeProp R.val.val}
|
||||
|
||||
/-- Those solutions which do not satisfy the condition `lineEqPropSol`. -/
|
||||
def notInLineEqSol : Type := {R : MSSMACC.Sols // ¬ lineEqPropSol R}
|
||||
def NotInLineEqSol : Type := {R : MSSMACC.Sols // ¬ LineEqPropSol R}
|
||||
|
||||
/-- Those solutions which satisfy the condition `lineEqPropSol` by not `inQuadSolProp`. -/
|
||||
def inLineEqSol : Type := {R : MSSMACC.Sols // lineEqPropSol R ∧ ¬ inQuadSolProp R}
|
||||
def InLineEqSol : Type := {R : MSSMACC.Sols // LineEqPropSol R ∧ ¬ InQuadSolProp R}
|
||||
|
||||
/-- Those solutions which satisfy the condition `lineEqPropSol` and `inQuadSolProp` but
|
||||
not `inCubeSolProp`. -/
|
||||
def inQuadSol : Type := {R : MSSMACC.Sols // lineEqPropSol R ∧ inQuadSolProp R ∧ ¬ inCubeSolProp R}
|
||||
def InQuadSol : Type := {R : MSSMACC.Sols // LineEqPropSol R ∧ InQuadSolProp R ∧ ¬ InCubeSolProp R}
|
||||
|
||||
/-- Those solutions which satisfy the condition all the conditions `lineEqPropSol`, `inQuadSolProp`
|
||||
and `inCubeSolProp`. -/
|
||||
def inQuadCubeSol : Type :=
|
||||
{R : MSSMACC.Sols // lineEqPropSol R ∧ inQuadSolProp R ∧ inCubeSolProp R}
|
||||
def InQuadCubeSol : Type :=
|
||||
{R : MSSMACC.Sols // LineEqPropSol R ∧ InQuadSolProp R ∧ InCubeSolProp R}
|
||||
|
||||
/-- Given a `R` perpendicular to `Y₃` and `B₃` a quadratic solution. -/
|
||||
def toSolNSQuad (R : MSSMACC.AnomalyFreePerp) : MSSMACC.QuadSols :=
|
||||
|
@ -244,7 +244,7 @@ def toSolNS : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := f
|
|||
def toSolNSProj (T : MSSMACC.Sols) : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ :=
|
||||
(proj T.1.1, (lineEqCoeff T)⁻¹, 0, 0)
|
||||
|
||||
lemma toSolNS_proj (T : notInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val := by
|
||||
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 = _
|
||||
|
@ -263,7 +263,7 @@ lemma toSolNS_proj (T : notInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val :=
|
|||
simp
|
||||
|
||||
/-- Given a element of `inLineEq × ℚ × ℚ × ℚ`, a solution to the ACCs. -/
|
||||
def inLineEqToSol : inLineEq × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, c₁, c₂, c₃) =>
|
||||
def inLineEqToSol : InLineEq × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, c₁, c₂, c₃) =>
|
||||
AnomalyFreeMk'' (lineQuad R.val c₁ c₂ c₃)
|
||||
(by
|
||||
rw [lineQuad_cube]
|
||||
|
@ -271,7 +271,7 @@ def inLineEqToSol : inLineEq × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, c
|
|||
simp)
|
||||
|
||||
/-- On elements of `inLineEqSol` a right-inverse to `inLineEqSol`. -/
|
||||
def inLineEqProj (T : inLineEqSol) : inLineEq × ℚ × ℚ × ℚ :=
|
||||
def inLineEqProj (T : InLineEqSol) : InLineEq × ℚ × ℚ × ℚ :=
|
||||
(⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
|
||||
(quadCoeff T.val)⁻¹ * quadBiLin B₃.val T.val.val,
|
||||
(quadCoeff T.val)⁻¹ * (- quadBiLin Y₃.val T.val.val),
|
||||
|
@ -279,14 +279,14 @@ def inLineEqProj (T : inLineEqSol) : inLineEq × ℚ × ℚ × ℚ :=
|
|||
quadBiLin B₃.val T.val.val * (dot B₃.val T.val.val - dot Y₃.val T.val.val)
|
||||
- quadBiLin Y₃.val T.val.val * (dot Y₃.val T.val.val - 2 * dot B₃.val T.val.val)))
|
||||
|
||||
lemma inLineEqTo_smul (R : inLineEq) (c₁ c₂ c₃ d : ℚ) :
|
||||
lemma inLineEqTo_smul (R : InLineEq) (c₁ c₂ c₃ d : ℚ) :
|
||||
inLineEqToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inLineEqToSol (R, c₁, c₂, c₃) := by
|
||||
apply ACCSystem.Sols.ext
|
||||
change (lineQuad _ _ _ _).val = _
|
||||
rw [lineQuad_smul]
|
||||
rfl
|
||||
|
||||
lemma inLineEqToSol_proj (T : inLineEqSol) : inLineEqToSol (inLineEqProj T) = T.val := by
|
||||
lemma inLineEqToSol_proj (T : InLineEqSol) : inLineEqToSol (inLineEqProj T) = T.val := by
|
||||
rw [inLineEqProj, inLineEqTo_smul]
|
||||
apply ACCSystem.Sols.ext
|
||||
change _ • (lineQuad _ _ _ _).val = _
|
||||
|
@ -306,7 +306,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]
|
||||
|
@ -314,7 +314,7 @@ def inQuadToSol : inQuad × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a₁,
|
|||
simp)
|
||||
(lineCube_cube R.val.val a₁ a₂ a₃)
|
||||
|
||||
lemma inQuadToSol_smul (R : inQuad) (c₁ c₂ c₃ d : ℚ) :
|
||||
lemma inQuadToSol_smul (R : InQuad) (c₁ c₂ c₃ d : ℚ) :
|
||||
inQuadToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadToSol (R, c₁, c₂, c₃) := by
|
||||
apply ACCSystem.Sols.ext
|
||||
change (lineCube _ _ _ _).val = _
|
||||
|
@ -322,7 +322,7 @@ lemma inQuadToSol_smul (R : inQuad) (c₁ c₂ c₃ d : ℚ) :
|
|||
rfl
|
||||
|
||||
/-- On elements of `inQuadSol` a right-inverse to `inQuadToSol`. -/
|
||||
def inQuadProj (T : inQuadSol) : inQuad × ℚ × ℚ × ℚ :=
|
||||
def inQuadProj (T : InQuadSol) : InQuad × ℚ × ℚ × ℚ :=
|
||||
(⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
|
||||
(inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩,
|
||||
(cubicCoeff T.val)⁻¹ * (cubeTriLin T.val.val T.val.val B₃.val),
|
||||
|
@ -332,7 +332,7 @@ def inQuadProj (T : inQuadSol) : inQuad × ℚ × ℚ × ℚ :=
|
|||
- cubeTriLin T.val.val T.val.val Y₃.val
|
||||
* (dot Y₃.val T.val.val - 2 * dot B₃.val T.val.val)))
|
||||
|
||||
lemma inQuadToSol_proj (T : inQuadSol) : inQuadToSol (inQuadProj T) = T.val := by
|
||||
lemma inQuadToSol_proj (T : InQuadSol) : inQuadToSol (inQuadProj T) = T.val := by
|
||||
rw [inQuadProj, inQuadToSol_smul]
|
||||
apply ACCSystem.Sols.ext
|
||||
change _ • (planeY₃B₃ _ _ _ _).val = _
|
||||
|
@ -352,7 +352,7 @@ lemma inQuadToSol_proj (T : inQuadSol) : inQuadToSol (inQuadProj T) = T.val := b
|
|||
simp
|
||||
|
||||
/-- Given a element of `inQuadCube × ℚ × ℚ × ℚ`, a solution to the ACCs. -/
|
||||
def inQuadCubeToSol : inQuadCube × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, b₁, b₂, b₃) =>
|
||||
def inQuadCubeToSol : InQuadCube × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, b₁, b₂, b₃) =>
|
||||
AnomalyFreeMk' (planeY₃B₃ R.val.val.val b₁ b₂ b₃)
|
||||
(by
|
||||
rw [planeY₃B₃_quad]
|
||||
|
@ -363,7 +363,7 @@ def inQuadCubeToSol : inQuadCube × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R
|
|||
rw [R.prop.1, R.prop.2.1, R.prop.2.2]
|
||||
simp)
|
||||
|
||||
lemma inQuadCubeToSol_smul (R : inQuadCube) (c₁ c₂ c₃ d : ℚ) :
|
||||
lemma inQuadCubeToSol_smul (R : InQuadCube) (c₁ c₂ c₃ d : ℚ) :
|
||||
inQuadCubeToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadCubeToSol (R, c₁, c₂, c₃):= by
|
||||
apply ACCSystem.Sols.ext
|
||||
change (planeY₃B₃ _ _ _ _).val = _
|
||||
|
@ -371,7 +371,7 @@ lemma inQuadCubeToSol_smul (R : inQuadCube) (c₁ c₂ c₃ d : ℚ) :
|
|||
rfl
|
||||
|
||||
/-- On elements of `inQuadCubeSol` a right-inverse to `inQuadCubeToSol`. -/
|
||||
def inQuadCubeProj (T : inQuadCubeSol) : inQuadCube × ℚ × ℚ × ℚ :=
|
||||
def inQuadCubeProj (T : InQuadCubeSol) : InQuadCube × ℚ × ℚ × ℚ :=
|
||||
(⟨⟨⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
|
||||
(inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1⟩,
|
||||
(inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2⟩,
|
||||
|
@ -379,7 +379,7 @@ def inQuadCubeProj (T : inQuadCubeSol) : inQuadCube × ℚ × ℚ × ℚ :=
|
|||
(dot Y₃.val B₃.val)⁻¹ * (2 * dot B₃.val T.val.val - dot Y₃.val T.val.val),
|
||||
(dot Y₃.val B₃.val)⁻¹ * 1)
|
||||
|
||||
lemma inQuadCubeToSol_proj (T : inQuadCubeSol) :
|
||||
lemma inQuadCubeToSol_proj (T : InQuadCubeSol) :
|
||||
inQuadCubeToSol (inQuadCubeProj T) = T.val := by
|
||||
rw [inQuadCubeProj, inQuadCubeToSol_smul]
|
||||
apply ACCSystem.Sols.ext
|
||||
|
@ -396,64 +396,64 @@ 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
|
||||
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
|
||||
if h₂ : LineEqProp R ∧ InQuadProp R then
|
||||
inQuadToSol (⟨⟨R, h₂.1⟩, h₂.2⟩, a, b, c)
|
||||
else
|
||||
if h₁ : lineEqProp R then
|
||||
if h₁ : LineEqProp R then
|
||||
inLineEqToSol (⟨R, h₁⟩, a, b, c)
|
||||
else
|
||||
toSolNS ⟨R, a, b, c⟩
|
||||
|
||||
lemma toSol_toSolNSProj (T : notInLineEqSol) :
|
||||
lemma toSol_toSolNSProj (T : NotInLineEqSol) :
|
||||
∃ X, toSol X = T.val := by
|
||||
use toSolNSProj T.val
|
||||
have h1 : ¬ lineEqProp (toSolNSProj T.val).1 :=
|
||||
have h1 : ¬ LineEqProp (toSolNSProj T.val).1 :=
|
||||
(linEqPropSol_iff_proj_linEqProp T.val).mpr.mt T.prop
|
||||
rw [toSol]
|
||||
simp_all
|
||||
exact toSolNS_proj T
|
||||
|
||||
lemma toSol_inLineEq (T : inLineEqSol) : ∃ X, toSol X = T.val := by
|
||||
lemma toSol_inLineEq (T : InLineEqSol) : ∃ X, toSol X = T.val := by
|
||||
let X := inLineEqProj T
|
||||
use ⟨X.1.val, X.2.1, X.2.2⟩
|
||||
have : ¬ inQuadProp X.1.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mpr.mt T.prop.2
|
||||
have : lineEqProp X.1.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
|
||||
have : ¬ InQuadProp X.1.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mpr.mt T.prop.2
|
||||
have : LineEqProp X.1.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
|
||||
rw [toSol]
|
||||
simp_all
|
||||
exact inLineEqToSol_proj T
|
||||
|
||||
lemma toSol_inQuad (T : inQuadSol) : ∃ X, toSol X = T.val := by
|
||||
lemma toSol_inQuad (T : InQuadSol) : ∃ X, toSol X = T.val := by
|
||||
let X := inQuadProj T
|
||||
use ⟨X.1.val.val, X.2.1, X.2.2⟩
|
||||
have : ¬ inCubeProp X.1.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mpr.mt T.prop.2.2
|
||||
have : inQuadProp X.1.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1
|
||||
have : lineEqProp X.1.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
|
||||
have : ¬ InCubeProp X.1.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mpr.mt T.prop.2.2
|
||||
have : InQuadProp X.1.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1
|
||||
have : LineEqProp X.1.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
|
||||
rw [toSol]
|
||||
simp_all
|
||||
exact inQuadToSol_proj T
|
||||
|
||||
lemma toSol_inQuadCube (T : inQuadCubeSol) : ∃ X, toSol X = T.val := by
|
||||
lemma toSol_inQuadCube (T : InQuadCubeSol) : ∃ X, toSol X = T.val := by
|
||||
let X := inQuadCubeProj T
|
||||
use ⟨X.1.val.val.val, X.2.1, X.2.2⟩
|
||||
have : inCubeProp X.1.val.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2
|
||||
have : inQuadProp X.1.val.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1
|
||||
have : lineEqProp X.1.val.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
|
||||
have : InCubeProp X.1.val.val.val := (inCubeSolProp_iff_proj_inCubeProp T.val).mp T.prop.2.2
|
||||
have : InQuadProp X.1.val.val.val := (inQuadSolProp_iff_proj_inQuadProp T.val).mp T.prop.2.1
|
||||
have : LineEqProp X.1.val.val.val := (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1
|
||||
rw [toSol]
|
||||
simp_all
|
||||
exact inQuadCubeToSol_proj T
|
||||
|
||||
theorem toSol_surjective : Function.Surjective toSol := by
|
||||
intro T
|
||||
by_cases h₁ : ¬ lineEqPropSol T
|
||||
by_cases h₁ : ¬ LineEqPropSol T
|
||||
exact toSol_toSolNSProj ⟨T, h₁⟩
|
||||
simp at h₁
|
||||
by_cases h₂ : ¬ inQuadSolProp T
|
||||
by_cases h₂ : ¬ InQuadSolProp T
|
||||
exact toSol_inLineEq ⟨T, And.intro h₁ h₂⟩
|
||||
simp at h₂
|
||||
by_cases h₃ : ¬ inCubeSolProp T
|
||||
by_cases h₃ : ¬ InCubeSolProp T
|
||||
exact toSol_inQuad ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩
|
||||
simp at h₃
|
||||
exact toSol_inQuadCube ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩
|
||||
|
|
|
@ -27,14 +27,14 @@ open BigOperators
|
|||
|
||||
/-- The group of family permutations is `S₃⁶`-/
|
||||
@[simp]
|
||||
def permGroup := Fin 6 → Equiv.Perm (Fin 3)
|
||||
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!]
|
||||
def chargeMap (f : permGroup) : MSSMCharges.charges →ₗ[ℚ] MSSMCharges.charges where
|
||||
def chargeMap (f : PermGroup) : MSSMCharges.Charges →ₗ[ℚ] MSSMCharges.Charges where
|
||||
toFun S := toSpecies.symm (fun i => toSMSpecies i S ∘ f i, Prod.snd (toSpecies S))
|
||||
map_add' S T := by
|
||||
simp only
|
||||
|
@ -52,16 +52,16 @@ def chargeMap (f : permGroup) : MSSMCharges.charges →ₗ[ℚ] MSSMCharges.char
|
|||
rw [(toSMSpecies i).map_smul, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
|
||||
rfl
|
||||
|
||||
lemma chargeMap_toSpecies (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) :
|
||||
lemma chargeMap_toSpecies (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
|
||||
toSMSpecies j (chargeMap f S) = toSMSpecies j S ∘ f j := by
|
||||
erw [toSMSpecies_toSpecies_inv]
|
||||
|
||||
/-- The representation of `permGroup` acting on the vector space of charges. -/
|
||||
@[simp]
|
||||
def repCharges : Representation ℚ (permGroup) (MSSMCharges).charges where
|
||||
def repCharges : Representation ℚ PermGroup (MSSMCharges).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]
|
||||
|
@ -81,56 +81,56 @@ def repCharges : Representation ℚ (permGroup) (MSSMCharges).charges where
|
|||
erw [toSMSpecies_toSpecies_inv]
|
||||
rfl
|
||||
|
||||
lemma repCharges_toSMSpecies (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) :
|
||||
lemma repCharges_toSMSpecies (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
|
||||
toSMSpecies j (repCharges f S) = toSMSpecies j S ∘ f⁻¹ j := by
|
||||
erw [toSMSpecies_toSpecies_inv]
|
||||
|
||||
lemma toSpecies_sum_invariant (m : ℕ) (f : permGroup) (S : MSSMCharges.charges) (j : Fin 6) :
|
||||
lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
|
||||
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j (repCharges f S)) i =
|
||||
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j S) i := by
|
||||
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)
|
||||
|
|
|
@ -26,7 +26,7 @@ open BigOperators
|
|||
|
||||
/-- $Y_3$ is the charge which is hypercharge in all families, but with the third
|
||||
family of the opposite sign. -/
|
||||
def Y₃AsCharge : MSSMACC.charges := toSpecies.symm
|
||||
def Y₃AsCharge : MSSMACC.Charges := toSpecies.symm
|
||||
⟨fun s => fun i =>
|
||||
match s, i with
|
||||
| 0, 0 => 1
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue