2024-04-17 16:23:40 -04:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
|
|
|
|
import HepLean.AnomalyCancellation.MSSMNu.LineY3B3
|
2024-04-22 06:57:46 -04:00
|
|
|
|
import HepLean.AnomalyCancellation.MSSMNu.OrthogY3B3.PlaneWithY3B3
|
2024-04-17 16:23:40 -04:00
|
|
|
|
import Mathlib.Tactic.Polyrith
|
|
|
|
|
/-!
|
2024-04-22 06:57:46 -04:00
|
|
|
|
# From charges perpendicular to `Y₃` and `B₃` to solutions
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
The main aim of this file is to take charge assignments perpendicular to `Y₃` and `B₃` and
|
|
|
|
|
produce solutions to the anomaly cancellation conditions. In this regard we will define
|
|
|
|
|
a surjective map `toSol` from `MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ` to `MSSMACC.Sols`.
|
|
|
|
|
|
|
|
|
|
To define `toSols` we define a series of other maps from various subtypes of
|
|
|
|
|
`MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ` to `MSSMACC.Sols`. And show that these maps form a
|
|
|
|
|
surjection on certain subtypes of `MSSMACC.Sols`.
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
# References
|
|
|
|
|
|
|
|
|
|
The main reference for the material in this file is:
|
|
|
|
|
|
|
|
|
|
- https://arxiv.org/pdf/2107.07926.pdf
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
universe v u
|
|
|
|
|
|
|
|
|
|
namespace MSSMACC
|
|
|
|
|
open MSSMCharges
|
|
|
|
|
open MSSMACCs
|
|
|
|
|
open BigOperators
|
|
|
|
|
|
2024-04-19 16:10:29 -04:00
|
|
|
|
namespace AnomalyFreePerp
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- 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. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def LineEqProp (R : MSSMACC.AnomalyFreePerp) : Prop := α₁ R = 0 ∧ α₂ R = 0 ∧ α₃ R = 0
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (LineEqProp R) := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
apply And.decidable
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- 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`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def LineEqPropSol (R : MSSMACC.Sols) : Prop :=
|
2024-04-22 09:48:44 -04:00
|
|
|
|
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
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- A rational which appears in `toSolNS` acting on sols, and which been zero is
|
|
|
|
|
equivalent to satisfying `lineEqPropSol`. -/
|
2024-04-22 08:41:50 -04:00
|
|
|
|
def lineEqCoeff (T : MSSMACC.Sols) : ℚ := dot Y₃.val B₃.val * α₃ (proj T.1.1)
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
|
|
|
|
lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) :
|
2024-06-26 11:54:02 -04:00
|
|
|
|
LineEqPropSol T ↔ lineEqCoeff T = 0 := by
|
|
|
|
|
rw [LineEqPropSol, lineEqCoeff, α₃]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
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]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
rw [show dot Y₃.val B₃.val = 108 by rfl]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, OfNat.ofNat_ne_zero, false_or]
|
2024-04-22 06:57:46 -04:00
|
|
|
|
ring_nf
|
|
|
|
|
rw [mul_comm _ 1259712, mul_comm _ 1259712, ← mul_sub]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp
|
|
|
|
|
|
|
|
|
|
lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) :
|
2024-06-26 11:54:02 -04:00
|
|
|
|
LineEqPropSol R ↔ LineEqProp (proj R.1.1) := by
|
|
|
|
|
rw [lineEqPropSol_iff_lineEqCoeff_zero, lineEqCoeff, LineEqProp]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
apply Iff.intro
|
|
|
|
|
intro h
|
2024-04-22 08:41:50 -04:00
|
|
|
|
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp at h
|
|
|
|
|
rw [α₁_proj, α₂_proj, h]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
simp only [neg_zero, Fin.isValue, Fin.reduceFinMk, zero_mul, and_self]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
intro h
|
|
|
|
|
rw [h.2.2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp
|
|
|
|
|
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies
|
|
|
|
|
entirely in the quadratic surface. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InQuadProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
|
2024-04-22 08:41:50 -04:00
|
|
|
|
quadBiLin R.val R.val = 0 ∧ quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InQuadProp R) := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
apply And.decidable
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
|
|
|
|
|
lies entirely in the quadratic surface. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InQuadSolProp (R : MSSMACC.Sols) : Prop :=
|
2024-04-22 08:41:50 -04:00
|
|
|
|
quadBiLin Y₃.val R.val = 0 ∧ quadBiLin B₃.val R.val = 0
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- A rational which has two properties. It is zero for a solution `T` if and only if
|
|
|
|
|
that solution satisfies `inQuadSolProp`. It appears in the definition of `inQuadProj`. -/
|
2024-04-19 16:10:29 -04:00
|
|
|
|
def quadCoeff (T : MSSMACC.Sols) : ℚ :=
|
2024-04-22 08:41:50 -04:00
|
|
|
|
2 * dot Y₃.val B₃.val ^ 2 *
|
|
|
|
|
(quadBiLin Y₃.val T.val ^ 2 + quadBiLin B₃.val T.val ^ 2)
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : InQuadSolProp T ↔ quadCoeff T = 0 := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
apply Iff.intro
|
|
|
|
|
intro h
|
|
|
|
|
rw [quadCoeff, h.1, h.2]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
|
2024-04-22 07:00:17 -04:00
|
|
|
|
zero_pow, add_zero, mul_zero]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
intro h
|
|
|
|
|
rw [quadCoeff] at h
|
2024-04-22 08:41:50 -04:00
|
|
|
|
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp only [ Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
|
|
|
|
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
|
|
|
|
|
apply (add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)).mp at h
|
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero,
|
|
|
|
|
not_false_eq_true, pow_eq_zero_iff] at h
|
|
|
|
|
exact h
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- 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₃`. -/
|
2024-04-19 16:10:29 -04:00
|
|
|
|
lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
|
2024-06-26 11:54:02 -04:00
|
|
|
|
InQuadSolProp R ↔ InQuadProp (proj R.1.1) := by
|
|
|
|
|
rw [InQuadSolProp, InQuadProp]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [quad_proj, quad_Y₃_proj, quad_B₃_proj]
|
|
|
|
|
apply Iff.intro
|
|
|
|
|
intro h
|
|
|
|
|
rw [h.1, h.2]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
intro h
|
2024-04-22 08:41:50 -04:00
|
|
|
|
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk , mul_eq_zero,
|
2024-04-19 16:10:29 -04:00
|
|
|
|
OfNat.ofNat_ne_zero, or_self, false_or] at h
|
|
|
|
|
rw [h.2.1, h.2.2]
|
|
|
|
|
simp
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- A condition which is satisfied if the plane spanned by `R`, `Y₃` and `B₃` lies
|
|
|
|
|
entirely in the cubic surface. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InCubeProp (R : MSSMACC.AnomalyFreePerp) : Prop :=
|
2024-04-22 09:48:44 -04:00
|
|
|
|
cubeTriLin R.val R.val R.val = 0 ∧ cubeTriLin R.val R.val B₃.val = 0 ∧
|
|
|
|
|
cubeTriLin R.val R.val Y₃.val = 0
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
instance (R : MSSMACC.AnomalyFreePerp) : Decidable (InCubeProp R) := by
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply And.decidable
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- A condition which is satisfied if the plane spanned by the solutions `R`, `Y₃` and `B₃`
|
|
|
|
|
lies entirely in the cubic surface. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InCubeSolProp (R : MSSMACC.Sols) : Prop :=
|
2024-04-22 09:48:44 -04:00
|
|
|
|
cubeTriLin R.val R.val B₃.val = 0 ∧ cubeTriLin R.val R.val Y₃.val = 0
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- 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`. -/
|
2024-04-19 16:10:29 -04:00
|
|
|
|
def cubicCoeff (T : MSSMACC.Sols) : ℚ :=
|
2024-04-22 09:48:44 -04:00
|
|
|
|
3 * (dot Y₃.val B₃.val) ^ 3 * (cubeTriLin T.val T.val Y₃.val ^ 2 +
|
|
|
|
|
cubeTriLin T.val T.val B₃.val ^ 2 )
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-19 16:10:29 -04:00
|
|
|
|
lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
|
2024-06-26 11:54:02 -04:00
|
|
|
|
InCubeSolProp T ↔ cubicCoeff T = 0 := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
apply Iff.intro
|
|
|
|
|
intro h
|
|
|
|
|
rw [cubicCoeff, h.1, h.2]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
|
2024-04-22 07:00:17 -04:00
|
|
|
|
zero_pow, add_zero, mul_zero]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
intro h
|
|
|
|
|
rw [cubicCoeff] at h
|
2024-04-22 08:41:50 -04:00
|
|
|
|
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp only [ Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
|
|
|
|
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
|
|
|
|
|
apply (add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)).mp at h
|
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero,
|
|
|
|
|
not_false_eq_true, pow_eq_zero_iff] at h
|
|
|
|
|
exact h.symm
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-19 16:10:29 -04:00
|
|
|
|
lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) :
|
2024-06-26 11:54:02 -04:00
|
|
|
|
InCubeSolProp R ↔ InCubeProp (proj R.1.1) := by
|
|
|
|
|
rw [InCubeSolProp, InCubeProp]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [cube_proj, cube_proj_proj_Y₃, cube_proj_proj_B₃]
|
|
|
|
|
apply Iff.intro
|
|
|
|
|
intro h
|
|
|
|
|
rw [h.1, h.2]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
intro h
|
2024-04-22 08:41:50 -04:00
|
|
|
|
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
|
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
|
2024-04-19 16:10:29 -04:00
|
|
|
|
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
|
|
|
|
|
rw [h.2.1, h.2.2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
|
|
|
|
|
`lineEqProp`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InLineEq : Type := {R : MSSMACC.AnomalyFreePerp // LineEqProp R}
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
|
|
|
|
|
`lineEqProp` and `inQuadProp`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InQuad : Type := {R : InLineEq // InQuadProp R.val}
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
|
|
|
|
/-- Those charge assignments perpendicular to `Y₃` and `B₃` which satisfy the condition
|
|
|
|
|
`lineEqProp`, `inQuadProp` and `inCubeProp`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InQuadCube : Type := {R : InQuad // InCubeProp R.val.val}
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
|
|
|
|
/-- Those solutions which do not satisfy the condition `lineEqPropSol`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def NotInLineEqSol : Type := {R : MSSMACC.Sols // ¬ LineEqPropSol R}
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
|
|
|
|
/-- Those solutions which satisfy the condition `lineEqPropSol` by not `inQuadSolProp`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InLineEqSol : Type := {R : MSSMACC.Sols // LineEqPropSol R ∧ ¬ InQuadSolProp R}
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
|
|
|
|
/-- Those solutions which satisfy the condition `lineEqPropSol` and `inQuadSolProp` but
|
|
|
|
|
not `inCubeSolProp`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InQuadSol : Type := {R : MSSMACC.Sols // LineEqPropSol R ∧ InQuadSolProp R ∧ ¬ InCubeSolProp R}
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
|
|
|
|
/-- Those solutions which satisfy the condition all the conditions `lineEqPropSol`, `inQuadSolProp`
|
|
|
|
|
and `inCubeSolProp`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def InQuadCubeSol : Type :=
|
|
|
|
|
{R : MSSMACC.Sols // LineEqPropSol R ∧ InQuadSolProp R ∧ InCubeSolProp R}
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
|
|
|
|
/-- Given a `R` perpendicular to `Y₃` and `B₃` a quadratic solution. -/
|
2024-04-19 16:10:29 -04:00
|
|
|
|
def toSolNSQuad (R : MSSMACC.AnomalyFreePerp) : MSSMACC.QuadSols :=
|
|
|
|
|
lineQuad R
|
2024-04-22 09:48:44 -04:00
|
|
|
|
(3 * cubeTriLin R.val R.val Y₃.val)
|
|
|
|
|
(3 * cubeTriLin R.val R.val B₃.val)
|
|
|
|
|
(cubeTriLin R.val R.val R.val)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-19 16:10:29 -04:00
|
|
|
|
lemma toSolNSQuad_cube (R : MSSMACC.AnomalyFreePerp) :
|
|
|
|
|
accCube (toSolNSQuad R).val = 0 := by
|
|
|
|
|
rw [toSolNSQuad]
|
|
|
|
|
rw [lineQuad_val]
|
|
|
|
|
rw [planeY₃B₃_cubic]
|
|
|
|
|
ring
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-19 16:10:29 -04:00
|
|
|
|
lemma toSolNSQuad_eq_planeY₃B₃_on_α (R : MSSMACC.AnomalyFreePerp) :
|
|
|
|
|
(toSolNSQuad R).1 = planeY₃B₃ R (α₁ R) (α₂ R) (α₃ R) := by
|
|
|
|
|
change (planeY₃B₃ _ _ _ _) = _
|
|
|
|
|
apply planeY₃B₃_eq
|
|
|
|
|
rw [α₁, α₂, α₃]
|
|
|
|
|
ring_nf
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp
|
|
|
|
|
|
2024-06-08 01:50:38 +02:00
|
|
|
|
/-- Given a `R ` perpendicular to `Y₃` and `B₃`, an element of `Sols`. This map is
|
2024-04-22 06:57:46 -04:00
|
|
|
|
not surjective. -/
|
|
|
|
|
def toSolNS : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a, _ , _) =>
|
2024-04-19 16:10:29 -04:00
|
|
|
|
a • AnomalyFreeMk'' (toSolNSQuad R) (toSolNSQuad_cube R)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
|
|
|
|
/-- A map from `Sols` to `MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ` which on elements of
|
|
|
|
|
`notInLineEqSol` will produce a right inverse to `toSolNS`. -/
|
2024-04-19 16:10:29 -04:00
|
|
|
|
def toSolNSProj (T : MSSMACC.Sols) : MSSMACC.AnomalyFreePerp × ℚ × ℚ × ℚ :=
|
|
|
|
|
(proj T.1.1, (lineEqCoeff T)⁻¹, 0, 0)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma toSolNS_proj (T : NotInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
apply ACCSystem.Sols.ext
|
|
|
|
|
rw [toSolNS, toSolNSProj]
|
2024-04-22 06:57:46 -04:00
|
|
|
|
change (lineEqCoeff T.val)⁻¹ • (toSolNSQuad _).1.1 = _
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [toSolNSQuad_eq_planeY₃B₃_on_α]
|
|
|
|
|
rw [planeY₃B₃_val]
|
|
|
|
|
rw [Y₃_plus_B₃_plus_proj]
|
|
|
|
|
rw [α₁_proj, α₂_proj]
|
|
|
|
|
ring_nf
|
|
|
|
|
simp only [zero_smul, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
have h1 : α₃ (proj T.val.toLinSols) * dot Y₃.val B₃.val = lineEqCoeff T.val := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [lineEqCoeff]
|
|
|
|
|
ring
|
|
|
|
|
rw [h1]
|
2024-04-22 06:57:46 -04:00
|
|
|
|
have h1 := (lineEqPropSol_iff_lineEqCoeff_zero T.val).mpr.mt T.prop
|
|
|
|
|
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h1]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- Given a element of `inLineEq × ℚ × ℚ × ℚ`, a solution to the ACCs. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def inLineEqToSol : InLineEq × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, c₁, c₂, c₃) =>
|
2024-04-19 16:10:29 -04:00
|
|
|
|
AnomalyFreeMk'' (lineQuad R.val c₁ c₂ c₃)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
(by
|
|
|
|
|
rw [lineQuad_cube]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [R.prop.1, R.prop.2.1, R.prop.2.2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp)
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- On elements of `inLineEqSol` a right-inverse to `inLineEqSol`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def inLineEqProj (T : InLineEqSol) : InLineEq × ℚ × ℚ × ℚ :=
|
2024-04-22 06:57:46 -04:00
|
|
|
|
(⟨proj T.val.1.1, (linEqPropSol_iff_proj_linEqProp T.val).mp T.prop.1 ⟩,
|
2024-04-22 08:41:50 -04:00
|
|
|
|
(quadCoeff T.val)⁻¹ * quadBiLin B₃.val T.val.val,
|
|
|
|
|
(quadCoeff T.val)⁻¹ * (- quadBiLin Y₃.val T.val.val),
|
2024-04-22 06:57:46 -04:00
|
|
|
|
(quadCoeff T.val)⁻¹ * (
|
2024-04-22 08:41:50 -04:00
|
|
|
|
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)))
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma inLineEqTo_smul (R : InLineEq) (c₁ c₂ c₃ d : ℚ) :
|
2024-04-19 16:10:29 -04:00
|
|
|
|
inLineEqToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inLineEqToSol (R, c₁, c₂, c₃) := by
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply ACCSystem.Sols.ext
|
|
|
|
|
change (lineQuad _ _ _ _).val = _
|
|
|
|
|
rw [lineQuad_smul]
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma inLineEqToSol_proj (T : InLineEqSol) : inLineEqToSol (inLineEqProj T) = T.val := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [inLineEqProj, inLineEqTo_smul]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply ACCSystem.Sols.ext
|
2024-04-19 16:10:29 -04:00
|
|
|
|
change _ • (lineQuad _ _ _ _).val = _
|
2024-04-17 16:23:40 -04:00
|
|
|
|
rw [lineQuad_val]
|
|
|
|
|
rw [planeY₃B₃_val]
|
|
|
|
|
rw [Y₃_plus_B₃_plus_proj]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [quad_proj, quad_Y₃_proj, quad_B₃_proj]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
ring_nf
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp only [zero_smul, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
have h1 : (quadBiLin Y₃.val T.val.val ^ 2 * dot Y₃.val B₃.val ^ 2 * 2 +
|
|
|
|
|
dot Y₃.val B₃.val ^ 2 * quadBiLin B₃.val T.val.val ^ 2 * 2) = quadCoeff T.val := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [quadCoeff]
|
|
|
|
|
ring
|
|
|
|
|
rw [h1]
|
2024-04-22 06:57:46 -04:00
|
|
|
|
have h2 := (inQuadSolProp_iff_quadCoeff_zero T.val).mpr.mt T.prop.2
|
|
|
|
|
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- Given a element of `inQuad × ℚ × ℚ × ℚ`, a solution to the ACCs. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def inQuadToSol : InQuad × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, a₁, a₂, a₃) =>
|
2024-04-19 16:10:29 -04:00
|
|
|
|
AnomalyFreeMk' (lineCube R.val.val a₁ a₂ a₃)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
(by
|
|
|
|
|
erw [planeY₃B₃_quad]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [R.prop.1, R.prop.2.1, R.prop.2.2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp)
|
2024-04-19 16:10:29 -04:00
|
|
|
|
(lineCube_cube R.val.val a₁ a₂ a₃)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma inQuadToSol_smul (R : InQuad) (c₁ c₂ c₃ d : ℚ) :
|
2024-04-19 16:10:29 -04:00
|
|
|
|
inQuadToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadToSol (R, c₁, c₂, c₃) := by
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply ACCSystem.Sols.ext
|
|
|
|
|
change (lineCube _ _ _ _).val = _
|
|
|
|
|
rw [lineCube_smul]
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- On elements of `inQuadSol` a right-inverse to `inQuadToSol`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def inQuadProj (T : InQuadSol) : InQuad × ℚ × ℚ × ℚ :=
|
2024-04-22 06:57:46 -04:00
|
|
|
|
(⟨⟨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⟩,
|
2024-04-22 09:48:44 -04:00
|
|
|
|
(cubicCoeff T.val)⁻¹ * (cubeTriLin T.val.val T.val.val B₃.val),
|
|
|
|
|
(cubicCoeff T.val)⁻¹ * (- cubeTriLin T.val.val T.val.val Y₃.val),
|
2024-04-22 06:57:46 -04:00
|
|
|
|
(cubicCoeff T.val)⁻¹ *
|
2024-04-22 09:48:44 -04:00
|
|
|
|
(cubeTriLin T.val.val T.val.val B₃.val * (dot B₃.val T.val.val - dot Y₃.val T.val.val)
|
|
|
|
|
- cubeTriLin T.val.val T.val.val Y₃.val
|
2024-04-22 08:41:50 -04:00
|
|
|
|
* (dot Y₃.val T.val.val - 2 * dot B₃.val T.val.val)))
|
2024-04-22 06:57:46 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma inQuadToSol_proj (T : InQuadSol) : inQuadToSol (inQuadProj T) = T.val := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [inQuadProj, inQuadToSol_smul]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply ACCSystem.Sols.ext
|
2024-04-19 16:10:29 -04:00
|
|
|
|
change _ • (planeY₃B₃ _ _ _ _).val = _
|
2024-04-17 16:23:40 -04:00
|
|
|
|
rw [planeY₃B₃_val]
|
|
|
|
|
rw [Y₃_plus_B₃_plus_proj]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [cube_proj, cube_proj_proj_B₃, cube_proj_proj_Y₃]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
ring_nf
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp only [zero_smul, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add]
|
2024-04-22 09:48:44 -04:00
|
|
|
|
have h1 : (cubeTriLin T.val.val T.val.val Y₃.val ^ 2 * dot Y₃.val B₃.val ^ 3 * 3 +
|
|
|
|
|
dot Y₃.val B₃.val ^ 3 * cubeTriLin T.val.val T.val.val B₃.val ^ 2
|
2024-04-22 06:57:46 -04:00
|
|
|
|
* 3) = cubicCoeff T.val := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [cubicCoeff]
|
|
|
|
|
ring
|
|
|
|
|
rw [h1]
|
2024-04-22 06:57:46 -04:00
|
|
|
|
have h2 := (inCubeSolProp_iff_cubicCoeff_zero T.val).mpr.mt T.prop.2.2
|
|
|
|
|
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- Given a element of `inQuadCube × ℚ × ℚ × ℚ`, a solution to the ACCs. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def inQuadCubeToSol : InQuadCube × ℚ × ℚ × ℚ → MSSMACC.Sols := fun (R, b₁, b₂, b₃) =>
|
2024-04-19 16:10:29 -04:00
|
|
|
|
AnomalyFreeMk' (planeY₃B₃ R.val.val.val b₁ b₂ b₃)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
(by
|
|
|
|
|
rw [planeY₃B₃_quad]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [R.val.prop.1, R.val.prop.2.1, R.val.prop.2.2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp)
|
|
|
|
|
(by
|
|
|
|
|
rw [planeY₃B₃_cubic]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [R.prop.1, R.prop.2.1, R.prop.2.2]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma inQuadCubeToSol_smul (R : InQuadCube) (c₁ c₂ c₃ d : ℚ) :
|
2024-04-19 16:10:29 -04:00
|
|
|
|
inQuadCubeToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadCubeToSol (R, c₁, c₂, c₃):= by
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply ACCSystem.Sols.ext
|
|
|
|
|
change (planeY₃B₃ _ _ _ _).val = _
|
|
|
|
|
rw [planeY₃B₃_smul]
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- On elements of `inQuadCubeSol` a right-inverse to `inQuadCubeToSol`. -/
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def inQuadCubeProj (T : InQuadCubeSol) : InQuadCube × ℚ × ℚ × ℚ :=
|
2024-04-22 06:57:46 -04:00
|
|
|
|
(⟨⟨⟨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⟩,
|
2024-04-22 08:41:50 -04:00
|
|
|
|
(dot Y₃.val B₃.val)⁻¹ * (dot Y₃.val T.val.val - dot B₃.val T.val.val),
|
|
|
|
|
(dot Y₃.val B₃.val)⁻¹ * (2 * dot B₃.val T.val.val - dot Y₃.val T.val.val),
|
|
|
|
|
(dot Y₃.val B₃.val)⁻¹ * 1)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma inQuadCubeToSol_proj (T : InQuadCubeSol) :
|
2024-04-22 06:57:46 -04:00
|
|
|
|
inQuadCubeToSol (inQuadCubeProj T) = T.val := by
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [inQuadCubeProj, inQuadCubeToSol_smul]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply ACCSystem.Sols.ext
|
2024-04-19 16:10:29 -04:00
|
|
|
|
change _ • (planeY₃B₃ _ _ _ _).val = _
|
2024-04-17 16:23:40 -04:00
|
|
|
|
rw [planeY₃B₃_val]
|
|
|
|
|
rw [Y₃_plus_B₃_plus_proj]
|
|
|
|
|
ring_nf
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp only [Fin.isValue, Fin.reduceFinMk, zero_smul, add_zero, zero_add]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel]
|
|
|
|
|
simp only [one_smul]
|
2024-04-22 08:41:50 -04:00
|
|
|
|
rw [show dot Y₃.val B₃.val = 108 by rfl]
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-22 06:57:46 -04:00
|
|
|
|
/-- 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) =>
|
2024-06-26 11:54:02 -04:00
|
|
|
|
if h₃ : LineEqProp R ∧ InQuadProp R ∧ InCubeProp R then
|
2024-04-19 16:10:29 -04:00
|
|
|
|
inQuadCubeToSol (⟨⟨⟨R, h₃.1⟩, h₃.2.1⟩, h₃.2.2⟩, a, b, c)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
else
|
2024-06-26 11:54:02 -04:00
|
|
|
|
if h₂ : LineEqProp R ∧ InQuadProp R then
|
2024-04-19 16:10:29 -04:00
|
|
|
|
inQuadToSol (⟨⟨R, h₂.1⟩, h₂.2⟩, a, b, c)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
else
|
2024-06-26 11:54:02 -04:00
|
|
|
|
if h₁ : LineEqProp R then
|
2024-04-19 16:10:29 -04:00
|
|
|
|
inLineEqToSol (⟨R, h₁⟩, a, b, c)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
else
|
2024-04-19 16:10:29 -04:00
|
|
|
|
toSolNS ⟨R, a, b, c⟩
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma toSol_toSolNSProj (T : NotInLineEqSol) :
|
2024-04-22 06:57:46 -04:00
|
|
|
|
∃ X, toSol X = T.val := by
|
|
|
|
|
use toSolNSProj T.val
|
2024-06-26 11:54:02 -04:00
|
|
|
|
have h1 : ¬ LineEqProp (toSolNSProj T.val).1 :=
|
2024-04-22 06:57:46 -04:00
|
|
|
|
(linEqPropSol_iff_proj_linEqProp T.val).mpr.mt T.prop
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [toSol]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp_all
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact toSolNS_proj T
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma toSol_inLineEq (T : InLineEqSol) : ∃ X, toSol X = T.val := by
|
2024-04-22 06:57:46 -04:00
|
|
|
|
let X := inLineEqProj T
|
2024-04-19 16:10:29 -04:00
|
|
|
|
use ⟨X.1.val, X.2.1, X.2.2⟩
|
2024-06-26 11:54:02 -04:00
|
|
|
|
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
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [toSol]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp_all
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact inLineEqToSol_proj T
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma toSol_inQuad (T : InQuadSol) : ∃ X, toSol X = T.val := by
|
2024-04-22 06:57:46 -04:00
|
|
|
|
let X := inQuadProj T
|
2024-04-19 16:10:29 -04:00
|
|
|
|
use ⟨X.1.val.val, X.2.1, X.2.2⟩
|
2024-06-26 11:54:02 -04:00
|
|
|
|
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
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [toSol]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
simp_all
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact inQuadToSol_proj T
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma toSol_inQuadCube (T : InQuadCubeSol) : ∃ X, toSol X = T.val := by
|
2024-04-22 06:57:46 -04:00
|
|
|
|
let X := inQuadCubeProj T
|
2024-04-19 16:10:29 -04:00
|
|
|
|
use ⟨X.1.val.val.val, X.2.1, X.2.2⟩
|
2024-06-26 11:54:02 -04:00
|
|
|
|
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
|
2024-04-19 16:10:29 -04:00
|
|
|
|
rw [toSol]
|
|
|
|
|
simp_all
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact inQuadCubeToSol_proj T
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-04-19 16:10:29 -04:00
|
|
|
|
theorem toSol_surjective : Function.Surjective toSol := by
|
2024-04-17 16:23:40 -04:00
|
|
|
|
intro T
|
2024-06-26 11:54:02 -04:00
|
|
|
|
by_cases h₁ : ¬ LineEqPropSol T
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact toSol_toSolNSProj ⟨T, h₁⟩
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp at h₁
|
2024-06-26 11:54:02 -04:00
|
|
|
|
by_cases h₂ : ¬ InQuadSolProp T
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact toSol_inLineEq ⟨T, And.intro h₁ h₂⟩
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp at h₂
|
2024-06-26 11:54:02 -04:00
|
|
|
|
by_cases h₃ : ¬ InCubeSolProp T
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact toSol_inQuad ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩
|
2024-04-19 16:10:29 -04:00
|
|
|
|
simp at h₃
|
2024-04-22 06:57:46 -04:00
|
|
|
|
exact toSol_inQuadCube ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩
|
2024-04-19 16:10:29 -04:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
end AnomalyFreePerp
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
|
|
|
|
end MSSMACC
|