refactor: Simp lemmas

This commit is contained in:
jstoobysmith 2024-10-12 07:57:35 +00:00
parent cd1b30c069
commit 1651b265e7
16 changed files with 116 additions and 86 deletions

View file

@ -31,7 +31,7 @@ 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]
simp only [HSMul.hSMul, SMul.smul, eq_ratCast, Rat.cast_eq_id, id_eq]
rw [← Finset.mul_sum]
/-- The symmetric trilinear form used to define the cubic anomaly. -/
@ -40,7 +40,7 @@ def accCubeTriLinSymm {n : } : TriLinearSymm (PureU1Charges n).Charges := Tri
(fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i)
(by
intro a S L T
simp [HSMul.hSMul]
simp only [PureU1Charges_numberCharges, HSMul.hSMul, ACCSystemCharges.chargesModule_smul]
rw [Finset.mul_sum]
apply Fintype.sum_congr
intro i
@ -108,7 +108,7 @@ open BigOperators
lemma pureU1_linear {n : } (S : (PureU1 n.succ).LinSols) :
∑ i, S.val i = 0 := by
have hS := S.linearSol
simp at hS
simp only [succ_eq_add_one, PureU1_numberLinear, PureU1_linearACCs] at hS
exact hS 0
lemma pureU1_cube {n : } (S : (PureU1 n.succ).Sols) :
@ -120,7 +120,7 @@ lemma pureU1_cube {n : } (S : (PureU1 n.succ).Sols) :
lemma pureU1_last {n : } (S : (PureU1 n.succ).LinSols) :
S.val (Fin.last n) = - ∑ i : Fin n, S.val i.castSucc := by
have hS := pureU1_linear S
simp at hS
simp only [succ_eq_add_one, PureU1_numberCharges] at hS
rw [Fin.sum_univ_castSucc] at hS
linear_combination hS
@ -134,7 +134,7 @@ lemma pureU1_anomalyFree_ext {n : } {S T : (PureU1 n.succ).LinSols}
obtain ⟨j, hj⟩ := hiCast
rw [← hj]
exact h j
· simp at hi
· simp only [succ_eq_add_one, PureU1_numberCharges, ne_eq, Decidable.not_not] at hi
rw [hi, pureU1_last, pureU1_last]
simp only [neg_inj]
apply Finset.sum_congr

View file

@ -113,7 +113,7 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ)).LinSols}
have h1 := Pa_δ!₄ f g
have h2 := Pa_δ!₁ f g (Fin.last n)
have h3 := Pa_δ!₂ f g (Fin.last n)
simp at h1 h2 h3
simp only [Fin.succ_last, Nat.succ_eq_add_one] at h1 h2 h3
have hl : f (Fin.succ (Fin.last n)) = - Pa f g δ!₄ := by
simp_all only [Fin.succ_last, neg_neg]
erw [hl] at h2
@ -133,7 +133,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols}
obtain ⟨g, f, hfg⟩ := span_basis S
have h1 := lineInCubicPerm_swap LIC (Fin.last n) g f hfg
rw [P_P_P!_accCube' g f hfg] at h1
simp at h1
simp only [Nat.succ_eq_add_one, neg_add_rev, mul_eq_zero] at h1
cases h1 <;> rename_i h1
· apply Or.inl
linear_combination h1
@ -148,7 +148,8 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
δ!₄ ?_ ?_ ?_ ?_
· simp [Fin.ext_iff, δ!₂, δ!₁]
· simp [Fin.ext_iff, δ!₂, δ!₄]
· simp [Fin.ext_iff, δ!₁, δ!₄]
· simp only [Nat.succ_eq_add_one, δ!₁, δ!₄, Fin.isValue, ne_eq, Fin.ext_iff, Fin.coe_cast,
Fin.coe_natAdd, Fin.coe_castAdd, Fin.val_last, Fin.coe_fin_one, add_zero, add_right_inj]
omega
· exact fun M => lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)

View file

