refactor: Lint

This commit is contained in:
jstoobysmith 2024-07-19 17:00:32 -04:00
parent 52e591fa7a
commit 9f27a3a9fd
46 changed files with 176 additions and 168 deletions

View file

@ -234,7 +234,7 @@ lemma Sols.ext {χ : ACCSystem} {S T : χ.Sols} (h : S.val = T.val) :
/-- A charge `S` is a solution if it extends to a solution. -/
def IsSolution (χ : ACCSystem) (S : χ.Charges) : Prop :=
∃ (sol : χ.Sols), sol.val = S
∃ (sol : χ.Sols), sol.val = S
/-- An instance giving the properties and structures to define an action of `` on `Sols`. -/
instance solsMulAction (χ : ACCSystem) : MulAction χ.Sols where

View file

@ -43,8 +43,8 @@ instance {χ : ACCSystem} (G : ACCSystemGroupAction χ) : Group G.group := G.gro
def linSolMap {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g : G.group) :
χ.LinSols →ₗ[] χ.LinSols where
toFun S := ⟨G.rep g S.val, by
intro i
rw [G.linearInvariant, S.linearSol]⟩
intro i
rw [G.linearInvariant, S.linearSol]⟩
map_add' S T := by
apply ACCSystemLinear.LinSols.ext
exact (G.rep g).map_add' _ _
@ -82,9 +82,9 @@ lemma rep_linSolRep_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (g :
instance quadSolAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) :
MulAction G.group χ.QuadSols where
smul f S := ⟨G.linSolRep f S.1, by
intro i
simp only [linSolRep_apply_apply_val]
rw [G.quadInvariant, S.quadSol]⟩
intro i
simp only [linSolRep_apply_apply_val]
rw [G.quadInvariant, S.quadSol]⟩
mul_smul f1 f2 S := by
apply ACCSystemQuad.QuadSols.ext
change (G.rep.toFun (f1 * f2)) S.val = _
@ -107,9 +107,9 @@ lemma rep_quadSolAction_commute {χ : ACCSystem} (G : ACCSystemGroupAction χ) (
/-- The group action acting on solutions to the anomaly cancellation conditions. -/
instance solAction {χ : ACCSystem} (G : ACCSystemGroupAction χ) : MulAction G.group χ.Sols where
smul g S := ⟨G.quadSolAction.toFun S.1 g, by
simp only [MulAction.toFun_apply]
change χ.cubicACC (G.rep g S.val) = 0
rw [G.cubicInvariant, S.cubicSol]⟩
simp only [MulAction.toFun_apply]
change χ.cubicACC (G.rep g S.val) = 0
rw [G.cubicInvariant, S.cubicSol]⟩
mul_smul f1 f2 S := by
apply ACCSystem.Sols.ext
change (G.rep.toFun (f1 * f2)) S.val = _

View file

@ -56,7 +56,7 @@ def toSplitSMPlusH : MSSMCharges.Charges ≃ (Fin 18 → ) × (Fin 2 → )
/-- An equivalence between `(Fin 18 → )` and `(Fin 6 → Fin 3 → )`. -/
@[simps!]
def toSpeciesMaps' : (Fin 18 → ) ≃ (Fin 6 → Fin 3 → ) :=
((Equiv.curry _ _ _).symm.trans
((Equiv.curry _ _ _).symm.trans
((@finProdFinEquiv 6 3).arrowCongr (Equiv.refl ))).symm
/-- An equivalence between `MSSMCharges.charges` and `(Fin 6 → Fin 3 → ) × (Fin 2 → ))`.
@ -367,7 +367,7 @@ lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.Charges) :
rw [add_comm (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L) _]
rw [add_assoc]
rw [← add_assoc _ _ (2 * Hd S * Hd R * Hd L + 2 * Hu S * Hu R * Hu L +
(2 * Hd T * Hd R * Hd L + 2 * Hu T * Hu R * Hu L))]
(2 * Hd T * Hd R * Hd L + 2 * Hu T * Hu R * Hu L))]
congr 1
rw [← Finset.sum_add_distrib]
apply Fintype.sum_congr

View file

