docs: More acc documentation

This commit is contained in:
jstoobysmith 2024-11-27 14:35:02 +00:00
parent d582979013
commit 371e402570
3 changed files with 11 additions and 1 deletions

View file

@ -36,16 +36,19 @@ namespace SMNoGrav
variable {n : }
/-- The charges in `(SMNoGrav n).LinSols` satisfy the `SU(2)` anomaly-equation. -/
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
have hS := S.linearSol
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
exact hS 0
/-- The charges in `(SMNoGrav n).LinSols` satisfy the `SU(3)` anomaly-equation. -/
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
have hS := S.linearSol
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
exact hS 1
/-- The charges in `(SMNoGrav n).Sols` satisfy the cubic anomaly-equation. -/
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
exact S.cubicSol