PhysLean/HepLean/AnomalyCancellation/SMNu/PlusU1/BMinusL.lean
2024-04-18 09:26:45 -04:00

117 lines
3 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
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
simp at i
match i with
| 0 => rfl
| 1 => rfl
| 2 => rfl
| 3 => rfl
quadSol := by
intro i
simp at i
match i with
| 0 => rfl
cubicSol := by rfl
/-- $B - L$ in the $n$-family case. -/
@[simps!]
def BL (n : ) : (PlusU1 n).Sols :=
familyUniversalAF n BL₁
namespace BL
variable {n : }
lemma on_quadBiLin (S : (PlusU1 n).charges) :
quadBiLin ((BL n).val, S) = 1/2 * accYY S + 3/2 * accSU2 S - 2 * accSU3 S := by
erw [familyUniversal_quadBiLin]
rw [accYY_decomp, accSU2_decomp, accSU3_decomp]
simp
ring
lemma on_quadBiLin_AFL (S : (PlusU1 n).LinSols) : quadBiLin ((BL n).val, S.val) = 0 := by
rw [on_quadBiLin]
rw [YYsol S, SU2Sol S, SU3Sol S]
simp
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]
simp
/-- 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)
lemma addQuad_zero (S : (PlusU1 n).QuadSols) (a : ): addQuad S a 0 = a • S := by
simp [addQuad, linearToQuad]
rfl
lemma on_cubeTriLin (S : (PlusU1 n).charges) :
cubeTriLin ((BL n).val, (BL n).val, S) = 9 * accGrav S - 24 * accSU3 S := by
erw [familyUniversal_cubeTriLin']
rw [accGrav_decomp, accSU3_decomp]
simp
ring
lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
cubeTriLin ((BL n).val, (BL n).val, S.val) = 0 := by
rw [on_cubeTriLin]
rw [gravSol S, SU3Sol S]
simp
lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ) :
accCube (a • S.val + b • (BL n).val) =
a ^ 2 * (a * accCube S.val + 3 * b * cubeTriLin (S.val, S.val, (BL n).val)) := by
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]
simp
ring
end BL
end PlusU1
end SMRHN