PhysLean/HepLean/AnomalyCancellation/PureU1/VectorLike.lean

40 lines
1.1 KiB
Text
Raw Normal View History

2024-04-18 10:09:08 -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:09:08 -04:00
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.PureU1.Sort
import Mathlib.Logic.Equiv.Fin
/-!
# Vector like charges
For the `n`-even case we define the property of a charge assignment being vector like.
-/
2024-07-09 19:22:16 -04:00
/-! TODO: Define vector like ACC in the `n`-odd case. -/
2024-04-18 10:09:08 -04:00
universe v u
open Nat
2024-07-12 11:23:02 -04:00
open Finset
2024-04-18 10:09:08 -04:00
open BigOperators
namespace PureU1
variable {n : }
/--
Given a natural number `n`, this lemma proves that `n + n` is equal to `2 * n`.
-/
lemma split_equal (n : ) : n + n = 2 * n := (Nat.two_mul n).symm
lemma split_odd (n : ) : n + 1 + n = 2 * n + 1 := by omega
/-- A charge configuration for n even is vector like if when sorted the `i`th element
2024-07-12 11:23:02 -04:00
is equal to the negative of the `n + i`th element. -/
2024-04-18 10:09:08 -04:00
@[simp]
2024-07-12 11:23:02 -04:00
def VectorLikeEven (S : (PureU1 (2 * n)).Charges) : Prop :=
∀ (i : Fin n), (sort S) (Fin.cast (split_equal n) (Fin.castAdd n i))
= - (sort S) (Fin.cast (split_equal n) (Fin.natAdd n i))
2024-04-18 10:09:08 -04:00
end PureU1