chore: Bump to 4.11.0

This commit is contained in:
jstoobysmith 2024-09-04 06:28:46 -04:00
parent b6162217b7
commit 17f09022db
48 changed files with 404 additions and 137 deletions

View file

@ -74,6 +74,12 @@ lemma doublePoint_B₃_B₃ (R : MSSMACC.LinSols) : cubeTriLin B₃.val B₃.val
have h0 := hLin 0
have h2 := hLin 2
simp [Fin.sum_univ_three] at h0 h2
linear_combination 9 * h0 - 24 * h2
have h1 {a b : } (ha : a = 0) (hb : b = 0) : 9 * a - 24 * b = 0 := by
rw [ha, hb]
simp
have h1' := h1 h0 h2
ring_nf at h1'
ring_nf
exact h1'
end MSSMACC

View file

@ -6,7 +6,7 @@ Authors: Joseph Tooby-Smith
import HepLean.AnomalyCancellation.MSSMNu.Basic
import HepLean.AnomalyCancellation.MSSMNu.Y3
import HepLean.AnomalyCancellation.MSSMNu.B3
import Mathlib.Tactic.Polyrith
import Mathlib.Tactic.LinearCombination'
/-!
# The line through B₃ and Y₃
@ -65,21 +65,22 @@ lemma doublePoint_Y₃_B₃ (R : MSSMACC.LinSols) :
cubeTriLin Y₃.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, Y₃_val]
rw [B₃AsCharge, Y₃AsCharge]
rw [Fin.sum_univ_three, B₃_val, Y₃_val, B₃AsCharge, Y₃AsCharge]
repeat rw [toSMSpecies_toSpecies_inv]
rw [Hd_toSpecies_inv, Hu_toSpecies_inv]
rw [Hd_toSpecies_inv, Hu_toSpecies_inv]
rw [Hd_toSpecies_inv, Hu_toSpecies_inv, Hd_toSpecies_inv, Hu_toSpecies_inv]
simp only [mul_one, Fin.isValue, toSMSpecies_apply, one_mul, mul_neg, neg_neg, neg_mul, zero_mul,
add_zero, neg_zero, Hd_apply, Fin.reduceFinMk, Hu_apply]
have hLin := R.linearSol
simp at hLin
simp only [MSSMACC_numberLinear, MSSMACC_linearACCs, Nat.reduceMul, Fin.isValue,
Fin.reduceFinMk] at hLin
have h1 := hLin 1
have h2 := hLin 2
have h3 := hLin 3
simp [Fin.sum_univ_three] at h1 h2 h3
linear_combination -(12 * h2) + 9 * h1 + 3 * h3
simp only [Fin.isValue, Fin.sum_univ_three, Prod.mk_zero_zero, LinearMap.coe_mk, AddHom.coe_mk,
Prod.mk_one_one] at h1 h2 h3
linear_combination (norm := ring_nf) -(12 * h2) + 9 * h1 + 3 * h3
simp only [Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one, add_add_sub_cancel, add_neg_cancel]
lemma lineY₃B₃_doublePoint (R : MSSMACC.LinSols) (a b : ) :
cubeTriLin (lineY₃B₃ a b).val (lineY₃B₃ a b).val R.val = 0 := by

View file

@ -101,7 +101,7 @@ lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : InQuadSolProp T ↔
· rw [quadCoeff, show dot Y₃.val B₃.val = 108 by rfl] at h
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
apply (add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)).mp at h
apply (add_eq_zero_iff_of_nonneg (sq_nonneg _) (sq_nonneg _)).mp at h
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero,
not_false_eq_true, pow_eq_zero_iff] at h
exact h
@ -148,7 +148,7 @@ lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
· rw [cubicCoeff, show dot Y₃.val B₃.val = 108 by rfl] at h
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero, ne_eq,
not_false_eq_true, pow_eq_zero_iff, or_self, false_or] at h
apply (add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)).mp at h
apply (add_eq_zero_iff_of_nonneg (sq_nonneg _) (sq_nonneg _)).mp at h
simp only [Fin.isValue, Fin.reduceFinMk, ne_eq, OfNat.ofNat_ne_zero,
not_false_eq_true, pow_eq_zero_iff] at h
exact h.symm
@ -240,7 +240,7 @@ lemma toSolNS_proj (T : NotInLineEqSol) : toSolNS (toSolNSProj T.val) = T.val :=
ring
rw [h1]
have h1 := (lineEqPropSol_iff_lineEqCoeff_zero T.val).mpr.mt T.prop
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h1]
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h1]
exact MulAction.one_smul T.1.val
/-- A solution to the ACCs, given an element of `inLineEq × × × `. -/
@ -283,7 +283,7 @@ lemma inLineEqToSol_proj (T : InLineEqSol) : inLineEqToSol (inLineEqProj T) = T.
ring
rw [h1]
have h2 := (inQuadSolProp_iff_quadCoeff_zero T.val).mpr.mt T.prop.2
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h2]
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h2]
exact MulAction.one_smul T.1.val
/-- Given an element of `inQuad × × × `, a solution to the ACCs. -/
@ -328,7 +328,7 @@ lemma inQuadToSol_proj (T : InQuadSol) : inQuadToSol (inQuadProj T) = T.val := b
ring
rw [h1]
have h2 := (inCubeSolProp_iff_cubicCoeff_zero T.val).mpr.mt T.prop.2.2
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h2]
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel h2]
exact MulAction.one_smul T.1.val
/-- Given a element of `inQuadCube × × × `, a solution to the ACCs. -/
@ -368,7 +368,7 @@ lemma inQuadCubeToSol_proj (T : InQuadCubeSol) :
rw [Y₃_plus_B₃_plus_proj]
ring_nf
simp only [Fin.isValue, Fin.reduceFinMk, zero_smul, add_zero, zero_add]
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel]
rw [← MulAction.mul_smul, mul_comm, mul_inv_cancel]
· exact MulAction.one_smul (T.1).val
· rw [show dot Y₃.val B₃.val = 108 by rfl]
exact Ne.symm (OfNat.zero_ne_ofNat 108)

