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

@ -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₃]