docs: Improve docs for basisLinear

This commit is contained in:
jstoobysmith 2024-11-28 11:18:25 +00:00
parent 2ada8df0c6
commit 970b3a7436
4 changed files with 51 additions and 41 deletions

View file

@ -150,8 +150,7 @@ lemma sum_of_charges {n : } (f : Fin k → (PureU1 n).Charges) (j : Fin n) :
· rfl
· rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
have hlt := hl (f ∘ Fin.castSucc)
erw [← hlt]
erw [← hl (f ∘ Fin.castSucc)]
rfl
/-- The `j`th charge of a sum of solutions to the linear ACC is equal to the sum of
@ -162,8 +161,7 @@ lemma sum_of_anomaly_free_linear {n : } (f : Fin k → (PureU1 n).LinSols) (j
· rfl
· rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
have hlt := hl (f ∘ Fin.castSucc)
erw [← hlt]
erw [← hl (f ∘ Fin.castSucc)]
rfl
end PureU1