79 lines
1.9 KiB
Text
79 lines
1.9 KiB
Text
/-
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
Released under Apache 2.0 license.
|
|
Authors: Joseph Tooby-Smith
|
|
-/
|
|
import HepLean.AnomalyCancellation.MSSMNu.Basic
|
|
/-!
|
|
# The definition of the solution B₃ and properties thereof
|
|
|
|
We define `B₃` and show that it is a double point of the cubic.
|
|
|
|
# References
|
|
|
|
The main reference for the material in this file is:
|
|
|
|
- https://arxiv.org/pdf/2107.07926.pdf
|
|
|
|
-/
|
|
|
|
universe v u
|
|
|
|
namespace MSSMACC
|
|
open MSSMCharges
|
|
open MSSMACCs
|
|
open BigOperators
|
|
|
|
/-- `B₃` is the charge which is $B-L$ in all families, but with the third
|
|
family of the opposite sign. -/
|
|
def B₃AsCharge : MSSMACC.Charges := toSpecies.symm
|
|
⟨fun s => fun i =>
|
|
match s, i with
|
|
| 0, 0 => 1
|
|
| 0, 1 => 1
|
|
| 0, 2 => - 1
|
|
| 1, 0 => - 1
|
|
| 1, 1 => - 1
|
|
| 1, 2 => 1
|
|
| 2, 0 => - 1
|
|
| 2, 1 => - 1
|
|
| 2, 2 => 1
|
|
| 3, 0 => - 3
|
|
| 3, 1 => - 3
|
|
| 3, 2 => 3
|
|
| 4, 0 => 3
|
|
| 4, 1 => 3
|
|
| 4, 2 => - 3
|
|
| 5, 0 => 3
|
|
| 5, 1 => 3
|
|
| 5, 2 => - 3,
|
|
fun s =>
|
|
match s with
|
|
| 0 => -3
|
|
| 1 => 3⟩
|
|
|
|
/-- `B₃` as a solution. -/
|
|
def B₃ : MSSMACC.Sols :=
|
|
MSSMACC.AnomalyFreeMk B₃AsCharge (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl)
|
|
|
|
lemma B₃_val : B₃.val = B₃AsCharge := by
|
|
rfl
|
|
|
|
lemma doublePoint_B₃_B₃ (R : MSSMACC.LinSols) : cubeTriLin B₃.val B₃.val R.val = 0 := by
|
|
simp only [cubeTriLin, TriLinearSymm.mk₃_toFun_apply_apply, cubeTriLinToFun,
|
|
MSSMSpecies_numberCharges, Fin.isValue, Fin.reduceFinMk]
|
|
rw [Fin.sum_univ_three]
|
|
rw [B₃_val]
|
|
rw [B₃AsCharge]
|
|
repeat rw [toSMSpecies_toSpecies_inv]
|
|
rw [Hd_toSpecies_inv, Hu_toSpecies_inv]
|
|
simp only [mul_one, Fin.isValue, toSMSpecies_apply, one_mul, mul_neg, neg_neg, neg_mul, Hd_apply,
|
|
Fin.reduceFinMk, Hu_apply]
|
|
have hLin := R.linearSol
|
|
simp at hLin
|
|
have h0 := hLin 0
|
|
have h2 := hLin 2
|
|
simp [Fin.sum_univ_three] at h0 h2
|
|
linear_combination 9 * h0 - 24 * h2
|
|
|
|
end MSSMACC
|