refactor: def of symmetric trilin function

This commit is contained in:
jstoobysmith 2024-04-22 09:48:44 -04:00
parent 748bcb61ae
commit e36c61b331
24 changed files with 279 additions and 246 deletions

View file

@ -471,7 +471,7 @@ lemma P!_accCube (f : Fin n → ) : accCube (2 * n +1) (P! f) = 0 := by
ring
lemma P_P_P!_accCube (g : Fin n → ) (j : Fin n) :
accCubeTriLinSymm.toFun (P g, P g, basis!AsCharges j)
accCubeTriLinSymm (P g) (P g) (basis!AsCharges j)
= (P g (δ!₁ j))^2 - (g j)^2 := by
simp [accCubeTriLinSymm]
rw [sum_δ!, basis!_on_δ!₃]

View file

@ -42,8 +42,8 @@ def lineInCubic (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S) :
∀ (g : Fin n → ) (f : Fin n → ) (_ : S.val = P g + P! f) (a b : ) ,
3 * a * b * (a * accCubeTriLinSymm (P g, P g, P! f)
+ b * accCubeTriLinSymm (P! f, P! f, P g)) = 0 := by
3 * a * b * (a * accCubeTriLinSymm (P g) (P g) (P! f)
+ b * accCubeTriLinSymm (P! f) (P! f) (P g)) = 0 := by
intro g f hS a b
have h1 := h g f hS a b
change accCubeTriLinSymm.toCubic (a • P g + b • P! f) = 0 at h1
@ -57,7 +57,7 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S)
lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : lineInCubic S) :
∀ (g : Fin n → ) (f : Fin n → ) (_ : S.val = P g + P! f),
accCubeTriLinSymm (P g, P g, P! f) = 0 := by
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
intro g f hS
linear_combination 2 / 3 * (lineInCubic_expand h g f hS 1 1 ) -
(lineInCubic_expand h g f hS 1 2 ) / 6
@ -91,7 +91,7 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ + 1)).LinSols}
(LIC : lineInCubicPerm S) :
∀ (j : Fin n.succ) (g f : Fin n.succ → ) (_ : S.val = Pa g f) ,
(S.val (δ!₂ j) - S.val (δ!₁ j))
* accCubeTriLinSymm.toFun (P g, P g, basis!AsCharges j) = 0 := by
* accCubeTriLinSymm (P g) (P g) (basis!AsCharges j) = 0 := by
intro j g f h
let S' := (FamilyPermutations (2 * n.succ + 1)).linSolRep
(pairSwap (δ!₁ j) (δ!₂ j)) S
@ -107,7 +107,7 @@ lemma lineInCubicPerm_swap {S : (PureU1 (2 * n.succ + 1)).LinSols}
lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
(f g : Fin n.succ.succ → ) (hS : S.val = Pa f g) :
accCubeTriLinSymm.toFun (P f, P f, basis!AsCharges 0) =
accCubeTriLinSymm (P f) (P f) (basis!AsCharges 0) =
(S.val (δ!₁ 0) + S.val (δ!₂ 0)) * (2 * S.val δ!₃ + S.val (δ!₁ 0) + S.val (δ!₂ 0)) := by
rw [P_P_P!_accCube f 0]
rw [← Pa_δa₁ f g]

View file

