docs: More acc documentation

This commit is contained in:
jstoobysmith 2024-11-27 14:35:02 +00:00
parent d582979013
commit 371e402570
3 changed files with 11 additions and 1 deletions

View file

@ -36,16 +36,19 @@ namespace SMNoGrav
variable {n : }
/-- The charges in `(SMNoGrav n).LinSols` satisfy the `SU(2)` anomaly-equation. -/
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
have hS := S.linearSol
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
exact hS 0
/-- The charges in `(SMNoGrav n).LinSols` satisfy the `SU(3)` anomaly-equation. -/
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
have hS := S.linearSol
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
exact hS 1
/-- The charges in `(SMNoGrav n).Sols` satisfy the cubic anomaly-equation. -/
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := by
exact S.cubicSol

View file

@ -25,6 +25,8 @@ open SMCharges
open SMACCs
open BigOperators
/-- For a set of 1 family SM charges satisfying all ACCs except the gravitational,
the charge of `Q` is zero if and only if `E` is zero. -/
lemma E_zero_iff_Q_zero {S : (SMNoGrav 1).Sols} : Q S.val (0 : Fin 1) = 0 ↔
E S.val (0 : Fin 1) = 0 := by
let S' := linearParameters.bijection.symm S.1.1
@ -34,6 +36,8 @@ lemma E_zero_iff_Q_zero {S : (SMNoGrav 1).Sols} : Q S.val (0 : Fin 1) = 0 ↔
rw [← hS'] at hC
exact Iff.intro (fun hQ => S'.cubic_zero_Q'_zero hC hQ) (fun hE => S'.cubic_zero_E'_zero hC hE)
/-- For a set of 1-family SM charges satisfying all ACCs except the gravitational,
if the `Q` charge is zero then the charges satisfy the gravitional ACCs. -/
lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) :
accGrav S.val = 0 := by
rw [accGrav]
@ -48,6 +52,8 @@ lemma accGrav_Q_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) = 0) :
simp_all
linear_combination 3 * h2
/-- For a set of 1-family SM charges satisfying all ACCs except the gravitational,
if the `Q` charge is not zero then the charges satisfy the gravitional ACCs. -/
lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0) :
accGrav S.val = 0 := by
have hE := E_zero_iff_Q_zero.mpr.mt hQ
@ -59,7 +65,7 @@ lemma accGrav_Q_neq_zero {S : (SMNoGrav 1).Sols} (hQ : Q S.val (0 : Fin 1) ≠ 0
rw [← hS'] at hC ⊢
exact S'.grav_of_cubic hC
/-- Any solution to the ACCs without gravity satisfies the gravitational ACC. -/
/-- Any solution to the 1-family ACCs without gravity satisfies the gravitational ACC. -/
theorem accGravSatisfied {S : (SMNoGrav 1).Sols} :
accGrav S.val = 0 := by
by_cases hQ : Q S.val (0 : Fin 1)= 0

View file

@ -7,6 +7,7 @@ import HepLean.AnomalyCancellation.SM.Basic
import Mathlib.Tactic.Polyrith
import Mathlib.RepresentationTheory.Basic
/-!
# Permutations of SM with no RHN.
We define the group of permutations for the SM charges with no RHN.