refactor: more simp to simp only

This commit is contained in:
jstoobysmith 2024-09-09 06:01:25 -04:00
parent 2c00ebba9b
commit 2792027a7f
9 changed files with 59 additions and 50 deletions

View file

@ -48,7 +48,7 @@ lemma exists_plane_exists_basis {n : } (hE : ExistsPlane n) :
have h3 : ∀ i, g (Sum.inl i) = 0 := hB1 (fun i => (g (Sum.inl i))) h2
rw [h2] at hg
have h4 : ∀ i, - g (Sum.inr i) = 0 := hE1 (fun i => (- g (Sum.inr i))) hg.symm
simp at h4
simp only [neg_eq_zero] at h4
intro i
match i with
| Sum.inl i => exact h3 i
@ -57,7 +57,7 @@ lemma exists_plane_exists_basis {n : } (hE : ExistsPlane n) :
theorem plane_exists_dim_le_7 {n : } (hn : ExistsPlane n) : n ≤ 7 := by
obtain ⟨B, hB⟩ := exists_plane_exists_basis hn
have h1 := LinearIndependent.fintype_card_le_finrank hB
simp at h1
simp only [Fintype.card_sum, Fintype.card_fin] at h1
rw [show FiniteDimensional.finrank (PlusU1 3).Charges = 18 from
FiniteDimensional.finrank_fin_fun ] at h1
exact Nat.le_of_add_le_add_left h1

View file

@ -100,7 +100,8 @@ def quadSolToSolInv {n : } : (PlusU1 n).Sols → (PlusU1 n).QuadSols ×
lemma quadSolToSolInv_1 (S : (PlusU1 n).Sols) :
(quadSolToSolInv S).1 = S.1 := by
simp [quadSolToSolInv]
simp only [quadSolToSolInv, α₁, BL_val, SMνACCs.cubeTriLin_toFun_apply_apply, Fin.isValue,
neg_mul, neg_eq_zero, mul_eq_zero, OfNat.ofNat_ne_zero, false_or]
split <;> rfl
lemma quadSolToSolInv_α₁_α₂_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 = 0) :