PhysLean/HepLean/AnomalyCancellation/SM/Permutations.lean

115 lines
4.1 KiB
Text
Raw Normal View History

/-
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.
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SM.Basic
2024-06-25 07:06:32 -04:00
import Mathlib.RepresentationTheory.Basic
/-!
2024-11-27 14:35:02 +00:00
# Permutations of SM with no RHN.
We define the group of permutations for the SM charges with no RHN.
-/
universe v u
open Nat
2024-07-12 11:23:02 -04:00
open Finset
namespace SM
open SMCharges
open SMACCs
open BigOperators
/-- The group of `Sₙ` permutations for each species. -/
@[simp]
2024-07-12 11:23:02 -04:00
def PermGroup (n : ) := ∀ (_ : Fin 5), Equiv.Perm (Fin n)
variable {n : }
2024-11-11 10:07:25 +00:00
/-- The type `PermGroup n` inherits the instance of a group from it's
2024-11-12 10:00:23 +00:00
target space `Equiv.Perm`. -/
@[simp]
instance : Group (PermGroup n) := Pi.group
/-- The image of an element of `permGroup n` under the representation on charges. -/
@[simps!]
def chargeMap (f : PermGroup n) : (SMCharges n).Charges →ₗ[] (SMCharges n).Charges where
2024-07-12 16:22:06 -04:00
toFun S := toSpeciesEquiv.symm (fun i => toSpecies i S ∘ f i)
2024-06-13 08:10:08 -04:00
map_add' _ _ := rfl
map_smul' _ _ := rfl
/-- The representation of `(permGroup n)` acting on the vector space of charges. -/
@[simp]
def repCharges {n : } : Representation (PermGroup n) (SMCharges n).Charges where
toFun f := chargeMap f⁻¹
2024-07-12 16:22:06 -04:00
map_mul' f g := by
simp only [PermGroup, mul_inv_rev]
apply LinearMap.ext
intro S
rw [charges_eq_toSpecies_eq]
intro i
simp only [chargeMap_apply, Pi.mul_apply, Pi.inv_apply, Equiv.Perm.coe_mul, LinearMap.mul_apply]
rw [toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
rfl
map_one' := by
apply LinearMap.ext
intro S
rw [charges_eq_toSpecies_eq]
intro i
erw [toSMSpecies_toSpecies_inv]
rfl
2025-02-10 10:51:44 +00:00
/-- The species charges of a set of charges acted on by a family permutation is the permutation
2024-11-27 06:49:56 +00:00
of those species charges with the corresponding part of the family permutation. -/
lemma repCharges_toSpecies (f : PermGroup n) (S : (SMCharges n).Charges) (j : Fin 5) :
toSpecies j (repCharges f S) = toSpecies j S ∘ f⁻¹ j := by
erw [toSMSpecies_toSpecies_inv]
2024-11-27 06:49:56 +00:00
/-- The sum over every charge in any species to some power `m` is invariant under the group
action. -/
lemma toSpecies_sum_invariant (m : ) (f : PermGroup n) (S : (SMCharges n).Charges) (j : Fin 5) :
∑ i, ((fun a => a ^ m) ∘ toSpecies j (repCharges f S)) i =
∑ i, ((fun a => a ^ m) ∘ toSpecies j S) i := by
2024-06-13 08:10:08 -04:00
rw [repCharges_toSpecies]
exact Fintype.sum_equiv (f⁻¹ j) (fun x => ((fun a => a ^ m) ∘ (toSpecies j) S ∘ ⇑(f⁻¹ j)) x)
2024-07-19 17:00:32 -04:00
(fun x => ((fun a => a ^ m) ∘ (toSpecies j) S) x) (congrFun rfl)
2024-06-13 08:10:08 -04:00
2025-02-10 10:51:44 +00:00
/-- The gravitational anomaly equations is invariant under family permutations. -/
2024-07-12 11:23:02 -04:00
lemma accGrav_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
2024-11-27 06:49:56 +00:00
accGrav (repCharges f S) = accGrav S := accGrav_ext
(by simpa using toSpecies_sum_invariant 1 f S)
2024-11-27 06:49:56 +00:00
/-- The `SU(2)` anomaly equation is invariant under family permutations. -/
2024-07-12 11:23:02 -04:00
lemma accSU2_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
2024-11-27 06:49:56 +00:00
accSU2 (repCharges f S) = accSU2 S := accSU2_ext
(by simpa using toSpecies_sum_invariant 1 f S)
2024-11-27 06:49:56 +00:00
/-- The `SU(3)` anomaly equation is invariant under family permutations. -/
2024-07-12 11:23:02 -04:00
lemma accSU3_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accSU3 (repCharges f S) = accSU3 S :=
accSU3_ext
(by simpa using toSpecies_sum_invariant 1 f S)
2024-11-27 06:49:56 +00:00
/-- The `Y²` anomaly equation is invariant under family permutations. -/
2024-07-12 11:23:02 -04:00
lemma accYY_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accYY (repCharges f S) = accYY S :=
accYY_ext
(by simpa using toSpecies_sum_invariant 1 f S)
2024-11-27 06:49:56 +00:00
/-- The quadratic anomaly equation is invariant under family permutations. -/
2024-07-12 11:23:02 -04:00
lemma accQuad_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accQuad (repCharges f S) = accQuad S :=
accQuad_ext
(toSpecies_sum_invariant 2 f S)
2024-11-27 06:49:56 +00:00
/-- The cubic anomaly equation is invariant under family permutations. -/
2024-07-12 11:23:02 -04:00
lemma accCube_invariant (f : PermGroup n) (S : (SMCharges n).Charges) :
accCube (repCharges f S) = accCube S :=
2024-06-13 08:10:08 -04:00
accCube_ext (fun j => toSpecies_sum_invariant 3 f S j)
end SM