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

40 lines
887 B
Text
Raw Normal View History

2024-04-18 10:50:21 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
2024-07-12 16:39:44 -04:00
Released under Apache 2.0 license as described in the file LICENSE.
2024-04-18 10:50:21 -04:00
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 2 fermions
We define an equivalence between `LinSols` and `Sols`.
-/
universe v u
open Nat
2024-07-12 11:23:02 -04:00
open Finset
2024-04-18 10:50:21 -04:00
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]
2024-04-18 10:50:59 -04:00
simp only [Fin.sum_univ_two, Fin.isValue]
2024-06-13 16:49:01 -04:00
rw [show S.val (0 : Fin 2) = - S.val (1 : Fin 2) from eq_neg_of_add_eq_zero_left hLin]
2024-04-18 10:50:21 -04:00
ring⟩
invFun S := S.1.1
left_inv S := rfl
right_inv S := rfl
end Two
end PureU1