refactor: Quad Lin equations
This commit is contained in:
parent
772e78ca77
commit
b5dd319eed
14 changed files with 245 additions and 299 deletions
|
@ -208,21 +208,23 @@ lemma accYY_ext {S T : (SMCharges n).charges}
|
|||
|
||||
/-- The quadratic bilinear map. -/
|
||||
@[simps!]
|
||||
def quadBiLin : BiLinearSymm (SMCharges n).charges where
|
||||
toFun S := ∑ i, (Q S.1 i * Q S.2 i +
|
||||
def quadBiLin : BiLinearSymm (SMCharges n).charges := BiLinearSymm.mk₂
|
||||
(fun S => ∑ i, (Q S.1 i * Q S.2 i +
|
||||
- 2 * (U S.1 i * U S.2 i) +
|
||||
D S.1 i * D S.2 i +
|
||||
(- 1) * (L S.1 i * L S.2 i) +
|
||||
E S.1 i * E S.2 i)
|
||||
map_smul₁' a S T := by
|
||||
E S.1 i * E S.2 i))
|
||||
(by
|
||||
intro a S T
|
||||
simp only
|
||||
rw [Finset.mul_sum]
|
||||
apply Fintype.sum_congr
|
||||
intro i
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
ring
|
||||
map_add₁' S T R := by
|
||||
ring)
|
||||
(by
|
||||
intro S1 S2 T
|
||||
simp only
|
||||
rw [← Finset.sum_add_distrib]
|
||||
apply Fintype.sum_congr
|
||||
|
@ -230,12 +232,13 @@ def quadBiLin : BiLinearSymm (SMCharges n).charges where
|
|||
repeat erw [map_add]
|
||||
simp only [ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply, Fin.isValue, neg_mul,
|
||||
one_mul]
|
||||
ring
|
||||
swap' S T := by
|
||||
ring)
|
||||
(by
|
||||
intro S T
|
||||
simp only [SMSpecies_numberCharges, toSpecies_apply, Fin.isValue, neg_mul, one_mul]
|
||||
apply Fintype.sum_congr
|
||||
intro i
|
||||
ring
|
||||
ring)
|
||||
|
||||
/-- The quadratic anomaly cancellation condition. -/
|
||||
@[simp]
|
||||
|
@ -247,10 +250,10 @@ lemma accQuad_ext {S T : (SMCharges n).charges}
|
|||
(h : ∀ j, ∑ i, ((fun a => a^2) ∘ toSpecies j S) i =
|
||||
∑ i, ((fun a => a^2) ∘ toSpecies j T) i) :
|
||||
accQuad S = accQuad T := by
|
||||
simp only [accQuad, BiLinearSymm.toHomogeneousQuad_toFun, Fin.isValue]
|
||||
simp only [HomogeneousQuadratic, accQuad, BiLinearSymm.toHomogeneousQuad_apply]
|
||||
erw [← quadBiLin.toFun_eq_coe]
|
||||
rw [quadBiLin]
|
||||
simp only
|
||||
simp only [quadBiLin, BiLinearSymm.mk₂, AddHom.toFun_eq_coe, AddHom.coe_mk, LinearMap.coe_mk]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
ring_nf
|
||||
|
@ -301,7 +304,7 @@ lemma accCube_ext {S T : (SMCharges n).charges}
|
|||
(h : ∀ j, ∑ i, ((fun a => a^3) ∘ toSpecies j S) i =
|
||||
∑ i, ((fun a => a^3) ∘ toSpecies j T) i) :
|
||||
accCube S = accCube T := by
|
||||
simp only [accCube, TriLinearSymm.toCubic_toFun, Fin.isValue]
|
||||
simp only [HomogeneousCubic, accCube, TriLinearSymm.toCubic_apply]
|
||||
erw [← cubeTriLin.toFun_eq_coe]
|
||||
rw [cubeTriLin]
|
||||
simp only
|
||||
|
|
|
@ -84,9 +84,7 @@ lemma asLinear_val (S : linearParameters) : S.asLinear.val = S.asCharges := by
|
|||
|
||||
lemma cubic (S : linearParameters) :
|
||||
accCube (S.asCharges) = - 54 * S.Q'^3 - 18 * S.Q' * S.Y ^ 2 + S.E'^3 := by
|
||||
simp only [accCube, TriLinearSymm.toCubic_toFun, Finset.univ_unique,
|
||||
Fin.default_eq_zero, Fin.isValue, asCharges, SMNoGrav_numberCharges, neg_add_rev, neg_mul,
|
||||
Finset.sum_singleton]
|
||||
simp only [HomogeneousCubic, accCube, TriLinearSymm.toCubic_apply, neg_mul]
|
||||
erw [← TriLinearSymm.toFun_eq_coe]
|
||||
rw [cubeTriLin]
|
||||
simp only [SMSpecies_numberCharges, Finset.univ_unique, Fin.default_eq_zero, Fin.isValue,
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue