PhysLean/HepLean/AnomalyCancellation/SMNu/FamilyMaps.lean

228 lines
9.3 KiB
Text
Raw Normal View History

2024-04-18 08:40:46 -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 08:40:46 -04:00
Authors: Joseph Tooby-Smith
-/
import HepLean.AnomalyCancellation.SMNu.Basic
/-!
2024-07-12 11:23:02 -04:00
# Family maps for the Standard Model for RHN ACCs
2024-04-18 08:40:46 -04:00
2024-05-20 00:19:34 +02:00
We define the a series of maps between the charges for different numbers of families.
2024-04-18 08:40:46 -04:00
-/
universe v u
namespace SMRHN
open SMνCharges
open SMνACCs
open BigOperators
/-- Given a map of for a generic species, the corresponding map for charges. -/
@[simps!]
def chargesMapOfSpeciesMap {n m : } (f : (SMνSpecies n).Charges →ₗ[] (SMνSpecies m).Charges) :
(SMνCharges n).Charges →ₗ[] (SMνCharges m).Charges where
2024-04-18 08:40:46 -04:00
toFun S := toSpeciesEquiv.symm (fun i => (LinearMap.comp f (toSpecies i)) S)
map_add' S T := by
rw [charges_eq_toSpecies_eq]
intro i
2024-08-30 22:37:00 +02:00
rw [map_add, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv,
map_add]
2024-04-18 08:40:46 -04:00
map_smul' a S := by
rw [charges_eq_toSpecies_eq]
intro i
2024-08-30 22:37:00 +02:00
rw [map_smul, toSMSpecies_toSpecies_inv, toSMSpecies_toSpecies_inv, map_smul]
2024-04-18 08:40:46 -04:00
rfl
lemma chargesMapOfSpeciesMap_toSpecies {n m : }
(f : (SMνSpecies n).Charges →ₗ[] (SMνSpecies m).Charges)
(S : (SMνCharges n).Charges) (j : Fin 6) :
2024-07-12 11:23:02 -04:00
toSpecies j (chargesMapOfSpeciesMap f S) = (LinearMap.comp f (toSpecies j)) S := by
2024-04-18 08:40:46 -04:00
erw [toSMSpecies_toSpecies_inv]
2024-07-12 11:23:02 -04:00
/-- The projection of the `m`-family charges onto the first `n`-family charges for species. -/
2024-04-18 08:40:46 -04:00
@[simps!]
def speciesFamilyProj {m n : } (h : n ≤ m) :
(SMνSpecies m).Charges →ₗ[] (SMνSpecies n).Charges where
2024-04-18 08:40:46 -04:00
toFun S := S ∘ Fin.castLE h
2024-06-13 16:55:47 -04:00
map_add' _ _ := rfl
map_smul' _ _ := rfl
2024-04-18 08:40:46 -04:00
2024-07-12 11:23:02 -04:00
/-- The projection of the `m`-family charges onto the first `n`-family charges. -/
def familyProjection {m n : } (h : n ≤ m) : (SMνCharges m).Charges →ₗ[] (SMνCharges n).Charges :=
2024-04-18 08:40:46 -04:00
chargesMapOfSpeciesMap (speciesFamilyProj h)
/-- For species, the embedding of the `m`-family charges onto the `n`-family charges, with all
2024-07-12 11:23:02 -04:00
other charges zero. -/
2024-04-18 08:40:46 -04:00
@[simps!]
def speciesEmbed (m n : ) :
(SMνSpecies m).Charges →ₗ[] (SMνSpecies n).Charges where
2024-04-18 08:40:46 -04:00
toFun S := fun i =>
if hi : i.val < m then
S ⟨i, hi⟩
else
0
map_add' S T := by
funext i
2024-04-18 08:47:45 -04:00
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add]
2024-04-18 08:40:46 -04:00
by_cases hi : i.val < m
· erw [dif_pos hi, dif_pos hi, dif_pos hi]
· erw [dif_neg hi, dif_neg hi, dif_neg hi]
rfl
2024-04-18 08:40:46 -04:00
map_smul' a S := by
funext i
simp [HSMul.hSMul]
by_cases hi : i.val < m
· erw [dif_pos hi, dif_pos hi]
· erw [dif_neg hi, dif_neg hi]
exact Eq.symm (Rat.mul_zero a)
2024-06-13 16:55:47 -04:00
2024-04-18 08:40:46 -04:00
/-- The embedding of the `m`-family charges onto the `n`-family charges, with all
2024-07-12 11:23:02 -04:00
other charges zero. -/
def familyEmbedding (m n : ) : (SMνCharges m).Charges →ₗ[] (SMνCharges n).Charges :=
2024-04-18 08:40:46 -04:00
chargesMapOfSpeciesMap (speciesEmbed m n)
2024-05-20 00:19:34 +02:00
/-- For species, the embedding of the `1`-family charges into the `n`-family charges in
2024-04-18 08:40:46 -04:00
a universal manor. -/
@[simps!]
def speciesFamilyUniversial (n : ) :
(SMνSpecies 1).Charges →ₗ[] (SMνSpecies n).Charges where
2024-08-30 22:37:00 +02:00
toFun S _ := S ⟨0, Nat.zero_lt_succ 0⟩
map_add' _ _ := rfl
map_smul' _ _ := rfl
2024-04-18 08:40:46 -04:00
2024-05-20 00:19:34 +02:00
/-- The embedding of the `1`-family charges into the `n`-family charges in
2024-04-18 08:40:46 -04:00
a universal manor. -/
def familyUniversal (n : ) : (SMνCharges 1).Charges →ₗ[] (SMνCharges n).Charges :=
2024-04-18 08:40:46 -04:00
chargesMapOfSpeciesMap (speciesFamilyUniversial n)
lemma toSpecies_familyUniversal {n : } (j : Fin 6) (S : (SMνCharges 1).Charges)
2024-04-18 08:40:46 -04:00
(i : Fin n) : toSpecies j (familyUniversal n S) i = toSpecies j S ⟨0, by simp⟩ := by
erw [chargesMapOfSpeciesMap_toSpecies]
rfl
lemma sum_familyUniversal {n : } (m : ) (S : (SMνCharges 1).Charges) (j : Fin 6) :
2024-04-18 08:40:46 -04:00
∑ i, ((fun a => a ^ m) ∘ toSpecies j (familyUniversal n S)) i =
n * (toSpecies j S ⟨0, by simp⟩) ^ m := by
2024-04-18 08:47:45 -04:00
simp only [SMνSpecies_numberCharges, Function.comp_apply, toSpecies_apply, Fin.zero_eta,
Fin.isValue]
have h1 : (n : ) * (toSpecies j S ⟨0, by simp⟩) ^ m =
∑ _i : Fin n, (toSpecies j S ⟨0, by simp⟩) ^ m := by
2024-04-18 08:40:46 -04:00
rw [Fin.sum_const]
simp
erw [h1]
refine Finset.sum_congr rfl (fun i _ => ?_)
erw [toSpecies_familyUniversal]
2024-04-18 08:40:46 -04:00
2024-07-12 11:23:02 -04:00
lemma sum_familyUniversal_one {n : } (S : (SMνCharges 1).Charges) (j : Fin 6) :
2024-04-18 08:40:46 -04:00
∑ i, toSpecies j (familyUniversal n S) i = n * (toSpecies j S ⟨0, by simp⟩) := by
simpa using @sum_familyUniversal n 1 S j
lemma sum_familyUniversal_two {n : } (S : (SMνCharges 1).Charges)
(T : (SMνCharges n).Charges) (j : Fin 6) :
2024-04-18 08:40:46 -04:00
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i) =
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i := by
2024-04-18 08:47:45 -04:00
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
2024-04-18 08:40:46 -04:00
rw [Finset.mul_sum]
refine Finset.sum_congr rfl (fun i _ => ?_)
erw [toSpecies_familyUniversal]
rfl
2024-04-18 08:40:46 -04:00
2024-07-12 11:23:02 -04:00
lemma sum_familyUniversal_three {n : } (S : (SMνCharges 1).Charges)
(T L : (SMνCharges n).Charges) (j : Fin 6) :
2024-04-18 08:40:46 -04:00
∑ i, (toSpecies j (familyUniversal n S) i * toSpecies j T i * toSpecies j L i) =
(toSpecies j S ⟨0, by simp⟩) * ∑ i, toSpecies j T i * toSpecies j L i := by
2024-04-18 08:47:45 -04:00
simp only [SMνSpecies_numberCharges, toSpecies_apply, Fin.zero_eta, Fin.isValue]
2024-04-18 08:40:46 -04:00
rw [Finset.mul_sum]
apply Finset.sum_congr
2024-08-30 10:43:29 -04:00
· rfl
· intro i _
erw [toSpecies_familyUniversal]
simp only [SMνSpecies_numberCharges, Fin.zero_eta, Fin.isValue, toSpecies_apply]
ring
2024-04-18 08:40:46 -04:00
lemma familyUniversal_accGrav (S : (SMνCharges 1).Charges) :
2024-07-12 11:23:02 -04:00
accGrav (familyUniversal n S) = n * (accGrav S) := by
2024-04-18 08:40:46 -04:00
rw [accGrav_decomp, accGrav_decomp]
repeat rw [sum_familyUniversal_one]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_accSU2 (S : (SMνCharges 1).Charges) :
2024-07-12 11:23:02 -04:00
accSU2 (familyUniversal n S) = n * (accSU2 S) := by
2024-04-18 08:40:46 -04:00
rw [accSU2_decomp, accSU2_decomp]
repeat rw [sum_familyUniversal_one]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_accSU3 (S : (SMνCharges 1).Charges) :
2024-07-12 11:23:02 -04:00
accSU3 (familyUniversal n S) = n * (accSU3 S) := by
2024-04-18 08:40:46 -04:00
rw [accSU3_decomp, accSU3_decomp]
repeat rw [sum_familyUniversal_one]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_accYY (S : (SMνCharges 1).Charges) :
2024-07-12 11:23:02 -04:00
accYY (familyUniversal n S) = n * (accYY S) := by
2024-04-18 08:40:46 -04:00
rw [accYY_decomp, accYY_decomp]
repeat rw [sum_familyUniversal_one]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_quadBiLin (S : (SMνCharges 1).Charges) (T : (SMνCharges n).Charges) :
2024-04-22 08:41:50 -04:00
quadBiLin (familyUniversal n S) T =
2024-04-18 08:40:46 -04:00
S (0 : Fin 6) * ∑ i, Q T i - 2 * S (1 : Fin 6) * ∑ i, U T i + S (2 : Fin 6) *∑ i, D T i -
2024-07-12 11:23:02 -04:00
S (3 : Fin 6) * ∑ i, L T i + S (4 : Fin 6) * ∑ i, E T i := by
2024-04-18 08:40:46 -04:00
rw [quadBiLin_decomp]
repeat rw [sum_familyUniversal_two]
repeat rw [toSpecies_one]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply, add_left_inj, sub_left_inj,
sub_right_inj]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_accQuad (S : (SMνCharges 1).Charges) :
2024-07-12 11:23:02 -04:00
accQuad (familyUniversal n S) = n * (accQuad S) := by
2024-04-18 08:40:46 -04:00
rw [accQuad_decomp]
repeat erw [sum_familyUniversal]
rw [accQuad_decomp]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_cubeTriLin (S : (SMνCharges 1).Charges) (T R : (SMνCharges n).Charges) :
cubeTriLin (familyUniversal n S) T R = 6 * S (0 : Fin 6) * ∑ i, (Q T i * Q R i) +
2024-07-12 11:23:02 -04:00
3 * S (1 : Fin 6) * ∑ i, (U T i * U R i) + 3 * S (2 : Fin 6) * ∑ i, (D T i * D R i)
2024-04-18 08:40:46 -04:00
+ 2 * S (3 : Fin 6) * ∑ i, (L T i * L R i) +
S (4 : Fin 6) * ∑ i, (E T i * E R i) + S (5 : Fin 6) * ∑ i, (N T i * N R i) := by
rw [cubeTriLin_decomp]
repeat rw [sum_familyUniversal_three]
repeat rw [toSpecies_one]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply, add_left_inj]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_cubeTriLin' (S T : (SMνCharges 1).Charges) (R : (SMνCharges n).Charges) :
cubeTriLin (familyUniversal n S) (familyUniversal n T) R =
2024-04-18 08:40:46 -04:00
6 * S (0 : Fin 6) * T (0 : Fin 6) * ∑ i, Q R i +
2024-07-19 17:00:32 -04:00
3 * S (1 : Fin 6) * T (1 : Fin 6) * ∑ i, U R i
2024-04-18 08:40:46 -04:00
+ 3 * S (2 : Fin 6) * T (2 : Fin 6) * ∑ i, D R i +
2 * S (3 : Fin 6) * T (3 : Fin 6) * ∑ i, L R i +
S (4 : Fin 6) * T (4 : Fin 6) * ∑ i, E R i + S (5 : Fin 6) * T (5 : Fin 6) * ∑ i, N R i := by
rw [familyUniversal_cubeTriLin]
repeat rw [sum_familyUniversal_two]
repeat rw [toSpecies_one]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, toSpecies_apply]
2024-04-18 08:40:46 -04:00
ring
lemma familyUniversal_accCube (S : (SMνCharges 1).Charges) :
2024-04-18 08:40:46 -04:00
accCube (familyUniversal n S) = n * (accCube S) := by
rw [accCube_decomp]
repeat erw [sum_familyUniversal]
rw [accCube_decomp]
2024-04-18 08:47:45 -04:00
simp only [Fin.isValue, SMνSpecies_numberCharges, Fin.zero_eta, toSpecies_apply,
Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton]
2024-04-18 08:40:46 -04:00
ring
end SMRHN