View file

@ -74,6 +74,9 @@ lemma doublePoint_Y₃_Y₃ (R : MSSMACC.LinSols) :
simp at hLin
have h3 := hLin 3
simp [Fin.sum_univ_three] at h3
linear_combination 6 * h3
have h5 := congrArg (HMul.hMul 6) h3
ring_nf at h5
ring_nf
exact h5
end MSSMACC

View file

@ -50,6 +50,7 @@ section charges
variable {S : (PureU1 n.succ).Charges} {A : (PureU1 n.succ).LinSols}
variable (hS : ConstAbsSorted S) (hA : ConstAbsSorted A.val)
include hS in
lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := by
have hSS := hS.2 i k hik
have ht := hS.1 i k
@ -58,11 +59,13 @@ lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := b
· exact h
· linarith
include hS in
lemma val_le_zero {i : Fin n.succ} (hi : S i ≤ 0) : S i = S (0 : Fin n.succ) := by
symm
apply lt_eq hS hi
exact Fin.zero_le i
include hS in
lemma gt_eq {k i: Fin n.succ} (hk : 0 ≤ S k) (hik : k ≤ i) : S i = S k := by
have hSS := hS.2 k i hik
have ht := hS.1 i k
@ -71,10 +74,12 @@ lemma gt_eq {k i: Fin n.succ} (hk : 0 ≤ S k) (hik : k ≤ i) : S i = S k := by
· exact h
· linarith
include hS in
lemma zero_gt (h0 : 0 ≤ S (0 : Fin n.succ)) (i : Fin n.succ) : S (0 : Fin n.succ) = S i := by
symm
refine gt_eq hS h0 (Fin.zero_le i)
include hS in
lemma opposite_signs_eq_neg {i j : Fin n.succ} (hi : S i ≤ 0) (hj : 0 ≤ S j) : S i = - S j := by
have hSS := hS.1 i j
rw [sq_eq_sq_iff_eq_or_eq_neg] at hSS
@ -83,6 +88,7 @@ lemma opposite_signs_eq_neg {i j : Fin n.succ} (hi : S i ≤ 0) (hj : 0 ≤ S j)
linarith
· exact h
include hS in
lemma is_zero (h0 : S (0 : Fin n.succ) = 0) : S = 0 := by
funext i
have ht := hS.1 i (0 : Fin n.succ)
@ -96,9 +102,11 @@ is defined as a element of `k ∈ Fin n` such that `S k.castSucc` and `S k.succ`
@[simp]
def Boundary (S : (PureU1 n.succ).Charges) (k : Fin n) : Prop := S k.castSucc < 0 ∧ 0 < S k.succ
include hS in
lemma boundary_castSucc {k : Fin n} (hk : Boundary S k) : S k.castSucc = S (0 : Fin n.succ) :=
(lt_eq hS (le_of_lt hk.left) (Fin.zero_le k.castSucc : 0 ≤ k.castSucc)).symm
include hS in
lemma boundary_succ {k : Fin n} (hk : Boundary S k) : S k.succ = - S (0 : Fin n.succ) := by
have hn := boundary_castSucc hS hk
rw [opposite_signs_eq_neg hS (le_of_lt hk.left) (le_of_lt hk.right)] at hn
@ -115,6 +123,7 @@ lemma boundary_accGrav' (k : Fin n) : accGrav n.succ S =
simp only [Fin.val_succ, mem_univ, RelIso.coe_fn_toEquiv]
· exact fun _ _ => rfl
include hS in
lemma boundary_accGrav'' (k : Fin n) (hk : Boundary S k) :
accGrav n.succ S = (2 * ↑↑k + 1 - ↑n) * S (0 : Fin n.succ) := by
rw [boundary_accGrav' k]
@ -136,6 +145,7 @@ lemma boundary_accGrav'' (k : Fin n) (hk : Boundary S k) :
def HasBoundary (S : (PureU1 n.succ).Charges) : Prop :=
∃ (k : Fin n), Boundary S k
include hS in
lemma not_hasBoundary_zero_le (hnot : ¬ (HasBoundary S)) (h0 : S (0 : Fin n.succ) < 0) :
∀ i, S (0 : Fin n.succ) = S i := by
intro ⟨i, hi⟩
@ -148,6 +158,7 @@ lemma not_hasBoundary_zero_le (hnot : ¬ (HasBoundary S)) (h0 : S (0 : Fin n.suc
erw [← hii] at hnott
exact (val_le_zero hS (hnott h0)).symm
include hS in
lemma not_hasBoundry_zero (hnot : ¬ (HasBoundary S)) (i : Fin n.succ) :
S (0 : Fin n.succ) = S i := by
by_cases hi : S (0 : Fin n.succ) < 0
@ -155,10 +166,12 @@ lemma not_hasBoundry_zero (hnot : ¬ (HasBoundary S)) (i : Fin n.succ) :
· simp at hi
exact zero_gt hS hi i
include hS in
lemma not_hasBoundary_grav (hnot : ¬ (HasBoundary S)) :
accGrav n.succ S = n.succ * S (0 : Fin n.succ) := by
simp [accGrav, ← not_hasBoundry_zero hS hnot]
include hA in
lemma AFL_hasBoundary (h : A.val (0 : Fin n.succ) ≠ 0) : HasBoundary A.val := by
by_contra hn
have h0 := not_hasBoundary_grav hA hn

View file

@ -125,7 +125,7 @@ theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : GenericCase S) :
rw [parameterizationAsLinear_val]
change S.val = _ • (_ • P g + _• P! f)
rw [anomalyFree_param _ _ hS]
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
· exact hS
· have h := h g f hS
rw [anomalyFree_param _ _ hS] at h

View file

@ -126,7 +126,7 @@ theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
rw [parameterizationAsLinear_val]
change S.val = _ • (_ • P g + _• P! f)
rw [anomalyFree_param _ _ hS]
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
· exact hS
· have h := h g f hS
rw [anomalyFree_param _ _ hS] at h

View file

@ -107,7 +107,7 @@ lemma cubic_zero_E'_zero (S : linearParameters) (hc : accCube (S.asCharges) = 0)
rw [h1] at hc
simp at hc
cases' hc with hc hc
· have h2 := (add_eq_zero_iff' (by nlinarith) (sq_nonneg S.Y)).mp hc
· have h2 := (add_eq_zero_iff_of_nonneg (by nlinarith) (sq_nonneg S.Y)).mp hc
simp at h2
exact h2.1
· exact hc

View file

@ -72,7 +72,7 @@ lemma genericToQuad_on_quad (S : (PlusU1 n).QuadSols) :
lemma genericToQuad_neq_zero (S : (PlusU1 n).QuadSols) (h : α₁ C S.1 ≠ 0) :
(α₁ C S.1)⁻¹ • genericToQuad C S.1 = S := by
rw [genericToQuad_on_quad, smul_smul, inv_mul_cancel h, one_smul]
rw [genericToQuad_on_quad, smul_smul, Rat.inv_mul_cancel _ h, one_smul]
/-- The construction of a `QuadSol` from a `LinSols` in the special case when `α₁ C S = 0` and
`α₂ S = 0`. -/

View file

@ -64,7 +64,7 @@ lemma generic_on_AF (S : (PlusU1 n).Sols) : generic S.1 = (α₁ S.1) • S := b
lemma generic_on_AF_α₁_ne_zero (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) :
(α₁ S.1)⁻¹ • generic S.1 = S := by
rw [generic_on_AF, smul_smul, inv_mul_cancel h, one_smul]
rw [generic_on_AF, smul_smul, inv_mul_cancel h, one_smul]
/-- The construction of a `Sol` from a `QuadSol` in the case when `α₁ S = 0` and `α₂ S = 0`. -/
def special (S : (PlusU1 n).QuadSols) (a b : ) (h1 : α₁ S = 0) (h2 : α₂ S = 0) :