PhysLean/HepLean/AnomalyCancellation/PureU1/LowDim/One.lean

36 lines
614 B
Text
Raw Normal View History

2024-04-18 10:50:21 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
2024-06-25 07:06:32 -04:00
import HepLean.AnomalyCancellation.PureU1.Basic
2024-04-18 10:50:21 -04:00
/-!
# The Pure U(1) case with 1 fermion
We show that in this case the charge must be zero.
-/
universe v u
open Nat
open Finset
namespace PureU1
variable {n : }
namespace One
theorem solEqZero (S : (PureU1 1).LinSols) : S = 0 := by
apply ACCSystemLinear.LinSols.ext
have hLin := pureU1_linear S
simp at hLin
funext i
simp at i
2024-06-13 16:49:01 -04:00
rw [show i = (0 : Fin 1) from Fin.fin_one_eq_zero i]
2024-04-18 10:50:21 -04:00
exact hLin
end One
end PureU1