@ -33,13 +33,13 @@ open VectorLikeOddPlane
show that this can be extended to a complete solution. -/
def parameterizationAsLinear (g f : Fin n → ) (a : ) :
(PureU1 (2 * n + 1)).LinSols :=
a • ((accCubeTriLinSymm (P! f, P! f, P g)) • P' g +
(- accCubeTriLinSymm (P g, P g, P! f)) • P!' f)
a • ((accCubeTriLinSymm (P! f) (P! f) (P g)) • P' g +
(- accCubeTriLinSymm (P g) (P g) (P! f)) • P!' f)
lemma parameterizationAsLinear_val (g f : Fin n → ) (a : ) :
(parameterizationAsLinear g f a).val =
a • ((accCubeTriLinSymm (P! f, P! f, P g)) • P g +
(- accCubeTriLinSymm (P g, P g, P! f)) • P! f) := by
a • ((accCubeTriLinSymm (P! f) (P! f) (P g)) • P g +
(- accCubeTriLinSymm (P g) (P g) (P! f)) • P! f) := by
rw [parameterizationAsLinear]
change a • (_ • (P' g).val + _ • (P!' f).val) = _
rw [P'_val, P!'_val]
@ -65,8 +65,8 @@ def parameterization (g f : Fin n → ) (a : ) :
lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
(g f : Fin n → ) (hS : S.val = P g + P! f) :
accCubeTriLinSymm (P g, P g, P! f) =
- accCubeTriLinSymm (P! f, P! f, P g) := by
accCubeTriLinSymm (P g) (P g) (P! f) =
- accCubeTriLinSymm (P! f) (P! f) (P g) := by
have hC := S.cubicSol
rw [hS] at hC
change (accCube (2 * n + 1)) (P g + P! f) = 0 at hC
@ -79,11 +79,11 @@ lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
In this case our parameterization above will be able to recover this point. -/
def genericCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
∀ (g f : Fin n.succ → ) (_ : S.val = P g + P! f) ,
accCubeTriLinSymm (P g, P g, P! f) ≠ 0
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
(hs : ∃ (g f : Fin n.succ → ), S.val = P g + P! f ∧
accCubeTriLinSymm (P g, P g, P! f) ≠ 0) : genericCase S := by
accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0) : genericCase S := by
intro g f hS hC
obtain ⟨g', f', hS', hC'⟩ := hs
rw [hS] at hS'
@ -95,11 +95,11 @@ lemma genericCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
In this case we will show that S is zero if it is true for all permutations. -/
def specialCase (S : (PureU1 (2 * n.succ + 1)).Sols) : Prop :=
∀ (g f : Fin n.succ → ) (_ : S.val = P g + P! f) ,
accCubeTriLinSymm (P g, P g, P! f) = 0
accCubeTriLinSymm (P g) (P g) (P! f) = 0
lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
(hs : ∃ (g f : Fin n.succ → ), S.val = P g + P! f ∧
accCubeTriLinSymm (P g, P g, P! f) = 0) : specialCase S := by
accCubeTriLinSymm (P g) (P g) (P! f) = 0) : specialCase S := by
intro g f hS
obtain ⟨g', f', hS', hC'⟩ := hs
rw [hS] at hS'
@ -110,8 +110,8 @@ lemma specialCase_exists (S : (PureU1 (2 * n.succ + 1)).Sols)
lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) :
genericCase S specialCase S := by
obtain ⟨g, f, h⟩ := span_basis S.1.1
have h1 : accCubeTriLinSymm (P g, P g, P! f) ≠ 0
accCubeTriLinSymm (P g, P g, P! f) = 0 := by
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
exact ne_or_eq _ _
cases h1 <;> rename_i h1
exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩)
@ -120,7 +120,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ + 1)).Sols) :
theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : genericCase S) :
∃ g f a, S = parameterization g f a := by
obtain ⟨g, f, hS⟩ := span_basis S.1.1
use g, f, (accCubeTriLinSymm (P! f, P! f, P g))⁻¹
use g, f, (accCubeTriLinSymm (P! f) (P! f) (P g))⁻¹
rw [parameterization]
apply ACCSystem.Sols.ext
rw [parameterizationAsLinear_val]
@ -149,7 +149,7 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
rw [h]
rw [anomalyFree_param _ _ hS] at h
simp at h
change accCubeTriLinSymm (P! f, P! f, P g) = 0 at h
change accCubeTriLinSymm (P! f) (P! f) (P g) = 0 at h
erw [h]
simp