2024-04-17 16:23:40 -04:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
|
|
|
|
import Mathlib.Tactic.Polyrith
|
2024-06-25 07:06:32 -04:00
|
|
|
|
import Mathlib.RepresentationTheory.Basic
|
2024-04-17 16:23:40 -04:00
|
|
|
|
/-!
|
|
|
|
|
# Permutations of MSSM charges and solutions
|
|
|
|
|
|
|
|
|
|
The three family MSSM charges has a family permutation of S₃⁶. This file defines this group
|
|
|
|
|
and its action on the MSSM.
|
|
|
|
|
|
|
|
|
|
-/
|
2024-04-17 16:26:43 -04:00
|
|
|
|
|
|
|
|
|
universe v u
|
|
|
|
|
|
2024-04-17 16:23:40 -04:00
|
|
|
|
open Nat
|
|
|
|
|
open Finset
|
|
|
|
|
|
|
|
|
|
namespace MSSM
|
|
|
|
|
|
|
|
|
|
open MSSMCharges
|
|
|
|
|
open MSSMACCs
|
|
|
|
|
open BigOperators
|
|
|
|
|
|
|
|
|
|
/-- The group of family permutations is `S₃⁶`-/
|
|
|
|
|
@[simp]
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def PermGroup := Fin 6 → Equiv.Perm (Fin 3)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
|
|
|
|
@[simp]
|
2024-06-26 11:54:02 -04:00
|
|
|
|
instance : Group PermGroup := Pi.group
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
|
|
|
|
/-- The image of an element of `permGroup` under the representation on charges. -/
|
|
|
|
|
@[simps!]
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def chargeMap (f : PermGroup) : MSSMCharges.Charges →ₗ[ℚ] MSSMCharges.Charges where
|
2024-04-17 16:23:40 -04:00
|
|
|
|
toFun S := toSpecies.symm (fun i => toSMSpecies i S ∘ f i, Prod.snd (toSpecies S))
|
|
|
|
|
map_add' S T := by
|
|
|
|
|
simp only
|
|
|
|
|
rw [charges_eq_toSpecies_eq]
|
2024-06-13 16:49:01 -04:00
|
|
|
|
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
2024-04-17 16:23:40 -04:00
|
|
|
|
intro i
|
|
|
|
|
rw [(toSMSpecies i).map_add]
|
|
|
|
|
rw [toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
|
|
|
|
|
rfl
|
|
|
|
|
map_smul' a S := by
|
|
|
|
|
simp only
|
|
|
|
|
rw [charges_eq_toSpecies_eq]
|
2024-06-13 16:49:01 -04:00
|
|
|
|
apply And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
2024-04-17 16:23:40 -04:00
|
|
|
|
intro i
|
|
|
|
|
rw [(toSMSpecies i).map_smul, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv]
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma chargeMap_toSpecies (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
toSMSpecies j (chargeMap f S) = toSMSpecies j S ∘ f j := by
|
|
|
|
|
erw [toSMSpecies_toSpecies_inv]
|
|
|
|
|
|
|
|
|
|
/-- The representation of `permGroup` acting on the vector space of charges. -/
|
|
|
|
|
@[simp]
|
2024-06-26 11:54:02 -04:00
|
|
|
|
def repCharges : Representation ℚ PermGroup (MSSMCharges).Charges where
|
2024-04-17 16:23:40 -04:00
|
|
|
|
toFun f := chargeMap f⁻¹
|
|
|
|
|
map_mul' f g :=by
|
2024-06-26 11:54:02 -04:00
|
|
|
|
simp only [PermGroup, mul_inv_rev]
|
2024-04-17 16:23:40 -04:00
|
|
|
|
apply LinearMap.ext
|
|
|
|
|
intro S
|
|
|
|
|
rw [charges_eq_toSpecies_eq]
|
2024-06-13 16:49:01 -04:00
|
|
|
|
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
2024-04-17 16:23:40 -04:00
|
|
|
|
intro i
|
|
|
|
|
simp only [ Pi.mul_apply, Pi.inv_apply, Equiv.Perm.coe_mul, LinearMap.mul_apply]
|
|
|
|
|
rw [chargeMap_toSpecies, chargeMap_toSpecies]
|
|
|
|
|
simp only [Pi.mul_apply, Pi.inv_apply]
|
|
|
|
|
rw [chargeMap_toSpecies]
|
|
|
|
|
rfl
|
|
|
|
|
map_one' := by
|
|
|
|
|
apply LinearMap.ext
|
|
|
|
|
intro S
|
|
|
|
|
rw [charges_eq_toSpecies_eq]
|
2024-06-13 16:49:01 -04:00
|
|
|
|
refine And.intro ?_ $ Prod.mk.inj_iff.mp rfl
|
2024-04-17 16:23:40 -04:00
|
|
|
|
intro i
|
|
|
|
|
erw [toSMSpecies_toSpecies_inv]
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma repCharges_toSMSpecies (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
toSMSpecies j (repCharges f S) = toSMSpecies j S ∘ f⁻¹ j := by
|
|
|
|
|
erw [toSMSpecies_toSpecies_inv]
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma toSpecies_sum_invariant (m : ℕ) (f : PermGroup) (S : MSSMCharges.Charges) (j : Fin 6) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j (repCharges f S)) i =
|
|
|
|
|
∑ i, ((fun a => a ^ m) ∘ toSMSpecies j S) i := by
|
2024-06-13 16:49:01 -04:00
|
|
|
|
rw [repCharges_toSMSpecies]
|
|
|
|
|
exact Equiv.sum_comp (f⁻¹ j) ((fun a => a ^ m) ∘ (toSMSpecies j) S)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma Hd_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
Hd (repCharges f S) = Hd S := rfl
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma Hu_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
Hu (repCharges f S) = Hu S := rfl
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma accGrav_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
accGrav (repCharges f S) = accGrav S :=
|
|
|
|
|
accGrav_ext
|
|
|
|
|
(by simpa using toSpecies_sum_invariant 1 f S)
|
|
|
|
|
(Hd_invariant f S)
|
|
|
|
|
(Hu_invariant f S)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma accSU2_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
accSU2 (repCharges f S) = accSU2 S :=
|
|
|
|
|
accSU2_ext
|
|
|
|
|
(by simpa using toSpecies_sum_invariant 1 f S)
|
|
|
|
|
(Hd_invariant f S)
|
|
|
|
|
(Hu_invariant f S)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma accSU3_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
accSU3 (repCharges f S) = accSU3 S :=
|
|
|
|
|
accSU3_ext
|
|
|
|
|
(by simpa using toSpecies_sum_invariant 1 f S)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma accYY_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
accYY (repCharges f S) = accYY S :=
|
|
|
|
|
accYY_ext
|
|
|
|
|
(by simpa using toSpecies_sum_invariant 1 f S)
|
|
|
|
|
(Hd_invariant f S)
|
|
|
|
|
(Hu_invariant f S)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma accQuad_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
accQuad (repCharges f S) = accQuad S :=
|
|
|
|
|
accQuad_ext
|
|
|
|
|
(toSpecies_sum_invariant 2 f S)
|
|
|
|
|
(Hd_invariant f S)
|
|
|
|
|
(Hu_invariant f S)
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma accCube_invariant (f : PermGroup) (S : MSSMCharges.Charges) :
|
2024-04-17 16:23:40 -04:00
|
|
|
|
accCube (repCharges f S) = accCube S :=
|
|
|
|
|
accCube_ext
|
2024-06-13 16:49:01 -04:00
|
|
|
|
(fun j => toSpecies_sum_invariant 3 f S j)
|
2024-04-17 16:23:40 -04:00
|
|
|
|
(Hd_invariant f S)
|
|
|
|
|
(Hu_invariant f S)
|
|
|
|
|
|
|
|
|
|
end MSSM
|