@ -58,14 +58,18 @@ lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : LineIn
erw [sq_eq_sq_iff_eq_or_eq_neg] at h
rw [LineInPlaneCond] at hS
simp only [LineInPlaneProp] at hS
simp [not_or] at h
simp only [Nat.succ_eq_add_one, Fin.succ_last, not_or] at h
have h1 (i : Fin n) : S.val i.castSucc.castSucc =
- (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by
have h1S := hS (Fin.last n).castSucc ((Fin.last n).succ) i.castSucc.castSucc
(by simp; rw [Fin.ext_iff]; exact Nat.ne_add_one ↑(Fin.last n).castSucc)
(by simp; rw [Fin.ext_iff]; simp; omega)
(by simp; rw [Fin.ext_iff]; simp; omega)
simp_all
(by simp only [Fin.ext_iff, Nat.succ_eq_add_one, Fin.succ_last, ne_eq]
exact Nat.ne_add_one ↑(Fin.last n).castSucc)
(by simp only [Nat.succ_eq_add_one, Fin.succ_last, ne_eq, Fin.ext_iff, Fin.val_last,
Fin.coe_castSucc]
omega)
(by simp only [Nat.succ_eq_add_one, ne_eq, Fin.ext_iff, Fin.coe_castSucc, Fin.val_last]
omega)
simp_all only [Nat.succ_eq_add_one, ne_eq, Fin.succ_last, false_or, neg_add_rev]
field_simp
linear_combination h1S
have h2 := pureU1_last S
@ -152,7 +156,7 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
rw [h4, h5] at hcube
have h6 : S.val (3 : Fin 4) ^ 3 = 0 := by
linear_combination -1 * hcube / 24
simp at h6
simp only [Fin.isValue, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff] at h6
simp_all
lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols}

View file

@ -60,7 +60,7 @@ lemma parameterizationCharge_cube (g f : Fin n → ) (a : ) :
/-- Given a `g f : Fin n → ` and a `a : ` we form a solution. -/
def parameterization (g f : Fin n → ) (a : ) :
(PureU1 (2 * n + 1)).Sols :=
⟨⟨parameterizationAsLinear g f a, by intro i; simp at i; exact Fin.elim0 i⟩,
⟨⟨parameterizationAsLinear g f a, fun i => Fin.elim0 i⟩,
parameterizationCharge_cube g f a⟩
lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
@ -125,13 +125,10 @@ theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
apply ACCSystem.Sols.ext
rw [parameterizationAsLinear_val]
change S.val = _ • (_ • P g + _• P! f)
rw [anomalyFree_param _ _ hS]
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel₀, one_smul]
rw [anomalyFree_param _ _ hS, neg_neg, ← smul_add, smul_smul, inv_mul_cancel₀, one_smul]
· exact hS
· have h := h g f hS
rw [anomalyFree_param _ _ hS] at h
simp at h
exact h
· simpa only [Nat.succ_eq_add_one, accCubeTriLinSymm_toFun_apply_apply, ne_eq,
anomalyFree_param _ _ hS, neg_eq_zero] using h g f hS
lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
(h : SpecialCase S) :
@ -139,15 +136,13 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
intro g f hS a b
erw [TriLinearSymm.toCubic_add]
rw [HomogeneousCubic.map_smul, HomogeneousCubic.map_smul]
erw [P_accCube]
erw [P!_accCube]
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₃]
rw [h]
accCubeTriLinSymm.map_smul₃, h]
rw [anomalyFree_param _ _ hS] at h
simp at h
simp only [Nat.succ_eq_add_one, accCubeTriLinSymm_toFun_apply_apply, neg_eq_zero] at h
change accCubeTriLinSymm (P! f) (P! f) (P g) = 0 at h
erw [h]
simp

View file

@ -33,7 +33,7 @@ def Y₁ : (PlusU1 1).Sols where
| (5 : Fin 6) => 0
linearSol := by
intro i
simp at i
simp only [PlusU1_numberLinear] at i
match i with
| 0 => rfl
| 1 => rfl
@ -41,7 +41,7 @@ def Y₁ : (PlusU1 1).Sols where
| 3 => rfl
quadSol := by
intro i
simp at i
simp only [PlusU1_numberQuadratic] at i
match i with
| 0 => rfl
cubicSol := by rfl