PhysLean/HepLean/AnomalyCancellation/PureU1/LowDim/Two.lean
2024-06-25 07:06:32 -04:00

39 lines
855 B
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Basic
/-!
# The Pure U(1) case with 2 fermions
We define an equivalence between `LinSols` and `Sols`.
-/
universe v u
open Nat
open Finset
namespace PureU1
variable {n : }
namespace Two
/-- An equivalence between `LinSols` and `Sols`. -/
def equiv : (PureU1 2).LinSols ≃ (PureU1 2).Sols where
toFun S := ⟨⟨S, by intro i; simp at i; exact Fin.elim0 i⟩, by
have hLin := pureU1_linear S
simp at hLin
erw [accCube_explicit]
simp only [Fin.sum_univ_two, Fin.isValue]
rw [show S.val (0 : Fin 2) = - S.val (1 : Fin 2) from eq_neg_of_add_eq_zero_left hLin]
ring⟩
invFun S := S.1.1
left_inv S := rfl
right_inv S := rfl
end Two
end PureU1