refactor: Replace some simp with simp only
This commit is contained in:
parent
da5e0e3f00
commit
49d089d4cd
17 changed files with 56 additions and 41 deletions
|
@ -71,9 +71,11 @@ lemma doublePoint_Y₃_Y₃ (R : MSSMACC.LinSols) :
|
|||
simp only [mul_one, Fin.isValue, toSMSpecies_apply, one_mul, mul_neg, neg_mul, neg_neg, mul_zero,
|
||||
zero_mul, add_zero, Hd_apply, Fin.reduceFinMk, Hu_apply]
|
||||
have hLin := R.linearSol
|
||||
simp at hLin
|
||||
simp only [MSSMACC_numberLinear, MSSMACC_linearACCs, Nat.reduceMul, Fin.isValue,
|
||||
Fin.reduceFinMk] at hLin
|
||||
have h3 := hLin 3
|
||||
simp [Fin.sum_univ_three] at h3
|
||||
simp only [Fin.isValue, Fin.sum_univ_three, Prod.mk_zero_zero, Prod.mk_one_one, LinearMap.coe_mk,
|
||||
AddHom.coe_mk] at h3
|
||||
linear_combination (norm := ring_nf) 6 * h3
|
||||
simp only [Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one, add_add_sub_cancel, add_neg_cancel]
|
||||
|
||||
|
|
|
@ -51,7 +51,8 @@ def permCharges {n : ℕ} : Representation ℚ (PermGroup n) (PureU1 n).Charges
|
|||
|
||||
lemma accGrav_invariant {n : ℕ} (f : (PermGroup n)) (S : (PureU1 n).Charges) :
|
||||
PureU1.accGrav n (permCharges f S) = accGrav n S := by
|
||||
simp [accGrav, permCharges]
|
||||
simp only [accGrav, PermGroup, permCharges, MonoidHom.coe_mk, OneHom.coe_mk, LinearMap.coe_mk,
|
||||
AddHom.coe_mk, chargeMap_apply]
|
||||
apply (Equiv.Perm.sum_comp _ _ _ ?_)
|
||||
simp
|
||||
|
||||
|
@ -144,7 +145,7 @@ lemma pairSwap_other {n : ℕ} (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) :
|
|||
|
||||
lemma pairSwap_inv_other {n : ℕ} {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) :
|
||||
(pairSwap i j).invFun k = k := by
|
||||
simp [pairSwap]
|
||||
simp only [pairSwap, Equiv.invFun_as_coe, Equiv.coe_fn_symm_mk]
|
||||
split
|
||||
· rename_i h
|
||||
exact False.elim (hik (id (Eq.symm h)))
|
||||
|
@ -197,8 +198,8 @@ lemma permTwo_fst : (permTwo hij hij').toFun i' = i := by
|
|||
have ht := Equiv.extendSubtype_apply_of_mem
|
||||
((permTwoInj hij').toEquivRange.symm.trans
|
||||
(permTwoInj hij).toEquivRange) i' (permTwoInj_fst hij')
|
||||
simp at ht
|
||||
simp [ht, permTwoInj_fst_apply]
|
||||
simp only [Equiv.trans_apply, Function.Embedding.toEquivRange_apply] at ht
|
||||
simp only [Equiv.toFun_as_coe, ht, permTwoInj_fst_apply, Fin.isValue]
|
||||
rfl
|
||||
|
||||
lemma permTwo_snd : (permTwo hij hij').toFun j' = j := by
|
||||
|
|
|
@ -269,7 +269,7 @@ def cubeTriLin : TriLinearSymm (SMCharges n).Charges := TriLinearSymm.mk₃
|
|||
apply Fintype.sum_congr
|
||||
intro i
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue]
|
||||
ring)
|
||||
(by
|
||||
intro S T R L
|
||||
|
|
|
@ -38,12 +38,12 @@ variable {n : ℕ}
|
|||
|
||||
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
|
||||
exact hS 1
|
||||
|
||||
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
|
||||
|
@ -55,7 +55,7 @@ def chargeToLinear (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accS
|
|||
(SMNoGrav n).LinSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [SMNoGrav_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact hSU2
|
||||
| 1 => exact hSU3⟩
|
||||
|
|
|
@ -37,17 +37,17 @@ variable {n : ℕ}
|
|||
|
||||
lemma gravSol (S : (SM n).LinSols) : accGrav S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SM_numberLinear, SM_linearACCs, Fin.isValue] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU2Sol (S : (SM n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SM_numberLinear, SM_linearACCs, Fin.isValue] at hS
|
||||
exact hS 1
|
||||
|
||||
lemma SU3Sol (S : (SM n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SM_numberLinear, SM_linearACCs, Fin.isValue] at hS
|
||||
exact hS 2
|
||||
|
||||
lemma cubeSol (S : (SM n).Sols) : accCube S.val = 0 := S.cubicSol
|
||||
|
@ -58,7 +58,7 @@ def chargeToLinear (S : (SM n).Charges) (hGrav : accGrav S = 0)
|
|||
(hSU2 : accSU2 S = 0) (hSU3 : accSU3 S = 0) : (SM n).LinSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [SM_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact hGrav
|
||||
| 1 => exact hSU2
|
||||
|
@ -100,14 +100,14 @@ def perm (n : ℕ) : ACCSystemGroupAction (SM n) where
|
|||
rep := repCharges
|
||||
linearInvariant := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [SM_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact accGrav_invariant
|
||||
| 1 => exact accSU2_invariant
|
||||
| 2 => exact accSU3_invariant
|
||||
quadInvariant := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [SM_numberQuadratic] at i
|
||||
exact Fin.elim0 i
|
||||
cubicInvariant := accCube_invariant
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue