refactor: Remove double empty lines

This commit is contained in:
jstoobysmith 2024-07-03 07:56:30 -04:00
parent ae18a2196d
commit f03d063c86
60 changed files with 0 additions and 232 deletions

View file

@ -190,7 +190,6 @@ lemma basis!_on_other {k : Fin n} {j : Fin (2 * n.succ)} (h1 : j ≠ δ!₁ k)
simp [basis!AsCharges]
simp_all only [ne_eq, ↓reduceIte]
lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) :
basis!AsCharges k (δ!₁ j) = 0 := by
simp [basis!AsCharges]
@ -310,7 +309,6 @@ lemma basis!_accCube (j : Fin n) :
simp [basis!_δ!₂_eq_minus_δ!₁]
ring
/-- The first part of the basis as `LinSols`. -/
@[simps!]
def basis (j : Fin n.succ) : (PureU1 (2 * n.succ)).LinSols :=
@ -587,7 +585,6 @@ theorem basis!_linear_independent : LinearIndependent (@basis! n) := by
rw [P!'_val] at h1
exact P!_zero f h1
theorem basisa_linear_independent : LinearIndependent (@basisa n) := by
apply Fintype.linearIndependent_iff.mpr
intro f h
@ -663,7 +660,6 @@ lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n → ) :
rw [← join_ext]
exact Pa'_eq _ _
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) =
FiniteDimensional.finrank (PureU1 (2 * n.succ)).LinSols := by
erw [BasisLinear.finrank_AnomalyFreeLinear]
@ -675,7 +671,6 @@ noncomputable def basisaAsBasis :
Basis (Fin (succ n) ⊕ Fin n) (PureU1 (2 * succ n)).LinSols :=
basisOfLinearIndependentOfCardEqFinrank (@basisa_linear_independent n) basisa_card
lemma span_basis (S : (PureU1 (2 * n.succ)).LinSols) :
∃ (g : Fin n.succ → ) (f : Fin n → ), S.val = P g + P! f := by
have h := (mem_span_range_iff_exists_fun ).mp (Basis.mem_span basisaAsBasis S)