chore: Double line linter

This commit is contained in:
jstoobysmith 2024-07-03 07:41:06 -04:00
parent bd9f796110
commit ae18a2196d
8 changed files with 75 additions and 17 deletions

View file

@ -22,7 +22,6 @@ universe v u
open Nat
open BigOperators
/-- The vector space of charges corresponding to the MSSM fermions. -/
@[simps!]
def MSSMCharges : ACCSystemCharges := ACCSystemChargesMk 20
@ -314,7 +313,6 @@ def quadBiLin : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂ (
ring
ring)
/-- The quadratic ACC. -/
@[simp]
def accQuad : HomogeneousQuadratic MSSMCharges.Charges := quadBiLin.toHomogeneousQuad
@ -335,7 +333,6 @@ lemma accQuad_ext {S T : (MSSMCharges).Charges}
repeat rw [h1]
rw [hd, hu]
/-- The function underlying the symmetric trilinear form used to define the cubic ACC. -/
@[simp]
def cubeTriLinToFun
@ -363,7 +360,6 @@ lemma cubeTriLinToFun_map_smul₁ (a : ) (S T R : MSSMCharges.Charges) :
simp only [map_smul, Hd_apply, Fin.reduceFinMk, Fin.isValue, smul_eq_mul, Hu_apply]
ring
lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.Charges) :
cubeTriLinToFun (S + T, R, L) = cubeTriLinToFun (S, R, L) + cubeTriLinToFun (T, R, L) := by
simp only [cubeTriLinToFun]
@ -382,7 +378,6 @@ lemma cubeTriLinToFun_map_add₁ (S T R L : MSSMCharges.Charges) :
rw [Hd.map_add, Hu.map_add]
ring
lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.Charges) :
cubeTriLinToFun (S, T, R) = cubeTriLinToFun (T, S, R) := by
simp only [cubeTriLinToFun, MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply,