2024-04-18 09:26:45 -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:26:45 -04:00
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.Basic
|
|
|
|
|
import HepLean.AnomalyCancellation.SMNu.PlusU1.FamilyMaps
|
|
|
|
|
/-!
|
|
|
|
|
# B Minus L in SM with RHN.
|
|
|
|
|
|
|
|
|
|
Relavent definitions for the SM `B-L`.
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
universe v u
|
|
|
|
|
|
|
|
|
|
namespace SMRHN
|
|
|
|
|
namespace PlusU1
|
|
|
|
|
|
|
|
|
|
open SMνCharges
|
|
|
|
|
open SMνACCs
|
|
|
|
|
open BigOperators
|
|
|
|
|
|
|
|
|
|
variable {n : ℕ}
|
|
|
|
|
|
|
|
|
|
/-- $B - L$ in the 1-family case. -/
|
|
|
|
|
@[simps!]
|
|
|
|
|
def BL₁ : (PlusU1 1).Sols where
|
|
|
|
|
val := fun i =>
|
|
|
|
|
match i with
|
|
|
|
|
| (0 : Fin 6) => 1
|
|
|
|
|
| (1 : Fin 6) => -1
|
|
|
|
|
| (2 : Fin 6) => -1
|
|
|
|
|
| (3 : Fin 6) => -3
|
|
|
|
|
| (4 : Fin 6) => 3
|
|
|
|
|
| (5 : Fin 6) => 3
|
|
|
|
|
linearSol := by
|
|
|
|
|
intro i
|
2024-10-12 08:42:20 +00:00
|
|
|
|
simp only [PlusU1_numberLinear] at i
|
2024-04-18 09:26:45 -04:00
|
|
|
|
match i with
|
2024-11-02 08:50:17 +00:00
|
|
|
|
| 0 => with_unfolding_all rfl
|
|
|
|
|
| 1 => with_unfolding_all rfl
|
|
|
|
|
| 2 => with_unfolding_all rfl
|
|
|
|
|
| 3 => with_unfolding_all rfl
|
2024-04-18 09:26:45 -04:00
|
|
|
|
quadSol := by
|
|
|
|
|
intro i
|
2024-10-12 08:42:20 +00:00
|
|
|
|
simp only [PlusU1_numberQuadratic] at i
|
2024-04-18 09:26:45 -04:00
|
|
|
|
match i with
|
2024-11-02 08:50:17 +00:00
|
|
|
|
| 0 => with_unfolding_all rfl
|
|
|
|
|
cubicSol := by with_unfolding_all rfl
|
2024-04-18 09:26:45 -04:00
|
|
|
|
|
|
|
|
|
/-- $B - L$ in the $n$-family case. -/
|
|
|
|
|
@[simps!]
|
|
|
|
|
def BL (n : ℕ) : (PlusU1 n).Sols :=
|
|
|
|
|
familyUniversalAF n BL₁
|
|
|
|
|
|
|
|
|
|
namespace BL
|
|
|
|
|
|
|
|
|
|
variable {n : ℕ}
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma on_quadBiLin (S : (PlusU1 n).Charges) :
|
2024-04-22 08:41:50 -04:00
|
|
|
|
quadBiLin (BL n).val S = 1/2 * accYY S + 3/2 * accSU2 S - 2 * accSU3 S := by
|
2024-04-18 09:26:45 -04:00
|
|
|
|
erw [familyUniversal_quadBiLin]
|
|
|
|
|
rw [accYY_decomp, accSU2_decomp, accSU3_decomp]
|
2024-04-18 09:29:34 -04:00
|
|
|
|
simp only [Fin.isValue, BL₁_val, SMνSpecies_numberCharges, toSpecies_apply, one_mul, mul_neg,
|
|
|
|
|
mul_one, neg_mul, sub_neg_eq_add, one_div]
|
2024-04-18 09:26:45 -04:00
|
|
|
|
ring
|
|
|
|
|
|
2024-04-22 08:41:50 -04:00
|
|
|
|
lemma on_quadBiLin_AFL (S : (PlusU1 n).LinSols) : quadBiLin (BL n).val S.val = 0 := by
|
2024-08-31 18:08:55 +02:00
|
|
|
|
rw [on_quadBiLin, YYsol S, SU2Sol S, SU3Sol S]
|
2024-11-02 08:50:17 +00:00
|
|
|
|
with_unfolding_all rfl
|
2024-04-18 09:26:45 -04:00
|
|
|
|
|
|
|
|
|
lemma add_AFL_quad (S : (PlusU1 n).LinSols) (a b : ℚ) :
|
|
|
|
|
accQuad (a • S.val + b • (BL n).val) = a ^ 2 * accQuad S.val := by
|
|
|
|
|
erw [BiLinearSymm.toHomogeneousQuad_add, quadSol (b • (BL n)).1]
|
|
|
|
|
rw [quadBiLin.map_smul₁, quadBiLin.map_smul₂, quadBiLin.swap, on_quadBiLin_AFL]
|
|
|
|
|
erw [accQuad.map_smul]
|
|
|
|
|
simp
|
|
|
|
|
|
|
|
|
|
lemma add_quad (S : (PlusU1 n).QuadSols) (a b : ℚ) :
|
|
|
|
|
accQuad (a • S.val + b • (BL n).val) = 0 := by
|
|
|
|
|
rw [add_AFL_quad, quadSol S]
|
2024-08-30 10:43:29 -04:00
|
|
|
|
exact Rat.mul_zero (a ^ 2)
|
2024-04-18 09:26:45 -04:00
|
|
|
|
|
|
|
|
|
/-- The `QuadSol` obtained by adding $B-L$ to a `QuadSol`. -/
|
|
|
|
|
def addQuad (S : (PlusU1 n).QuadSols) (a b : ℚ) : (PlusU1 n).QuadSols :=
|
|
|
|
|
linearToQuad (a • S.1 + b • (BL n).1.1) (add_quad S a b)
|
|
|
|
|
|
2024-07-19 17:00:32 -04:00
|
|
|
|
lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ℚ) : addQuad S a 0 = a • S := by
|
2024-08-31 18:08:55 +02:00
|
|
|
|
simp only [addQuad, linearToQuad, zero_smul, add_zero]
|
2024-04-18 09:26:45 -04:00
|
|
|
|
rfl
|
|
|
|
|
|
2024-06-26 11:54:02 -04:00
|
|
|
|
lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
|
2024-04-22 09:48:44 -04:00
|
|
|
|
cubeTriLin (BL n).val (BL n).val S = 9 * accGrav S - 24 * accSU3 S := by
|
2024-04-18 09:26:45 -04:00
|
|
|
|
erw [familyUniversal_cubeTriLin']
|
|
|
|
|
rw [accGrav_decomp, accSU3_decomp]
|
2024-04-18 09:29:34 -04:00
|
|
|
|
simp only [Fin.isValue, BL₁_val, mul_one, SMνSpecies_numberCharges, toSpecies_apply, mul_neg,
|
|
|
|
|
neg_neg, neg_mul]
|
2024-04-18 09:26:45 -04:00
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
|
2024-04-22 09:48:44 -04:00
|
|
|
|
cubeTriLin (BL n).val (BL n).val S.val = 0 := by
|
2024-08-31 18:08:55 +02:00
|
|
|
|
rw [on_cubeTriLin, gravSol S, SU3Sol S]
|
2024-11-02 08:50:17 +00:00
|
|
|
|
with_unfolding_all rfl
|
2024-04-18 09:26:45 -04:00
|
|
|
|
|
|
|
|
|
lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ℚ) :
|
|
|
|
|
accCube (a • S.val + b • (BL n).val) =
|
2024-04-22 09:48:44 -04:00
|
|
|
|
a ^ 2 * (a * accCube S.val + 3 * b * cubeTriLin S.val S.val (BL n).val) := by
|
2024-04-18 09:26:45 -04:00
|
|
|
|
erw [TriLinearSymm.toCubic_add, cubeSol (b • (BL n)), accCube.map_smul]
|
|
|
|
|
repeat rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₃]
|
|
|
|
|
rw [on_cubeTriLin_AFL]
|
2024-04-22 09:48:44 -04:00
|
|
|
|
simp only [HomogeneousCubic, accCube, TriLinearSymm.toCubic_apply, Fin.isValue,
|
2024-04-22 08:41:50 -04:00
|
|
|
|
add_zero, BL_val, mul_zero]
|
2024-04-18 09:26:45 -04:00
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
end BL
|
|
|
|
|
end PlusU1
|
|
|
|
|
end SMRHN
|