refactor: Change case of type and props

This commit is contained in:
jstoobysmith 2024-06-26 11:54:02 -04:00
parent 18b83f582e
commit f7a638d32e
58 changed files with 695 additions and 696 deletions

View file

@ -23,11 +23,11 @@ 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)
def ExistsPlane (n : ) : Prop := ∃ (B : Fin n → (PlusU1 3).Charges),
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
lemma exists_plane_exists_basis {n : } (hE : ExistsPlane n) :
∃ (B : Fin 11 ⊕ Fin n → (PlusU1 3).Charges), LinearIndependent B := by
obtain ⟨E, hE1, hE2⟩ := hE
obtain ⟨B, hB1, hB2⟩ := eleven_dim_plane_of_no_sols_exists
let Y := Sum.elim B E
@ -59,11 +59,11 @@ lemma exists_plane_exists_basis {n : } (hE : existsPlane n) :
| Sum.inr i => exact h4 i
theorem plane_exists_dim_le_7 {n : } (hn : existsPlane n) : n ≤ 7 := by
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
rw [show FiniteDimensional.finrank (PlusU1 3).charges = 18 from
rw [show FiniteDimensional.finrank (PlusU1 3).Charges = 18 from
FiniteDimensional.finrank_fin_fun ] at h1
exact Nat.le_of_add_le_add_left h1