PhysLean/HepLean/AnomalyCancellation/PureU1/Basic.lean

165 lines
5 KiB
Text
Raw Normal View History

2024-04-18 09:53:05 -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 09:53:05 -04:00
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.Basic
/-!
# Pure U(1) ACC system.
2024-05-21 14:10:53 +02:00
We define the anomaly cancellation conditions for a pure U(1) gauge theory with `n` fermions.
2024-04-18 09:53:05 -04:00
-/
universe v u
open Nat
2024-07-12 11:23:02 -04:00
open Finset
2024-04-18 09:53:05 -04:00
namespace PureU1
open BigOperators
/-- The vector space of charges. -/
@[simps!]
def PureU1Charges (n : ) : ACCSystemCharges := ACCSystemChargesMk n
open BigOperators in
/-- The gravitational anomaly. -/
def accGrav (n : ) : ((PureU1Charges n).Charges →ₗ[] ) where
2024-04-18 09:53:05 -04:00
toFun S := ∑ i : Fin n, S i
map_add' S T := Finset.sum_add_distrib
map_smul' a S := by
2024-10-12 07:57:35 +00:00
simp only [HSMul.hSMul, SMul.smul, eq_ratCast, Rat.cast_eq_id, id_eq]
2024-07-19 17:00:32 -04:00
rw [← Finset.mul_sum]
2024-04-18 09:53:05 -04:00
/-- The symmetric trilinear form used to define the cubic anomaly. -/
@[simps!]
def accCubeTriLinSymm {n : } : TriLinearSymm (PureU1Charges n).Charges := TriLinearSymm.mk₃
2024-07-12 11:23:02 -04:00
(fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i)
(by
intro a S L T
2024-10-12 07:57:35 +00:00
simp only [PureU1Charges_numberCharges, HSMul.hSMul, ACCSystemCharges.chargesModule_smul]
2024-04-18 09:53:05 -04:00
rw [Finset.mul_sum]
apply Fintype.sum_congr
intro i
ring)
(by
intro S L T R
2024-04-18 09:56:51 -04:00
simp only [PureU1Charges_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add]
2024-04-18 09:53:05 -04:00
rw [← Finset.sum_add_distrib]
apply Fintype.sum_congr
intro i
2024-07-12 16:22:06 -04:00
ring)
(by
intro S L T
2024-04-18 09:56:51 -04:00
simp only [PureU1Charges_numberCharges]
2024-04-18 09:53:05 -04:00
apply Fintype.sum_congr
intro i
2024-07-12 16:22:06 -04:00
ring)
(by
intro S L T
2024-04-18 09:56:51 -04:00
simp only [PureU1Charges_numberCharges]
2024-04-18 09:53:05 -04:00
apply Fintype.sum_congr
intro i
2024-07-12 16:22:06 -04:00
ring)
2024-04-18 09:53:05 -04:00
/-- The cubic anomaly equation. -/
@[simp]
2024-07-12 11:23:02 -04:00
def accCube (n : ) : HomogeneousCubic ((PureU1Charges n).Charges) :=
2024-04-18 09:53:05 -04:00
(accCubeTriLinSymm).toCubic
2024-11-27 14:50:14 +00:00
/-- The cubic ACC for the pure-`U(1)` anomaly equations is equal to the sum of the cubed
charges. -/
lemma accCube_explicit (n : ) (S : (PureU1Charges n).Charges) :
2024-04-18 09:53:05 -04:00
accCube n S = ∑ i : Fin n, S i ^ 3:= by
rw [accCube, TriLinearSymm.toCubic]
2024-07-12 11:23:02 -04:00
change accCubeTriLinSymm S S S = _
2024-04-18 09:53:05 -04:00
rw [accCubeTriLinSymm]
simp only [PureU1Charges_numberCharges, TriLinearSymm.mk₃_toFun_apply_apply]
2024-11-27 14:50:14 +00:00
exact Finset.sum_congr rfl fun x _ => Eq.symm (pow_three' (S x))
2024-04-18 09:53:05 -04:00
end PureU1
/-- The ACC system for a pure $U(1)$ gauge theory with $n$ fermions. -/
@[simps!]
def PureU1 (n : ) : ACCSystem where
numberLinear := 1
linearACCs := fun i =>
match i with
| 0 => PureU1.accGrav n
numberQuadratic := 0
quadraticACCs := Fin.elim0
cubicACC := PureU1.accCube n
/-- An equivalence of vector spaces of charges when the number of fermions is equal. -/
2024-07-19 17:00:32 -04:00
def pureU1EqCharges {n m : } (h : n = m) :
2024-07-12 11:23:02 -04:00
(PureU1 n).Charges ≃ₗ[] (PureU1 m).Charges where
2024-04-18 09:53:05 -04:00
toFun f := f ∘ Fin.cast h.symm
invFun f := f ∘ Fin.cast h
map_add' _ _ := rfl
map_smul' _ _:= rfl
left_inv _ := rfl
right_inv _ := rfl
open BigOperators
2024-11-27 14:50:14 +00:00
/-- A solution to the pure U(1) accs satisfies the linear ACCs. -/
lemma pureU1_linear {n : } (S : (PureU1 n).LinSols) :
2024-04-18 09:53:05 -04:00
∑ i, S.val i = 0 := by
have hS := S.linearSol
2024-10-12 07:57:35 +00:00
simp only [succ_eq_add_one, PureU1_numberLinear, PureU1_linearACCs] at hS
2024-04-18 09:53:05 -04:00
exact hS 0
2024-11-27 14:50:14 +00:00
/-- A solution to the pure U(1) accs satisfies the cubic ACCs. -/
lemma pureU1_cube {n : } (S : (PureU1 n).Sols) :
2024-04-18 09:53:05 -04:00
∑ i, (S.val i) ^ 3 = 0 := by
2024-11-27 14:50:14 +00:00
rw [← PureU1.accCube_explicit]
exact S.cubicSol
2024-04-18 09:53:05 -04:00
2024-11-27 14:50:14 +00:00
/-- The last charge of a solution to the linear ACCs is equal to the negation of the sum
of the other charges. -/
2024-04-18 09:53:05 -04:00
lemma pureU1_last {n : } (S : (PureU1 n.succ).LinSols) :
S.val (Fin.last n) = - ∑ i : Fin n, S.val i.castSucc := by
have hS := pureU1_linear S
2024-10-12 07:57:35 +00:00
simp only [succ_eq_add_one, PureU1_numberCharges] at hS
2024-04-18 09:53:05 -04:00
rw [Fin.sum_univ_castSucc] at hS
linear_combination hS
/-- Two solutions to the Linear ACCs for `n.succ` areq equal if their first `n` charges are
2024-11-27 14:50:14 +00:00
equal. -/
2024-04-18 09:53:05 -04:00
lemma pureU1_anomalyFree_ext {n : } {S T : (PureU1 n.succ).LinSols}
(h : ∀ (i : Fin n), S.val i.castSucc = T.val i.castSucc) : S = T := by
apply ACCSystemLinear.LinSols.ext
funext i
2024-11-27 14:50:14 +00:00
rcases Fin.eq_castSucc_or_eq_last i with hi | hi
· obtain ⟨j, hj⟩ := hi
subst hj
exact h j
2024-11-27 14:50:14 +00:00
· rw [hi, pureU1_last, pureU1_last]
exact neg_inj.mpr (Finset.sum_congr rfl fun j _ => h j)
2024-04-18 09:53:05 -04:00
namespace PureU1
2024-11-27 14:50:14 +00:00
/-- The `j`th charge of a sum of pure-`U(1)` charges is equal to the sum of
their `j`th charges. -/
lemma sum_of_charges {n : } (f : Fin k → (PureU1 n).Charges) (j : Fin n) :
2024-04-18 09:53:05 -04:00
(∑ i : Fin k, (f i)) j = ∑ i : Fin k, (f i) j := by
induction k
2024-08-30 10:43:29 -04:00
· rfl
· rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
2024-11-28 11:18:25 +00:00
erw [← hl (f ∘ Fin.castSucc)]
2024-08-30 10:43:29 -04:00
rfl
2024-04-18 09:53:05 -04:00
2024-11-27 14:50:14 +00:00
/-- The `j`th charge of a sum of solutions to the linear ACC is equal to the sum of
their `j`th charges. -/
2024-04-18 09:53:05 -04:00
lemma sum_of_anomaly_free_linear {n : } (f : Fin k → (PureU1 n).LinSols) (j : Fin n) :
(∑ i : Fin k, (f i)).1 j = (∑ i : Fin k, (f i).1 j) := by
induction k
2024-08-30 10:43:29 -04:00
· rfl
· rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
2024-11-28 11:18:25 +00:00
erw [← hl (f ∘ Fin.castSucc)]
2024-08-30 10:43:29 -04:00
rfl
2024-04-18 09:53:05 -04:00
end PureU1