@ -109,8 +109,8 @@ lemma quad_self_proj (T : MSSMACC.Sols) :
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
((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₁]
@ -151,7 +151,7 @@ lemma cube_proj_proj_B₃ (T : MSSMACC.LinSols) :
rw [cubeTriLin.map_add₂, cubeTriLin.map_smul₁, cubeTriLin.map_smul₂]
rw [cubeTriLin.swap₁, cubeTriLin.swap₂, doublePoint_Y₃_B₃]
rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.swap₁, cubeTriLin.swap₂,
cubeTriLin.swap₁, doublePoint_B₃_B₃]
cubeTriLin.swap₁, doublePoint_B₃_B₃]
rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂]
ring

View file

@ -48,7 +48,7 @@ lemma planeY₃B₃_eq (R : MSSMACC.AnomalyFreePerp) (a b c : ) (h : a = a'
lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ) (hR' : R.val ≠ 0)
(h : (planeY₃B₃ R a b c).val = (planeY₃B₃ R a' b' c').val) :
a = a' ∧ b = b' ∧ c = c' := by
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

View file

@ -43,7 +43,7 @@ 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`. -/
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
@ -338,8 +338,7 @@ lemma inQuadToSol_proj (T : InQuadSol) : inQuadToSol (inQuadProj T) = T.val := b
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
* 3) = cubicCoeff T.val := by
dot Y₃.val B₃.val ^ 3 * cubeTriLin T.val.val T.val.val B₃.val ^ 2* 3) = cubicCoeff T.val := by
rw [cubicCoeff]
ring
rw [h1]
@ -360,7 +359,7 @@ def inQuadCubeToSol : InQuadCube × × × → MSSMACC.Sols := fun (R
simp)
lemma inQuadCubeToSol_smul (R : InQuadCube) (c₁ c₂ c₃ d : ) :
inQuadCubeToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadCubeToSol (R, c₁, c₂, c₃):= by
inQuadCubeToSol (R, (d * c₁), (d * c₂), (d * c₃)) = d • inQuadCubeToSol (R, c₁, c₂, c₃) := by
apply ACCSystem.Sols.ext
change (planeY₃B₃ _ _ _ _).val = _
rw [planeY₃B₃_smul]

View file

@ -31,8 +31,8 @@ def accGrav (n : ) : ((PureU1Charges n).Charges →ₗ[] ) where
toFun S := ∑ i : Fin n, S i
map_add' S T := Finset.sum_add_distrib
map_smul' a S := by
simp [HSMul.hSMul, SMul.smul]
rw [← Finset.mul_sum]
simp [HSMul.hSMul, SMul.smul]
rw [← Finset.mul_sum]
/-- The symmetric trilinear form used to define the cubic anomaly. -/
@[simps!]
@ -95,7 +95,7 @@ def PureU1 (n : ) : ACCSystem where
cubicACC := PureU1.accCube n
/-- An equivalence of vector spaces of charges when the number of fermions is equal. -/
def pureU1EqCharges {n m : } (h : n = m):
def pureU1EqCharges {n m : } (h : n = m) :
(PureU1 n).Charges ≃ₗ[] (PureU1 m).Charges where
toFun f := f ∘ Fin.cast h.symm
invFun f := f ∘ Fin.cast h

View file

@ -24,14 +24,11 @@ namespace BasisLinear
last position. -/
@[simp]
def asCharges (j : Fin n) : (PureU1 n.succ).Charges :=
(fun i =>
if i = j.castSucc then
1
else
if i = Fin.last n then
- 1
else
0)
(fun i =>
if i = j.castSucc then 1
else if i = Fin.last n then
- 1
else 0)
lemma asCharges_eq_castSucc (j : Fin n) :
asCharges j (Fin.castSucc j) = 1 := by
@ -123,7 +120,7 @@ def asBasis : Basis (Fin n) ((PureU1 n.succ).LinSols) where
repr := coordinateMap
instance : Module.Finite ((PureU1 n.succ).LinSols) :=
Module.Finite.of_basis asBasis
Module.Finite.of_basis asBasis
lemma finrank_AnomalyFreeLinear :
FiniteDimensional.finrank (((PureU1 n.succ).LinSols)) = n := by

View file

@ -96,7 +96,7 @@ lemma is_zero (h0 : S (0 : Fin n.succ) = 0) : S = 0 := by
/-- A boundary of `S : (PureU1 n.succ).charges` (assumed sorted, constAbs and non-zero)
is defined as a element of `k ∈ Fin n` such that `S k.castSucc` and `S k.succ` are different signs.
-/
-/
@[simp]
def Boundary (S : (PureU1 n.succ).Charges) (k : Fin n) : Prop := S k.castSucc < 0 ∧ 0 < S k.succ

View file

@ -550,7 +550,7 @@ def Pa' (f : (Fin n.succ) ⊕ (Fin n) → ) : (PureU1 (2 * n.succ)).LinSols :
∑ i, f i • basisa i
lemma Pa'_P'_P!' (f : (Fin n.succ) ⊕ (Fin n) → ) :
Pa' f = P' (f ∘ Sum.inl) + P!' (f ∘ Sum.inr):= by
Pa' f = P' (f ∘ Sum.inl) + P!' (f ∘ Sum.inr) := by
exact Fintype.sum_sum_type _
lemma P'_val (f : Fin n.succ → ) : (P' f).val = P f := by
@ -617,7 +617,7 @@ lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ) : Pa' f = Pa' f' ↔ f =
simp
have h2 : ∀ i, (f i + (- f' i)) = 0 := by
exact Fintype.linearIndependent_iff.mp (@basisa_linear_independent n)
(fun i => f i + -f' i) h1
(fun i => f i + -f' i) h1
have h2i := h2 i
linarith
intro h
@ -685,9 +685,9 @@ lemma span_basis (S : (PureU1 (2 * n.succ)).LinSols) :
rfl
lemma P!_in_span (f : Fin n → ) : P! f ∈ Submodule.span (Set.range basis!AsCharges) := by
rw [(mem_span_range_iff_exists_fun )]
use f
rfl
rw [(mem_span_range_iff_exists_fun )]
use f
rfl
lemma smul_basis!AsCharges_in_span (S : (PureU1 (2 * n.succ)).LinSols) (j : Fin n) :
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∈
@ -700,11 +700,9 @@ lemma smul_basis!AsCharges_in_span (S : (PureU1 (2 * n.succ)).LinSols) (j : Fin
lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
(hS : ((FamilyPermutations (2 * n.succ)).linSolRep
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') (g : Fin n.succ → ) (f : Fin n → )
(h : S.val = P g + P! f):
(g' : Fin n.succ → ) (f' : Fin n → ),
S'.val = P g' + P! f' ∧ P! f' = P! f +
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∧ g' = g := by
(h : S.val = P g + P! f) : ∃ (g' : Fin n.succ → ) (f' : Fin n → ),
S'.val = P g' + P! f' ∧ P! f' = P! f +
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∧ g' = g := by
let X := P! f + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j
have hX : X ∈ Submodule.span (Set.range (basis!AsCharges)) := by
apply Submodule.add_mem
@ -723,10 +721,8 @@ lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
exact hS
lemma vectorLikeEven_in_span (S : (PureU1 (2 * n.succ)).LinSols)
(hS : VectorLikeEven S.val) :
∃ (M : (FamilyPermutations (2 * n.succ)).group),
(FamilyPermutations (2 * n.succ)).linSolRep M S
∈ Submodule.span (Set.range basis) := by
(hS : VectorLikeEven S.val) : ∃ (M : (FamilyPermutations (2 * n.succ)).group),
(FamilyPermutations (2 * n.succ)).linSolRep M S ∈ Submodule.span (Set.range basis) := by
use (Tuple.sort S.val).symm
change sortAFL S ∈ Submodule.span (Set.range basis)
rw [mem_span_range_iff_exists_fun ]

View file

@ -49,23 +49,23 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S)
change accCubeTriLinSymm.toCubic (a • P g + b • P! f) = 0 at h1
simp only [TriLinearSymm.toCubic_add] at h1
simp only [HomogeneousCubic.map_smul,
accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂, accCubeTriLinSymm.map_smul₃] at h1
accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂, accCubeTriLinSymm.map_smul₃] at h1
erw [P_accCube, P!_accCube] at h1
rw [← h1]
ring
/--
This lemma states that for a given `S` of type `(PureU1 (2 * n.succ)).AnomalyFreeLinear` and
a proof `h` that the line through `S` lies on a cubic curve,
for any functions `g : Fin n.succ → ` and `f : Fin n → `, if `S.val = P g + P! f`,
then `accCubeTriLinSymm.toFun (P g, P g, P! f) = 0`.
This lemma states that for a given `S` of type `(PureU1 (2 * n.succ)).AnomalyFreeLinear` and
a proof `h` that the line through `S` lies on a cubic curve,
for any functions `g : Fin n.succ → ` and `f : Fin n → `, if `S.val = P g + P! f`,
then `accCubeTriLinSymm.toFun (P g, P g, P! f) = 0`.
-/
lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n.succ)).LinSols} (h : LineInCubic S) :
∀ (g : Fin n.succ → ) (f : Fin n → ) (_ : S.val = P g + P! f),
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
(lineInCubic_expand h g f hS 1 2) / 6
/-- A `LinSol` satisfies `LineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
def LineInCubicPerm (S : (PureU1 (2 * n.succ)).LinSols) : Prop :=

View file

@ -49,15 +49,15 @@ lemma parameterizationAsLinear_val (g : Fin n.succ → ) (f : Fin n → )
change a • (_ • (P' g).val + _ • (P!' f).val) = _
rw [P'_val, P!'_val]
lemma parameterizationCharge_cube (g : Fin n.succ → ) (f : Fin n → ) (a : ):
lemma parameterizationCharge_cube (g : Fin n.succ → ) (f : Fin n → ) (a : ) :
accCube (2* n.succ) (parameterizationAsLinear g f a).val = 0 := by
change accCubeTriLinSymm.toCubic _ = 0
rw [parameterizationAsLinear_val, HomogeneousCubic.map_smul,
TriLinearSymm.toCubic_add, HomogeneousCubic.map_smul, HomogeneousCubic.map_smul]
erw [P_accCube, P!_accCube]
rw [accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
ring
/-- The construction of a `Sol` from a `Fin n.succ → `, a `Fin n → ` and a ``. -/
@ -112,7 +112,7 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ)).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
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⟩)
@ -142,8 +142,8 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ)).Sols}
erw [P_accCube, P!_accCube]
have h := h g f hS
rw [accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
rw [h]
rw [anomalyFree_param _ _ hS] at h
simp at h

View file

@ -89,7 +89,7 @@ lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
intro a
linarith
have ht : S.val ((Fin.last n.succ.succ.succ).succ) =
- S.val ((Fin.last n.succ.succ.succ).castSucc) := by
- S.val ((Fin.last n.succ.succ.succ).castSucc) := by
rw [← mul_right_inj' hx]
simp [Nat.succ_eq_add_one]
simp at h

View file

@ -535,7 +535,7 @@ def Pa' (f : (Fin n) ⊕ (Fin n) → ) : (PureU1 (2 * n + 1)).LinSols :=
∑ i, f i • basisa i
lemma Pa'_P'_P!' (f : (Fin n) ⊕ (Fin n) → ) :
Pa' f = P' (f ∘ Sum.inl) + P!' (f ∘ Sum.inr):= by
Pa' f = P' (f ∘ Sum.inl) + P!' (f ∘ Sum.inr) := by
exact Fintype.sum_sum_type _
lemma P'_val (f : Fin n → ) : (P' f).val = P f := by
@ -602,7 +602,7 @@ lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n.succ) → ) : Pa' f = Pa' f' ↔
simp
have h2 : ∀ i, (f i + (- f' i)) = 0 := by
exact Fintype.linearIndependent_iff.mp (@basisa_linear_independent n)
(fun i => f i + -f' i) h1
(fun i => f i + -f' i) h1
have h2i := h2 i
linarith
intro h
@ -671,15 +671,15 @@ lemma span_basis (S : (PureU1 (2 * n.succ + 1)).LinSols) :
lemma span_basis_swap! {S : (PureU1 (2 * n.succ + 1)).LinSols} (j : Fin n.succ)
(hS : ((FamilyPermutations (2 * n.succ + 1)).linSolRep
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') (g f : Fin n.succ → ) (hS1 : S.val = P g + P! f):
(pairSwap (δ!₁ j) (δ!₂ j))) S = S') (g f : Fin n.succ → ) (hS1 : S.val = P g + P! f) :
∃ (g' f' : Fin n.succ → ),
S'.val = P g' + P! f' ∧ P! f' = P! f +
S'.val = P g' + P! f' ∧ P! f' = P! f +
(S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∧ g' = g := by
let X := P! f + (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j
have hf : P! f ∈ Submodule.span (Set.range basis!AsCharges) := by
rw [(mem_span_range_iff_exists_fun )]
use f
rfl
rw [(mem_span_range_iff_exists_fun )]
use f
rfl
have hP : (S.val (δ!₂ j) - S.val (δ!₁ j)) • basis!AsCharges j ∈
Submodule.span (Set.range basis!AsCharges) := by
apply Submodule.smul_mem

View file

@ -49,7 +49,7 @@ lemma lineInCubic_expand {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S)
change accCubeTriLinSymm.toCubic (a • P g + b • P! f) = 0 at h1
simp only [TriLinearSymm.toCubic_add] at h1
simp only [HomogeneousCubic.map_smul,
accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂, accCubeTriLinSymm.map_smul₃] at h1
accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂, accCubeTriLinSymm.map_smul₃] at h1
erw [P_accCube, P!_accCube] at h1
rw [← h1]
ring
@ -59,7 +59,7 @@ lemma line_in_cubic_P_P_P! {S : (PureU1 (2 * n + 1)).LinSols} (h : LineInCubic S
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
(lineInCubic_expand h g f hS 1 2) / 6
/-- A `LinSol` satisfies `lineInCubicPerm` if all its permutations satisfy `lineInCubic`. -/
def LineInCubicPerm (S : (PureU1 (2 * n + 1)).LinSols) : Prop :=
@ -119,8 +119,8 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
have h2 := Pa_δa₂ f g 0
rw [← hS] at h1 h2 h4
simp at h2
have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0):= by
linear_combination -(1 * h1) - 1 * h4 - 1 * h2
have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0) := by
linear_combination -(1 * h1) - 1 * h4 - 1 * h2
rw [h5]
rw [δa₄_δ!₂]
have h0 : (δa₂ (0 : Fin n.succ)) = δ!₁ 0 := by

View file

@ -44,7 +44,7 @@ lemma parameterizationAsLinear_val (g f : Fin n → ) (a : ) :
change a • (_ • (P' g).val + _ • (P!' f).val) = _
rw [P'_val, P!'_val]
lemma parameterizationCharge_cube (g f : Fin n → ) (a : ):
lemma parameterizationCharge_cube (g f : Fin n → ) (a : ) :
(accCube (2 * n + 1)) (parameterizationAsLinear g f a).val = 0 := by
change accCubeTriLinSymm.toCubic _ = 0
rw [parameterizationAsLinear_val]
@ -53,8 +53,8 @@ lemma parameterizationCharge_cube (g f : Fin n → ) (a : ):
rw [HomogeneousCubic.map_smul, HomogeneousCubic.map_smul]
erw [P_accCube g, P!_accCube f]
rw [accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
ring
/-- Given a `g f : Fin n → ` and a `a : ` we form a solution. -/
@ -111,7 +111,7 @@ 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
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⟩)
@ -143,8 +143,8 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
erw [P!_accCube]
have h := h g f hS
rw [accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
accCubeTriLinSymm.map_smul₃]
rw [h]
rw [anomalyFree_param _ _ hS] at h
simp at h

View file

@ -310,8 +310,8 @@ lemma Prop_two (P : × → Prop) {S : (PureU1 n).LinSols}
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply] at h1
simp at h1
change P
(S.val ((permTwo hij hab).toFun a),
S.val ((permTwo hij hab).toFun b)) at h1
(S.val ((permTwo hij hab).toFun a),
S.val ((permTwo hij hab).toFun b)) at h1
erw [permTwo_fst,permTwo_snd] at h1
exact h1
@ -325,11 +325,11 @@ lemma Prop_three (P : × × → Prop) {S : (PureU1 n).LinSols}
intro i j k hij hjk hik
have h1 := h (permThree hij hjk hik hab hbc hac).symm
rw [FamilyPermutations_anomalyFreeLinear_apply, FamilyPermutations_anomalyFreeLinear_apply,
FamilyPermutations_anomalyFreeLinear_apply] at h1
FamilyPermutations_anomalyFreeLinear_apply] at h1
simp at h1
change P
(S.val ((permThree hij hjk hik hab hbc hac).toFun a),
S.val ((permThree hij hjk hik hab hbc hac).toFun b),
(S.val ((permThree hij hjk hik hab hbc hac).toFun a),
S.val ((permThree hij hjk hik hab hbc hac).toFun b),
S.val ((permThree hij hjk hik hab hbc hac).toFun c)) at h1
erw [permThree_fst,permThree_snd, permThree_thd] at h1
exact h1

View file

@ -24,7 +24,7 @@ variable {n : }
/-- A charge is sorted if for all `i ≤ j`, then `S i ≤ S j`. -/
@[simp]
def Sorted {n : } (S : (PureU1 n).Charges) : Prop :=
∀ i j (_ : i ≤ j), S i ≤ S j
∀ i j (_ : i ≤ j), S i ≤ S j
/-- Given a charge assignment `S`, the corresponding sorted charge assignment. -/
@[simp]

View file

@ -46,7 +46,7 @@ lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) :
have h1 := SU2Sol S.1.1
have h2 := SU3Sol S.1.1
simp only [accSU2, SMSpecies_numberCharges, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
Finset.sum_singleton, LinearMap.coe_mk, AddHom.coe_mk, accSU3] at h1 h2
Finset.sum_singleton, LinearMap.coe_mk, AddHom.coe_mk, accSU3] at h1 h2
erw [hQ] at h1 h2
simp_all
linear_combination 3 * h2

View file

@ -116,8 +116,8 @@ lemma cubic_zero_E'_zero (S : linearParameters) (hc : accCube (S.asCharges) = 0)
def bijection : linearParameters ≃ (SMNoGrav 1).LinSols where
toFun S := S.asLinear
invFun S := ⟨SMCharges.Q S.val (0 : Fin 1), (SMCharges.U S.val (0 : Fin 1) -
SMCharges.D S.val (0 : Fin 1))/2,
SMCharges.E S.val (0 : Fin 1)⟩
SMCharges.D S.val (0 : Fin 1))/2,
SMCharges.E S.val (0 : Fin 1)⟩
left_inv S := by
apply linearParameters.ext
rfl
@ -134,7 +134,7 @@ def bijection : linearParameters ≃ (SMNoGrav 1).LinSols where
intro i
rw [asLinear_val]
funext j
have hj : j = (0 : Fin 1):= by
have hj : j = (0 : Fin 1) := by
simp only [Fin.isValue]
ext
simp
@ -284,7 +284,7 @@ lemma cubic_v_or_w_zero (S : linearParametersQENeqZero) (h : accCube (bijection
(FLTThree : FermatLastTheoremWith 3) :
S.v = 0 S.w = 0 := by
rw [S.cubic] at h
have h1 : (-1)^3 = (-1 : ):= by rfl
have h1 : (-1)^3 = (-1 : ) := by rfl
rw [← h1] at h
by_contra hn
simp [not_or] at hn

View file

@ -70,7 +70,7 @@ lemma toSpecies_sum_invariant (m : ) (f : PermGroup n) (S : (SMCharges n).Cha
∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by
rw [repCharges_toSpecies]
exact Fintype.sum_equiv (f⁻¹ j) (fun x => ((fun a => a ^ m) ∘ (toSpecies j) S ∘ ⇑(f⁻¹ j)) x)
(fun x => ((fun a => a ^ m) ∘ (toSpecies j) S) x) (congrFun rfl)
(fun x => ((fun a => a ^ m) ∘ (toSpecies j) S) x) (congrFun rfl)
lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accGrav (repCharges f S) = accGrav S :=

View file

@ -260,7 +260,7 @@ def quadBiLin : BiLinearSymm (SMνCharges n).Charges := BiLinearSymm.mk₂
lemma quadBiLin_decomp (S T : (SMνCharges n).Charges) :
quadBiLin S T = ∑ i, Q S i * Q T i - 2 * ∑ i, U S i * U T i +
∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by
∑ i, D S i * D T i - ∑ i, L S i * L T i + ∑ i, E S i * E T i := by
erw [← quadBiLin.toFun_eq_coe]
rw [quadBiLin]
simp only [quadBiLin, BiLinearSymm.mk₂, AddHom.toFun_eq_coe, AddHom.coe_mk, LinearMap.coe_mk]

View file

@ -213,7 +213,7 @@ lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).Charges) (T R : (SMνCharg
lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).Charges) (R : (SMνCharges n).Charges) :
cubeTriLin (familyUniversal n S) (familyUniversal n T) R =
6 * S (0 : Fin 6) * T (0 : Fin 6) * ∑ i, Q R i +
3 * S (1 : Fin 6) * T (1 : Fin 6) * ∑ i, U R i
3 * S (1 : Fin 6) * T (1 : Fin 6) * ∑ i, U R i
+ 3 * S (2 : Fin 6) * T (2 : Fin 6) * ∑ i, D R i +
2 * S (3 : Fin 6) * T (3 : Fin 6) * ∑ i, L R i +
S (4 : Fin 6) * T (4 : Fin 6) * ∑ i, E R i + S (5 : Fin 6) * T (5 : Fin 6) * ∑ i, N R i := by

View file

@ -86,7 +86,7 @@ lemma add_quad (S : (PlusU1 n).QuadSols) (a b : ) :
def addQuad (S : (PlusU1 n).QuadSols) (a b : ) : (PlusU1 n).QuadSols :=
linearToQuad (a • S.1 + b • (BL n).1.1) (add_quad S a b)
lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ): addQuad S a 0 = a • S := by
lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ) : addQuad S a 0 = a • S := by
simp [addQuad, linearToQuad]
rfl

View file

@ -24,7 +24,7 @@ open BigOperators
/-- A proposition which is true if for a given `n`, a plane of charges of dimension `n` exists
in which each point is a solution. -/
def ExistsPlane (n : ) : Prop := ∃ (B : Fin n → (PlusU1 3).Charges),
LinearIndependent B ∧ ∀ (f : Fin n → ), (PlusU1 3).IsSolution (∑ i, f i • B i)
LinearIndependent B ∧ ∀ (f : Fin n → ), (PlusU1 3).IsSolution (∑ i, f i • B i)
lemma exists_plane_exists_basis {n : } (hE : ExistsPlane n) :
∃ (B : Fin 11 ⊕ Fin n → (PlusU1 3).Charges), LinearIndependent B := by
@ -38,7 +38,7 @@ lemma exists_plane_exists_basis {n : } (hE : ExistsPlane n) :
rw [@add_eq_zero_iff_eq_neg] at hg
rw [← @Finset.sum_neg_distrib] at hg
have h1 : ∑ x : Fin n, -(g (Sum.inr x) • Y (Sum.inr x)) =
∑ x : Fin n, (-g (Sum.inr x)) • Y (Sum.inr x):= by
∑ x : Fin n, (-g (Sum.inr x)) • Y (Sum.inr x) := by
apply Finset.sum_congr
simp only
intro i _

View file

@ -84,7 +84,7 @@ lemma add_quad (S : (PlusU1 n).QuadSols) (a b : ) :
def addQuad (S : (PlusU1 n).QuadSols) (a b : ) : (PlusU1 n).QuadSols :=
linearToQuad (a • S.1 + b • (Y n).1.1) (add_quad S a b)
lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ): addQuad S a 0 = a • S := by
lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ) : addQuad S a 0 = a • S := by
simp [addQuad, linearToQuad]
rfl