refactor: Quad Lin equations

This commit is contained in:
jstoobysmith 2024-04-22 08:41:50 -04:00
parent 772e78ca77
commit b5dd319eed
14 changed files with 245 additions and 299 deletions

View file

@ -270,13 +270,13 @@ lemma accYY_ext {S T : MSSMCharges.charges}
rw [hd, hu]
rfl
/-- The symmetric bilinear form used to define the quadratic ACC. -/
@[simps!]
def quadBiLin : BiLinearSymm MSSMCharges.charges where
toFun S := ∑ i, (Q S.1 i * Q S.2 i + (- 2) * (U S.1 i * U S.2 i) +
def quadBiLin : BiLinearSymm MSSMCharges.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 + (- 1) * (L S.1 i * L S.2 i) + E S.1 i * E S.2 i) +
(- Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2)
map_smul₁' a S T := by
(- Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2))
(by
intro a S T
simp only
rw [mul_add]
congr 1
@ -287,8 +287,9 @@ def quadBiLin : BiLinearSymm MSSMCharges.charges where
simp only [HSMul.hSMul, SMul.smul, toSMSpecies_apply, Fin.isValue, neg_mul, one_mul]
ring
simp only [map_smul, Hd_apply, Fin.reduceFinMk, Fin.isValue, smul_eq_mul, neg_mul, Hu_apply]
ring
map_add₁' S T R := by
ring)
(by
intro S T R
simp only
rw [add_assoc, ← add_assoc (-Hd S * Hd R + Hu S * Hu R) _ _]
rw [add_comm (-Hd S * Hd R + Hu S * Hu R) _]
@ -303,15 +304,17 @@ def quadBiLin : BiLinearSymm MSSMCharges.charges where
one_mul]
ring
rw [Hd.map_add, Hu.map_add]
ring
swap' S L := by
ring)
(by
intro S L
simp only [MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, neg_mul, one_mul,
Hd_apply, Fin.reduceFinMk, Hu_apply]
congr 1
rw [Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
ring
ring
ring)
/-- The quadratic ACC. -/
@[simp]
@ -323,10 +326,9 @@ lemma accQuad_ext {S T : (MSSMCharges).charges}
∑ i, ((fun a => a^2) ∘ toSMSpecies j T) i)
(hd : Hd S = Hd T) (hu : Hu S = Hu T) :
accQuad S = accQuad T := by
simp only [accQuad, BiLinearSymm.toHomogeneousQuad_toFun]
simp only [HomogeneousQuadratic, accQuad, BiLinearSymm.toHomogeneousQuad_apply]
erw [← quadBiLin.toFun_eq_coe]
rw [quadBiLin]
simp only
simp only [quadBiLin, BiLinearSymm.mk₂, AddHom.toFun_eq_coe, AddHom.coe_mk, LinearMap.coe_mk]
repeat erw [Finset.sum_add_distrib]
repeat erw [← Finset.mul_sum]
ring_nf
@ -337,6 +339,7 @@ lemma accQuad_ext {S T : (MSSMCharges).charges}
repeat rw [h1]
rw [hd, hu]
/-- The function underlying the symmetric trilinear form used to define the cubic ACC. -/
@[simp]
def cubeTriLinToFun
@ -461,7 +464,7 @@ open MSSMCharges
lemma quadSol (S : MSSMACC.QuadSols) : accQuad S.val = 0 := by
have hS := S.quadSol
simp [MSSMACCs.accQuad, HomogeneousQuadratic.toFun] at hS
simp only [MSSMACC_numberQuadratic, HomogeneousQuadratic, MSSMACC_quadraticACCs] at hS
exact hS 0
/-- A solution from a charge satisfying the ACCs. -/
@ -523,30 +526,33 @@ lemma AnomalyFreeMk''_val (S : MSSMACC.QuadSols)
/-- The dot product on the vector space of charges. -/
@[simps!]
def dot : BiLinearSymm MSSMCharges.charges where
toFun S := ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i +
def dot : BiLinearSymm MSSMCharges.charges := BiLinearSymm.mk₂
(fun S => ∑ i, (Q S.1 i * Q S.2 i + U S.1 i * U S.2 i +
D S.1 i * D S.2 i + L S.1 i * L S.2 i + E S.1 i * E S.2 i
+ N S.1 i * N S.2 i) + Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2
map_smul₁' a S T := by
+ N S.1 i * N S.2 i) + Hd S.1 * Hd S.2 + Hu S.1 * Hu S.2)
(by
intro a S T
simp only [MSSMSpecies_numberCharges]
repeat rw [(toSMSpecies _).map_smul]
rw [Hd.map_smul, Hu.map_smul]
rw [Fin.sum_univ_three, Fin.sum_univ_three]
simp only [HSMul.hSMul, SMul.smul, Fin.isValue, toSMSpecies_apply, Hd_apply, Fin.reduceFinMk,
Hu_apply]
ring
map_add₁' S T R := by
ring)
(by
intro S1 S2 T
simp only [MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue,
ACCSystemCharges.chargesAddCommMonoid_add, map_add, Hd_apply, Fin.reduceFinMk, Hu_apply]
repeat erw [AddHom.map_add]
rw [Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
ring
swap' S L := by
ring)
(by
intro S T
simp only [MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply, Fin.reduceFinMk,
Hu_apply]
rw [Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
ring
ring)
end MSSMACC

View file

@ -40,7 +40,7 @@ lemma lineY₃B₃Charges_quad (a b : ) : accQuad (lineY₃B₃Charges a b).v
rw [quadBiLin.toHomogeneousQuad.map_smul]
rw [quadBiLin.map_smul₁, quadBiLin.map_smul₂]
erw [quadSol Y₃.1, quadSol B₃.1]
simp only [mul_zero, add_zero, quadBiLin_toFun, Fin.isValue, Fin.reduceFinMk, zero_add,
simp only [mul_zero, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add,
mul_eq_zero, OfNat.ofNat_ne_zero, false_or]
apply Or.inr ∘ Or.inr
rfl

View file

@ -29,42 +29,42 @@ open BigOperators
/-- The type of linear solutions orthogonal to $Y_3$ and $B_3$. -/
structure AnomalyFreePerp extends MSSMACC.LinSols where
perpY₃ : dot (Y₃.val, val) = 0
perpB₃ : dot (B₃.val, val) = 0
perpY₃ : dot Y₃.val val = 0
perpB₃ : dot B₃.val val = 0
/-- The projection of an object in `MSSMACC.AnomalyFreeLinear` onto the subspace
orthgonal to `Y₃` and`B₃`. -/
def proj (T : MSSMACC.LinSols) : MSSMACC.AnomalyFreePerp :=
⟨(dot (B₃.val, T.val) - dot (Y₃.val, T.val)) • Y₃.1.1
+ (dot (Y₃.val, T.val) - 2 * dot (B₃.val, T.val)) • B₃.1.1
+ dot (Y₃.val, B₃.val) • T,
⟨(dot B₃.val T.val - dot Y₃.val T.val) • Y₃.1.1
+ (dot Y₃.val T.val - 2 * dot B₃.val T.val) • B₃.1.1
+ dot Y₃.val B₃.val • T,
by
change dot (_, _ • Y₃.val + _ • B₃.val + _ • T.val) = 0
change dot _ (_ • Y₃.val + _ • B₃.val + _ • T.val) = 0
rw [dot.map_add₂, dot.map_add₂]
rw [dot.map_smul₂, dot.map_smul₂, dot.map_smul₂]
rw [show dot (Y₃.val, B₃.val) = 108 by rfl]
rw [show dot (Y₃.val, Y₃.val) = 216 by rfl]
rw [show dot Y₃.val B₃.val = 108 by rfl]
rw [show dot Y₃.val Y₃.val = 216 by rfl]
ring,
by
change dot (_, _ • Y₃.val + _ • B₃.val + _ • T.val) = 0
change dot _ (_ • Y₃.val + _ • B₃.val + _ • T.val) = 0
rw [dot.map_add₂, dot.map_add₂]
rw [dot.map_smul₂, dot.map_smul₂, dot.map_smul₂]
rw [show dot (Y₃.val, B₃.val) = 108 by rfl]
rw [show dot (B₃.val, Y₃.val) = 108 by rfl]
rw [show dot (B₃.val, B₃.val) = 108 by rfl]
rw [show dot Y₃.val B₃.val = 108 by rfl]
rw [show dot B₃.val Y₃.val = 108 by rfl]
rw [show dot B₃.val B₃.val = 108 by rfl]
ring⟩
lemma proj_val (T : MSSMACC.LinSols) :
(proj T).val = (dot (B₃.val, T.val) - dot (Y₃.val, T.val)) • Y₃.val +
( (dot (Y₃.val, T.val) - 2 * dot (B₃.val, T.val))) • B₃.val +
dot (Y₃.val, B₃.val) • T.val := by
(proj T).val = (dot B₃.val T.val - dot Y₃.val T.val) • Y₃.val +
( (dot Y₃.val T.val - 2 * dot B₃.val T.val)) • B₃.val +
dot Y₃.val B₃.val • T.val := by
rfl
lemma Y₃_plus_B₃_plus_proj (T : MSSMACC.LinSols) (a b c : ) :
a • Y₃.val + b • B₃.val + c • (proj T).val =
(a + c * (dot (B₃.val, T.val) - dot (Y₃.val, T.val))) • Y₃.val
+ (b + c * (dot (Y₃.val, T.val) - 2 * dot (B₃.val, T.val))) • B₃.val
+ (dot (Y₃.val, B₃.val) * c) • T.val:= by
(a + c * (dot B₃.val T.val - dot Y₃.val T.val)) • Y₃.val
+ (b + c * (dot Y₃.val T.val - 2 * dot B₃.val T.val)) • B₃.val
+ (dot Y₃.val B₃.val * c) • T.val:= by
rw [proj_val]
rw [DistribMulAction.smul_add, DistribMulAction.smul_add]
rw [add_assoc (_ • _ • Y₃.val), ← add_assoc (_ • Y₃.val + _ • B₃.val), add_assoc (_ • Y₃.val)]
@ -81,27 +81,27 @@ lemma Y₃_plus_B₃_plus_proj (T : MSSMACC.LinSols) (a b c : ) :
lemma quad_Y₃_proj (T : MSSMACC.LinSols) :
quadBiLin (Y₃.val, (proj T).val) = dot (Y₃.val, B₃.val) * quadBiLin (Y₃.val, T.val) := by
quadBiLin Y₃.val (proj T).val = dot Y₃.val B₃.val * quadBiLin Y₃.val T.val := by
rw [proj_val]
rw [quadBiLin.map_add₂, quadBiLin.map_add₂]
rw [quadBiLin.map_smul₂, quadBiLin.map_smul₂, quadBiLin.map_smul₂]
rw [show quadBiLin (Y₃.val, B₃.val) = 0 by rfl]
rw [show quadBiLin (Y₃.val, Y₃.val) = 0 by rfl]
rw [show quadBiLin Y₃.val B₃.val = 0 by rfl]
rw [show quadBiLin Y₃.val Y₃.val = 0 by rfl]
ring
lemma quad_B₃_proj (T : MSSMACC.LinSols) :
quadBiLin (B₃.val, (proj T).val) = dot (Y₃.val, B₃.val) * quadBiLin (B₃.val, T.val) := by
quadBiLin B₃.val (proj T).val = dot Y₃.val B₃.val * quadBiLin B₃.val T.val := by
rw [proj_val]
rw [quadBiLin.map_add₂, quadBiLin.map_add₂]
rw [quadBiLin.map_smul₂, quadBiLin.map_smul₂, quadBiLin.map_smul₂]
rw [show quadBiLin (B₃.val, Y₃.val) = 0 by rfl]
rw [show quadBiLin (B₃.val, B₃.val) = 0 by rfl]
rw [show quadBiLin B₃.val Y₃.val = 0 by rfl]
rw [show quadBiLin B₃.val B₃.val = 0 by rfl]
ring
lemma quad_self_proj (T : MSSMACC.Sols) :
quadBiLin (T.val, (proj T.1.1).val) =
(dot (B₃.val, T.val) - dot (Y₃.val, T.val)) * quadBiLin (Y₃.val, T.val) +
(dot (Y₃.val, T.val) - 2 * dot (B₃.val, T.val)) * quadBiLin (B₃.val, T.val) := by
quadBiLin T.val (proj T.1.1).val =
(dot B₃.val T.val - dot Y₃.val T.val) * quadBiLin Y₃.val T.val +
(dot Y₃.val T.val - 2 * dot B₃.val T.val) * quadBiLin B₃.val T.val := by
rw [proj_val]
rw [quadBiLin.map_add₂, quadBiLin.map_add₂]
rw [quadBiLin.map_smul₂, quadBiLin.map_smul₂, quadBiLin.map_smul₂]
@ -110,9 +110,9 @@ lemma quad_self_proj (T : MSSMACC.Sols) :
ring
lemma quad_proj (T : MSSMACC.Sols) :
quadBiLin ((proj T.1.1).val, (proj T.1.1).val) = 2 * (dot (Y₃.val, B₃.val)) *
((dot (B₃.val, T.val) - dot (Y₃.val, T.val)) * quadBiLin (Y₃.val, T.val) +
(dot (Y₃.val, T.val) - 2 * dot (B₃.val, T.val)) * quadBiLin (B₃.val, T.val) ) := by
quadBiLin (proj T.1.1).val (proj T.1.1).val = 2 * dot Y₃.val B₃.val *
((dot B₃.val T.val - dot Y₃.val T.val) * quadBiLin Y₃.val T.val +
(dot Y₃.val T.val - 2 * dot B₃.val T.val) * quadBiLin B₃.val T.val ) := by
nth_rewrite 1 [proj_val]
repeat rw [quadBiLin.map_add₁]
repeat rw [quadBiLin.map_smul₁]
@ -122,7 +122,7 @@ lemma quad_proj (T : MSSMACC.Sols) :
lemma cube_proj_proj_Y₃ (T : MSSMACC.LinSols) :
cubeTriLin ((proj T).val, (proj T).val, Y₃.val) =
(dot (Y₃.val, B₃.val))^2 * cubeTriLin (T.val, T.val, Y₃.val):= by
(dot Y₃.val B₃.val)^2 * cubeTriLin (T.val, T.val, Y₃.val) := by
rw [proj_val]
rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂]
erw [lineY₃B₃_doublePoint]
@ -144,7 +144,7 @@ lemma cube_proj_proj_Y₃ (T : MSSMACC.LinSols) :
lemma cube_proj_proj_B₃ (T : MSSMACC.LinSols) :
cubeTriLin ((proj T).val, (proj T).val, B₃.val) =
(dot (Y₃.val, B₃.val))^2 * cubeTriLin (T.val, T.val, B₃.val):= by
(dot Y₃.val B₃.val)^2 * cubeTriLin (T.val, T.val, B₃.val) := by
rw [proj_val]
rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂]
erw [lineY₃B₃_doublePoint]
@ -160,9 +160,9 @@ lemma cube_proj_proj_B₃ (T : MSSMACC.LinSols) :
lemma cube_proj_proj_self (T : MSSMACC.Sols) :
cubeTriLin ((proj T.1.1).val, (proj T.1.1).val, T.val) =
2 * dot (Y₃.val, B₃.val) *
((dot (B₃.val, T.val) - dot (Y₃.val, T.val)) * cubeTriLin (T.val, T.val, Y₃.val) +
( dot (Y₃.val, T.val)- 2 * dot (B₃.val, T.val)) * cubeTriLin (T.val, T.val, B₃.val)) := by
2 * dot Y₃.val B₃.val *
((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin (T.val, T.val, Y₃.val) +
( dot Y₃.val T.val- 2 * dot B₃.val T.val) * cubeTriLin (T.val, T.val, B₃.val)) := by
rw [proj_val]
rw [cubeTriLin.map_add₁, cubeTriLin.map_add₂]
erw [lineY₃B₃_doublePoint]
@ -177,9 +177,9 @@ lemma cube_proj_proj_self (T : MSSMACC.Sols) :
lemma cube_proj (T : MSSMACC.Sols) :
cubeTriLin ((proj T.1.1).val, (proj T.1.1).val, (proj T.1.1).val) =
3 * dot (Y₃.val, B₃.val) ^ 2 *
((dot (B₃.val, T.val) - dot (Y₃.val, T.val)) * cubeTriLin (T.val, T.val, Y₃.val) +
(dot (Y₃.val, T.val) - 2 * dot (B₃.val, T.val)) * cubeTriLin (T.val, T.val, B₃.val)) := by
3 * dot Y₃.val B₃.val ^ 2 *
((dot B₃.val T.val - dot Y₃.val T.val) * cubeTriLin (T.val, T.val, Y₃.val) +
(dot Y₃.val T.val - 2 * dot B₃.val T.val) * cubeTriLin (T.val, T.val, B₃.val)) := by
nth_rewrite 3 [proj_val]
repeat rw [cubeTriLin.map_add₃]
repeat rw [cubeTriLin.map_smul₃]

View file

@ -53,8 +53,8 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ) (hR' : R
(h : (planeY₃B₃ R a b c).val = (planeY₃B₃ R a' b' c').val) :
a = a' ∧ b = b' ∧ c = c' := by
rw [planeY₃B₃_val, planeY₃B₃_val] at h
have h1 := congrArg (fun S => dot (Y₃.val, S)) h
have h2 := congrArg (fun S => dot (B₃.val, S)) h
have h1 := congrArg (fun S => dot Y₃.val S) h
have h2 := congrArg (fun S => dot B₃.val S) h
simp only [ Fin.isValue, ACCSystemCharges.chargesAddCommMonoid_add, Fin.reduceFinMk] at h1 h2
erw [dot.map_add₂, dot.map_add₂] at h1 h2
erw [dot.map_add₂ Y₃.val (a' • Y₃.val + b' • B₃.val) (c' • R.val)] at h1
@ -64,10 +64,10 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ) (hR' : R
rw [dot.map_smul₂, dot.map_smul₂, dot.map_smul₂] at h1 h2
rw [R.perpY₃] at h1
rw [R.perpB₃] at h2
rw [show dot (Y₃.val, Y₃.val) = 216 by rfl] at h1
rw [show dot (B₃.val, B₃.val) = 108 by rfl] at h2
rw [show dot (Y₃.val, B₃.val) = 108 by rfl] at h1
rw [show dot (B₃.val, Y₃.val) = 108 by rfl] at h2
rw [show dot Y₃.val Y₃.val = 216 by rfl] at h1
rw [show dot B₃.val B₃.val = 108 by rfl] at h2
rw [show dot Y₃.val B₃.val = 108 by rfl] at h1
rw [show dot B₃.val Y₃.val = 108 by rfl] at h2
simp_all
have ha : a = a' := by
linear_combination h1 / 108 + -1 * h2 / 108
@ -102,15 +102,15 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ) (hR' : R
lemma planeY₃B₃_quad (R : MSSMACC.AnomalyFreePerp) (a b c : ) :
accQuad (planeY₃B₃ R a b c).val = c * (2 * a * quadBiLin (Y₃.val, R.val)
+ 2 * b * quadBiLin (B₃.val, R.val) + c * quadBiLin (R.val, R.val)) := by
accQuad (planeY₃B₃ R a b c).val = c * (2 * a * quadBiLin Y₃.val R.val
+ 2 * b * quadBiLin B₃.val R.val + c * quadBiLin R.val R.val) := by
rw [planeY₃B₃_val]
erw [BiLinearSymm.toHomogeneousQuad_add]
erw [lineY₃B₃Charges_quad]
rw [quadBiLin.toHomogeneousQuad.map_smul]
rw [quadBiLin.map_add₁, quadBiLin.map_smul₁, quadBiLin.map_smul₁]
rw [quadBiLin.map_smul₂, quadBiLin.map_smul₂]
rw [show (BiLinearSymm.toHomogeneousQuad quadBiLin) R.val = quadBiLin (R.val, R.val) by rfl]
rw [show (BiLinearSymm.toHomogeneousQuad quadBiLin) R.val = quadBiLin R.val R.val by rfl]
ring
lemma planeY₃B₃_cubic (R : MSSMACC.AnomalyFreePerp) (a b c : ) :
@ -130,9 +130,9 @@ lemma planeY₃B₃_cubic (R : MSSMACC.AnomalyFreePerp) (a b c : ) :
/-- The line in the plane spanned by $Y_3$, $B_3$ and $R$ which is in the quadratic,
as `LinSols`. -/
def lineQuadAFL (R : MSSMACC.AnomalyFreePerp) (c1 c2 c3 : ) : MSSMACC.LinSols :=
planeY₃B₃ R (c2 * quadBiLin (R.val, R.val) - 2 * c3 * quadBiLin (B₃.val, R.val))
(2 * c3 * quadBiLin (Y₃.val, R.val) - c1 * quadBiLin (R.val, R.val))
(2 * c1 * quadBiLin (B₃.val, R.val) - 2 * c2 * quadBiLin (Y₃.val, R.val))
planeY₃B₃ R (c2 * quadBiLin R.val R.val - 2 * c3 * quadBiLin B₃.val R.val)
(2 * c3 * quadBiLin Y₃.val R.val - c1 * quadBiLin R.val R.val)
(2 * c1 * quadBiLin B₃.val R.val - 2 * c2 * quadBiLin Y₃.val R.val)
lemma lineQuadAFL_quad (R : MSSMACC.AnomalyFreePerp) (c1 c2 c3 : ) :
accQuad (lineQuadAFL R c1 c2 c3).val = 0 := by
@ -146,10 +146,10 @@ def lineQuad (R : MSSMACC.AnomalyFreePerp) (c1 c2 c3 : ) : MSSMACC.QuadSols :
AnomalyFreeQuadMk' (lineQuadAFL R c1 c2 c3) (lineQuadAFL_quad R c1 c2 c3)
lemma lineQuad_val (R : MSSMACC.AnomalyFreePerp) (c1 c2 c3 : ) :
(lineQuad R c1 c2 c3).val = (planeY₃B₃ R
(c2 * quadBiLin (R.val, R.val) - 2 * c3 * quadBiLin (B₃.val, R.val))
(2 * c3 * quadBiLin (Y₃.val, R.val) - c1 * quadBiLin (R.val, R.val))
(2 * c1 * quadBiLin (B₃.val, R.val) - 2 * c2 * quadBiLin (Y₃.val, R.val))).val := by
(lineQuad R c1 c2 c3).val = (planeY₃B₃ R
(c2 * quadBiLin R.val R.val - 2 * c3 * quadBiLin B₃.val R.val)
(2 * c3 * quadBiLin Y₃.val R.val - c1 * quadBiLin R.val R.val)
(2 * c1 * quadBiLin B₃.val R.val - 2 * c2 * quadBiLin Y₃.val R.val)).val := by
rfl
lemma lineQuad_smul (R : MSSMACC.AnomalyFreePerp) (a b c d : ) :
@ -163,26 +163,26 @@ lemma lineQuad_smul (R : MSSMACC.AnomalyFreePerp) (a b c d : ) :
/-- A helper function to simplify following expressions. -/
def α₁ (T : MSSMACC.AnomalyFreePerp) : :=
(3 * cubeTriLin (T.val, T.val, B₃.val) * quadBiLin (T.val, T.val) -
2 * cubeTriLin (T.val, T.val, T.val) * quadBiLin (B₃.val, T.val))
(3 * cubeTriLin (T.val, T.val, B₃.val) * quadBiLin T.val T.val -
2 * cubeTriLin (T.val, T.val, T.val) * quadBiLin B₃.val T.val)
/-- A helper function to simplify following expressions. -/
def α₂ (T : MSSMACC.AnomalyFreePerp) : :=
(2 * cubeTriLin (T.val, T.val, T.val) * quadBiLin (Y₃.val, T.val) -
3 * cubeTriLin (T.val, T.val, Y₃.val) * quadBiLin (T.val, T.val))
/-- A helper function to simplify following expressions. -/
def α₂ (T : MSSMACC.AnomalyFreePerp) : :=
(2 * cubeTriLin (T.val, T.val, T.val) * quadBiLin Y₃.val T.val -
3 * cubeTriLin (T.val, T.val, Y₃.val) * quadBiLin T.val T.val)
/-- A helper function to simplify following expressions. -/
def α₃ (T : MSSMACC.AnomalyFreePerp) : :=
6 * ((cubeTriLin (T.val, T.val, Y₃.val)) * quadBiLin (B₃.val, T.val) -
(cubeTriLin (T.val, T.val, B₃.val)) * quadBiLin (Y₃.val, T.val))
/-- A helper function to simplify following expressions. -/
def α₃ (T : MSSMACC.AnomalyFreePerp) : :=
6 * ((cubeTriLin (T.val, T.val, Y₃.val)) * quadBiLin B₃.val T.val -
(cubeTriLin (T.val, T.val, B₃.val)) * quadBiLin Y₃.val T.val)
lemma lineQuad_cube (R : MSSMACC.AnomalyFreePerp) (c₁ c₂ c₃ : ) :
accCube (lineQuad R c₁ c₂ c₃).val =
- 4 * ( c₁ * quadBiLin (B₃.val, R.val) - c₂ * quadBiLin (Y₃.val, R.val)) ^ 2 *
( α₁ R * c₁ + α₂ R * c₂ + α₃ R * c₃ ) := by
rw [lineQuad_val]
rw [planeY₃B₃_cubic, α₁, α₂, α₃]
ring
lemma lineQuad_cube (R : MSSMACC.AnomalyFreePerp) (c₁ c₂ c₃ : ) :
accCube (lineQuad R c₁ c₂ c₃).val =
- 4 * ( c₁ * quadBiLin B₃.val R.val - c₂ * quadBiLin Y₃.val R.val) ^ 2 *
( α₁ R * c₁ + α₂ R * c₂ + α₃ R * c₃ ) := by
rw [lineQuad_val]
rw [planeY₃B₃_cubic, α₁, α₂, α₃]
ring
/-- The line in the plane spanned by $Y_3$, $B_3$ and $R$ which is in the cubic. -/
def lineCube (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ) :
@ -220,21 +220,21 @@ lemma lineCube_quad (R : MSSMACC.AnomalyFreePerp) (a₁ a₂ a₃ : ) :
section proj
lemma α₃_proj (T : MSSMACC.Sols) : α₃ (proj T.1.1) =
6 * dot (Y₃.val, B₃.val) ^ 3 * (
cubeTriLin (T.val, T.val, Y₃.val) * quadBiLin (B₃.val, T.val) -
cubeTriLin (T.val, T.val, B₃.val) * quadBiLin (Y₃.val, T.val)) := by
6 * dot Y₃.val B₃.val ^ 3 * (
cubeTriLin (T.val, T.val, Y₃.val) * quadBiLin B₃.val T.val -
cubeTriLin (T.val, T.val, B₃.val) * quadBiLin Y₃.val T.val) := by
rw [α₃]
rw [cube_proj_proj_Y₃, cube_proj_proj_B₃, quad_B₃_proj, quad_Y₃_proj]
ring
lemma α₂_proj (T : MSSMACC.Sols) : α₂ (proj T.1.1) =
- α₃ (proj T.1.1) * (dot (Y₃.val, T.val) - 2 * dot (B₃.val, T.val)) := by
- α₃ (proj T.1.1) * (dot Y₃.val T.val - 2 * dot B₃.val T.val) := by
rw [α₃_proj, α₂]
rw [cube_proj_proj_Y₃, quad_Y₃_proj, quad_proj, cube_proj]
ring
lemma α₁_proj (T : MSSMACC.Sols) : α₁ (proj T.1.1) =
- α₃ (proj T.1.1) * (dot (B₃.val, T.val) - dot (Y₃.val, T.val)) := by
- α₃ (proj T.1.1) * (dot B₃.val T.val - dot Y₃.val T.val) := by
rw [α₃_proj, α₁]
rw [cube_proj_proj_B₃, quad_B₃_proj, quad_proj, cube_proj]
ring

View file

@ -46,12 +46,12 @@ instance (R : MSSMACC.AnomalyFreePerp) : Decidable (lineEqProp R) := by
/-- 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 :=
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
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
/-- A rational which appears in `toSolNS` acting on sols, and which been zero is
equivalent to satisfying `lineEqPropSol`. -/
def lineEqCoeff (T : MSSMACC.Sols) : := dot (Y₃.val, B₃.val) * α₃ (proj T.1.1)
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
@ -59,7 +59,7 @@ lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) :
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]
rw [show dot (Y₃.val, B₃.val) = 108 by rfl]
rw [show dot Y₃.val B₃.val = 108 by rfl]
simp only [Fin.isValue, Fin.reduceFinMk, OfNat.ofNat_ne_zero, false_or]
ring_nf
rw [mul_comm _ 1259712, mul_comm _ 1259712, ← mul_sub]
@ -70,10 +70,10 @@ lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) :
rw [lineEqPropSol_iff_lineEqCoeff_zero, lineEqCoeff, lineEqProp]
apply Iff.intro
intro h
rw [show dot (Y₃.val, B₃.val) = 108 by rfl] at h
rw [show dot Y₃.val B₃.val = 108 by rfl] at h
simp at h
rw [α₁_proj, α₂_proj, h]
simp only [neg_zero, dot_toFun, Fin.isValue, Fin.reduceFinMk, zero_mul, and_self]
simp only [neg_zero, Fin.isValue, Fin.reduceFinMk, zero_mul, and_self]
intro h
rw [h.2.2]
simp
@ -82,7 +82,7 @@ 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 :=
quadBiLin (R.val, R.val) = 0 ∧ quadBiLin (Y₃.val, R.val) = 0 ∧ quadBiLin (B₃.val, R.val) = 0
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
apply And.decidable
@ -90,23 +90,23 @@ instance (R : MSSMACC.AnomalyFreePerp) : Decidable (inQuadProp R) := by
/-- 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 :=
quadBiLin (Y₃.val, R.val) = 0 ∧ quadBiLin (B₃.val, R.val) = 0
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
that solution satisfies `inQuadSolProp`. It appears in the definition of `inQuadProj`. -/
def quadCoeff (T : MSSMACC.Sols) : :=
2 * dot (Y₃.val, B₃.val) ^ 2 *
(quadBiLin (Y₃.val, T.val) ^ 2 + quadBiLin (B₃.val, T.val) ^ 2)
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
apply Iff.intro
intro h
rw [quadCoeff, h.1, h.2]
simp only [dot_toFun, Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
zero_pow, add_zero, mul_zero]
intro h
rw [quadCoeff] at h
rw [show dot (Y₃.val, B₃.val) = 108 by rfl] at h
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,
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
@ -123,10 +123,10 @@ lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
apply Iff.intro
intro h
rw [h.1, h.2]
simp only [dot_toFun, Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
simp only [Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
intro h
rw [show dot (Y₃.val, B₃.val) = 108 by rfl] at h
simp only [dot_toFun, Fin.isValue, Fin.reduceFinMk , mul_eq_zero,
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, or_self, false_or] at h
rw [h.2.1, h.2.2]
simp
@ -149,7 +149,7 @@ def inCubeSolProp (R : MSSMACC.Sols) : Prop :=
/-- A rational which has two properties. It is zero for a solution `T` if and only if
that solution satisfies `inCubeSolProp`. It appears in the definition of `inLineEqProj`. -/
def cubicCoeff (T : MSSMACC.Sols) : :=
3 * dot (Y₃.val, B₃.val) ^ 3 * (cubeTriLin (T.val, T.val, Y₃.val) ^ 2 +
3 * (dot Y₃.val B₃.val) ^ 3 * (cubeTriLin (T.val, T.val, Y₃.val) ^ 2 +
cubeTriLin (T.val, T.val, B₃.val) ^ 2 )
lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
@ -157,11 +157,11 @@ lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
apply Iff.intro
intro h
rw [cubicCoeff, h.1, h.2]
simp only [dot_toFun, Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true,
zero_pow, add_zero, mul_zero]
intro h
rw [cubicCoeff] at h
rw [show dot (Y₃.val, B₃.val) = 108 by rfl] at h
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,
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
@ -176,10 +176,10 @@ lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) :
apply Iff.intro
intro h
rw [h.1, h.2]
simp only [dot_toFun, Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
simp only [Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
intro h
rw [show dot (Y₃.val, B₃.val) = 108 by rfl] at h
simp only [dot_toFun, Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
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,
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
rw [h.2.1, h.2.2]
simp
@ -254,7 +254,7 @@ lemma toSolNS_proj (T : notInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val :=
rw [α₁_proj, α₂_proj]
ring_nf
simp only [zero_smul, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add]
have h1 : α₃ (proj T.val.toLinSols) * dot (Y₃.val, B₃.val) = lineEqCoeff T.val := by
have h1 : α₃ (proj T.val.toLinSols) * dot Y₃.val B₃.val = lineEqCoeff T.val := by
rw [lineEqCoeff]
ring
rw [h1]
@ -273,12 +273,11 @@ def inLineEqToSol : inLineEq × × × → MSSMACC.Sols := fun (R, c
/-- On elements of `inLineEqSol` a right-inverse to `inLineEqSol`. -/
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)),
(quadCoeff T.val)⁻¹ * quadBiLin B₃.val T.val.val,
(quadCoeff T.val)⁻¹ * (- quadBiLin Y₃.val T.val.val),
(quadCoeff T.val)⁻¹ * (
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))))
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 : ) :
inLineEqToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inLineEqToSol (R, c₁, c₂, c₃) := by
@ -297,8 +296,8 @@ lemma inLineEqToSol_proj (T : inLineEqSol) : inLineEqToSol (inLineEqProj T) = T.
rw [quad_proj, quad_Y₃_proj, quad_B₃_proj]
ring_nf
simp only [zero_smul, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add]
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
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
rw [quadCoeff]
ring
rw [h1]
@ -329,9 +328,9 @@ def inQuadProj (T : inQuadSol) : inQuad × × × :=
(cubicCoeff T.val)⁻¹ * (cubeTriLin (T.val.val, T.val.val, B₃.val)),
(cubicCoeff T.val)⁻¹ * (- cubeTriLin (T.val.val, T.val.val, Y₃.val)),
(cubicCoeff T.val)⁻¹ *
(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, B₃.val) * (dot B₃.val T.val.val - dot Y₃.val T.val.val)
- cubeTriLin (T.val.val, T.val.val, Y₃.val)
* (dot (Y₃.val, T.val.val) - 2 * dot (B₃.val, T.val.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
@ -343,8 +342,8 @@ lemma inQuadToSol_proj (T : inQuadSol) : inQuadToSol (inQuadProj T) = T.val := b
rw [cube_proj, cube_proj_proj_B₃, cube_proj_proj_Y₃]
ring_nf
simp only [zero_smul, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add]
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
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
* 3) = cubicCoeff T.val := by
rw [cubicCoeff]
ring
@ -378,9 +377,9 @@ 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⟩,
(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)
(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)
lemma inQuadCubeToSol_proj (T : inQuadCubeSol) :
inQuadCubeToSol (inQuadCubeProj T) = T.val := by
@ -393,7 +392,7 @@ lemma inQuadCubeToSol_proj (T : inQuadCubeSol) :
simp only [Fin.isValue, Fin.reduceFinMk, zero_smul, add_zero, zero_add]
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel]
simp only [one_smul]
rw [show dot (Y₃.val, B₃.val) = 108 by rfl]
rw [show dot Y₃.val B₃.val = 108 by rfl]
simp
/-- Given an element of `MSSMACC.AnomalyFreePerp × × × ` a solution. We will