refactor: More simps
This commit is contained in:
parent
1651b265e7
commit
4a396783ab
17 changed files with 138 additions and 87 deletions
|
@ -24,9 +24,10 @@ namespace One
|
|||
theorem solEqZero (S : (PureU1 1).LinSols) : S = 0 := by
|
||||
apply ACCSystemLinear.LinSols.ext
|
||||
have hLin := pureU1_linear S
|
||||
simp at hLin
|
||||
simp only [succ_eq_add_one, reduceAdd, PureU1_numberCharges, univ_unique, Fin.default_eq_zero,
|
||||
Fin.isValue, sum_singleton] at hLin
|
||||
funext i
|
||||
simp at i
|
||||
simp only [PureU1_numberCharges] at i
|
||||
rw [show i = (0 : Fin 1) from Fin.fin_one_eq_zero i]
|
||||
exact hLin
|
||||
|
||||
|
|
|
@ -25,7 +25,7 @@ lemma cube_for_linSol' (S : (PureU1 3).LinSols) :
|
|||
3 * S.val (0 : Fin 3) * S.val (1 : Fin 3) * S.val (2 : Fin 3) = 0 ↔
|
||||
(PureU1 3).cubicACC S.val = 0 := by
|
||||
have hL := pureU1_linear S
|
||||
simp at hL
|
||||
simp only [succ_eq_add_one, reduceAdd, PureU1_numberCharges] at hL
|
||||
rw [Fin.sum_univ_three] at hL
|
||||
change _ ↔ accCube _ _ = _
|
||||
rw [accCube_explicit, Fin.sum_univ_three]
|
||||
|
@ -47,7 +47,7 @@ lemma three_sol_zero (S : (PureU1 3).Sols) : S.val (0 : Fin 3) = 0 ∨ S.val (1
|
|||
def solOfLinear (S : (PureU1 3).LinSols)
|
||||
(hS : S.val (0 : Fin 3) = 0 ∨ S.val (1 : Fin 3) = 0 ∨ S.val (2 : Fin 3) = 0) :
|
||||
(PureU1 3).Sols :=
|
||||
⟨⟨S, by intro i; simp at i; exact Fin.elim0 i⟩,
|
||||
⟨⟨S, fun i => Fin.elim0 i⟩,
|
||||
(cube_for_linSol S).mp hS⟩
|
||||
|
||||
theorem solOfLinear_surjects (S : (PureU1 3).Sols) :
|
||||
|
|
|
@ -23,9 +23,10 @@ namespace Two
|
|||
|
||||
/-- An equivalence between `LinSols` and `Sols`. -/
|
||||
def equiv : (PureU1 2).LinSols ≃ (PureU1 2).Sols where
|
||||
toFun S := ⟨⟨S, by intro i; simp at i; exact Fin.elim0 i⟩, by
|
||||
toFun S := ⟨⟨S, fun i => Fin.elim0 i⟩, by
|
||||
have hLin := pureU1_linear S
|
||||
simp at hLin
|
||||
simp only [succ_eq_add_one, reduceAdd, PureU1_numberCharges, Fin.sum_univ_two,
|
||||
Fin.isValue] at hLin
|
||||
erw [accCube_explicit]
|
||||
simp only [Fin.sum_univ_two, Fin.isValue]
|
||||
rw [show S.val (0 : Fin 2) = - S.val (1 : Fin 2) from eq_neg_of_add_eq_zero_left hLin]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue