feat: More fixes

This commit is contained in:
jstoobysmith 2024-11-02 08:50:17 +00:00
parent 4df8663cbc
commit c9c9047a0c
31 changed files with 134 additions and 124 deletions

View file

@ -54,7 +54,9 @@ def B₃AsCharge : MSSMACC.Charges := toSpecies.symm
/-- `B₃` as a solution. -/
def B₃ : MSSMACC.Sols :=
MSSMACC.AnomalyFreeMk B₃AsCharge (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl)
MSSMACC.AnomalyFreeMk B₃AsCharge (by with_unfolding_all rfl)
(by with_unfolding_all rfl) (by with_unfolding_all rfl) (by with_unfolding_all rfl)
(by with_unfolding_all rfl) (by with_unfolding_all rfl)
lemma B₃_val : B₃.val = B₃AsCharge := by
rfl

View file

@ -315,7 +315,7 @@ def quadBiLin : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂
Hd_apply, Fin.reduceFinMk, Hu_apply]
congr 1
· rw [Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
simp only [reduceMul, Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one]
ring
· ring)
@ -390,7 +390,7 @@ lemma cubeTriLinToFun_swap1 (S T R : MSSMCharges.Charges) :
Fin.reduceFinMk, Hu_apply]
congr 1
· rw [Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
simp only [reduceMul, Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one]
ring
· ring
@ -400,7 +400,7 @@ lemma cubeTriLinToFun_swap2 (S T R : MSSMCharges.Charges) :
Fin.reduceFinMk, Hu_apply]
congr 1
· rw [Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
simp only [reduceMul, Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one]
ring
· ring
@ -537,14 +537,14 @@ def dot : BiLinearSymm MSSMCharges.Charges := BiLinearSymm.mk₂
simp only [MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue,
ACCSystemCharges.chargesAddCommMonoid_add, map_add, Hd_apply, Fin.reduceFinMk, Hu_apply]
rw [Fin.sum_univ_three, Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
simp only [reduceMul, Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one]
ring)
(by
intro S T
simp only [MSSMSpecies_numberCharges, toSMSpecies_apply, Fin.isValue, Hd_apply, Fin.reduceFinMk,
Hu_apply]
rw [Fin.sum_univ_three, Fin.sum_univ_three]
simp only [Fin.isValue]
simp only [reduceMul, Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one]
ring)
end MSSMACC

View file

@ -48,6 +48,9 @@ def YAsCharge : MSSMACC.Charges := toSpecies.invFun
/-- The hypercharge as an element of `MSSMACC.Sols`. -/
def Y : MSSMACC.Sols :=
MSSMACC.AnomalyFreeMk YAsCharge (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl)
MSSMACC.AnomalyFreeMk YAsCharge
(by with_unfolding_all rfl) (by with_unfolding_all rfl)
(by with_unfolding_all rfl) (by with_unfolding_all rfl) (by with_unfolding_all rfl)
(by with_unfolding_all rfl)
end MSSMACC

View file

@ -41,7 +41,7 @@ lemma lineY₃B₃Charges_quad (a b : ) : accQuad (lineY₃B₃Charges a b).v
simp only [mul_zero, add_zero, Fin.isValue, Fin.reduceFinMk, zero_add,
mul_eq_zero, OfNat.ofNat_ne_zero, false_or]
apply Or.inr ∘ Or.inr
rfl
with_unfolding_all rfl
lemma lineY₃B₃Charges_cubic (a b : ) : accCube (lineY₃B₃Charges a b).val = 0 := by
change accCube (a • Y₃.val + b • B₃.val) = 0
@ -52,8 +52,8 @@ lemma lineY₃B₃Charges_cubic (a b : ) : accCube (lineY₃B₃Charges a b).
rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₃]
rw [cubeTriLin.map_smul₁, cubeTriLin.map_smul₂, cubeTriLin.map_smul₃]
erw [Y₃.cubicSol, B₃.cubicSol]
rw [show cubeTriLin Y₃.val Y₃.val B₃.val = 0 by rfl]
rw [show cubeTriLin B₃.val B₃.val Y₃.val = 0 by rfl]
rw [show cubeTriLin Y₃.val Y₃.val B₃.val = 0 by with_unfolding_all rfl]
rw [show cubeTriLin B₃.val B₃.val Y₃.val = 0 by with_unfolding_all rfl]
simp
/-- The line through $Y_3$ and $B_3$ as `Sols`. -/

View file

@ -42,16 +42,16 @@ def proj (T : MSSMACC.LinSols) : MSSMACC.AnomalyFreePerp :=
change dot _ (_ • Y₃.val + _ • B₃.val + _ • T.val) = 0
rw [dot.map_add₂, dot.map_add₂]
rw [dot.map_smul₂, dot.map_smul₂, dot.map_smul₂]
rw [show dot Y₃.val B₃.val = 108 by rfl]
rw [show dot Y₃.val Y₃.val = 216 by rfl]
rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all rfl]
rw [show dot Y₃.val Y₃.val = 216 by with_unfolding_all rfl]
ring,
by
change dot _ (_ • Y₃.val + _ • B₃.val + _ • T.val) = 0
rw [dot.map_add₂, dot.map_add₂]
rw [dot.map_smul₂, dot.map_smul₂, dot.map_smul₂]
rw [show dot Y₃.val B₃.val = 108 by rfl]
rw [show dot B₃.val Y₃.val = 108 by rfl]
rw [show dot B₃.val B₃.val = 108 by rfl]
rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all rfl]
rw [show dot B₃.val Y₃.val = 108 by with_unfolding_all rfl]
rw [show dot B₃.val B₃.val = 108 by with_unfolding_all rfl]
ring⟩
lemma proj_val (T : MSSMACC.LinSols) :
@ -83,8 +83,8 @@ lemma quad_Y₃_proj (T : MSSMACC.LinSols) :
rw [proj_val]
rw [quadBiLin.map_add₂, quadBiLin.map_add₂]
rw [quadBiLin.map_smul₂, quadBiLin.map_smul₂, quadBiLin.map_smul₂]
rw [show quadBiLin Y₃.val B₃.val = 0 by rfl]
rw [show quadBiLin Y₃.val Y₃.val = 0 by rfl]
rw [show quadBiLin Y₃.val B₃.val = 0 by with_unfolding_all rfl]
rw [show quadBiLin Y₃.val Y₃.val = 0 by with_unfolding_all rfl]
ring
lemma quad_B₃_proj (T : MSSMACC.LinSols) :
@ -92,8 +92,8 @@ lemma quad_B₃_proj (T : MSSMACC.LinSols) :
rw [proj_val]
rw [quadBiLin.map_add₂, quadBiLin.map_add₂]
rw [quadBiLin.map_smul₂, quadBiLin.map_smul₂, quadBiLin.map_smul₂]
rw [show quadBiLin B₃.val Y₃.val = 0 by rfl]
rw [show quadBiLin B₃.val B₃.val = 0 by rfl]
rw [show quadBiLin B₃.val Y₃.val = 0 by with_unfolding_all rfl]
rw [show quadBiLin B₃.val B₃.val = 0 by with_unfolding_all rfl]
ring
lemma quad_self_proj (T : MSSMACC.Sols) :

View file

@ -61,10 +61,10 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ) (hR' : R
rw [dot.map_smul₂, dot.map_smul₂, dot.map_smul₂] at h1 h2
rw [R.perpY₃] at h1
rw [R.perpB₃] at h2
rw [show dot Y₃.val Y₃.val = 216 by rfl] at h1
rw [show dot B₃.val B₃.val = 108 by rfl] at h2
rw [show dot Y₃.val B₃.val = 108 by rfl] at h1
rw [show dot B₃.val Y₃.val = 108 by rfl] at h2
rw [show dot Y₃.val Y₃.val = 216 by with_unfolding_all rfl] at h1
rw [show dot B₃.val B₃.val = 108 by with_unfolding_all rfl] at h2
rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all rfl] at h1
rw [show dot B₃.val Y₃.val = 108 by with_unfolding_all rfl] at h2
simp_all
have ha : a = a' := by
linear_combination h1 / 108 + -1 * h2 / 108

View file

@ -57,7 +57,7 @@ lemma lineEqPropSol_iff_lineEqCoeff_zero (T : MSSMACC.Sols) :
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero, OfNat.ofNat_ne_zero,
false_or]
rw [cube_proj_proj_B₃, cube_proj_proj_Y₃, quad_B₃_proj, quad_Y₃_proj]
rw [show dot Y₃.val B₃.val = 108 by rfl]
rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all rfl]
simp only [Fin.isValue, Fin.reduceFinMk, OfNat.ofNat_ne_zero, false_or]
ring_nf
rw [mul_comm _ 1259712, mul_comm _ 1259712, ← mul_sub]
@ -67,7 +67,7 @@ lemma linEqPropSol_iff_proj_linEqProp (R : MSSMACC.Sols) :
LineEqPropSol R ↔ LineEqProp (proj R.1.1) := by
rw [lineEqPropSol_iff_lineEqCoeff_zero, lineEqCoeff, LineEqProp]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [show dot Y₃.val B₃.val = 108 by rfl] at h
· rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all rfl] at h
simp only [mul_eq_zero, OfNat.ofNat_ne_zero, false_or] at h
rw [α₁_proj, α₂_proj, h]
simp only [neg_zero, Fin.isValue, Fin.reduceFinMk, zero_mul, and_self]
@ -95,8 +95,8 @@ def quadCoeff (T : MSSMACC.Sols) : :=
lemma inQuadSolProp_iff_quadCoeff_zero (T : MSSMACC.Sols) : InQuadSolProp T ↔ quadCoeff T = 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [quadCoeff, h.1, h.2]
rfl
· rw [quadCoeff, show dot Y₃.val B₃.val = 108 by rfl] at h
with_unfolding_all rfl
· rw [quadCoeff, show dot Y₃.val B₃.val = 108 by with_unfolding_all 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_of_nonneg (sq_nonneg _) (sq_nonneg _)).mp at h
@ -112,7 +112,7 @@ lemma inQuadSolProp_iff_proj_inQuadProp (R : MSSMACC.Sols) :
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [h.1, h.2]
simp only [Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
· rw [show dot Y₃.val B₃.val = 108 by rfl] at h
· rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all rfl] at h
simp only [Fin.isValue, Fin.reduceFinMk, mul_eq_zero,
OfNat.ofNat_ne_zero, or_self, false_or] at h
rw [h.2.1, h.2.2]
@ -141,8 +141,8 @@ lemma inCubeSolProp_iff_cubicCoeff_zero (T : MSSMACC.Sols) :
InCubeSolProp T ↔ cubicCoeff T = 0 := by
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [cubicCoeff, h.1, h.2]
rfl
· rw [cubicCoeff, show dot Y₃.val B₃.val = 108 by rfl] at h
with_unfolding_all rfl
· rw [cubicCoeff, show dot Y₃.val B₃.val = 108 by with_unfolding_all 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_of_nonneg (sq_nonneg _) (sq_nonneg _)).mp at h
@ -157,7 +157,7 @@ lemma inCubeSolProp_iff_proj_inCubeProp (R : MSSMACC.Sols) :
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [h.1, h.2]
simp only [Fin.isValue, Fin.reduceFinMk, mul_zero, add_zero, and_self]
· rw [show dot Y₃.val B₃.val = 108 by rfl] at h
· rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all 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
rw [h.2.1, h.2.2]
@ -367,7 +367,7 @@ lemma inQuadCubeToSol_proj (T : InQuadCubeSol) :
simp only [Fin.isValue, Fin.reduceFinMk, zero_smul, add_zero, zero_add]
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]
· rw [show dot Y₃.val B₃.val = 108 by with_unfolding_all rfl]
exact Ne.symm (OfNat.zero_ne_ofNat 108)
/-- A solution from an element of `MSSMACC.AnomalyFreePerp × × × `. We will

View file

@ -54,7 +54,9 @@ def Y₃AsCharge : MSSMACC.Charges := toSpecies.symm
/-- $Y_3$ as a solution. -/
def Y₃ : MSSMACC.Sols :=
MSSMACC.AnomalyFreeMk Y₃AsCharge (by rfl) (by rfl) (by rfl) (by rfl) (by rfl) (by rfl)
MSSMACC.AnomalyFreeMk Y₃AsCharge
(by with_unfolding_all rfl) (by with_unfolding_all rfl) (by with_unfolding_all rfl)
(by with_unfolding_all rfl) (by with_unfolding_all rfl) (by with_unfolding_all rfl)
lemma Y₃_val : Y₃.val = Y₃AsCharge := by
rfl

View file

@ -110,7 +110,7 @@ lemma δ!₃_δ₁0 : @δ!₃ n = δ₁ 0 := rfl
lemma δ!₄_δ₂Last: @δ!₄ n = δ₂ (Fin.last n) := by
rw [Fin.ext_iff]
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ₂, Fin.natAdd_last, Fin.val_last]
omega
@ -169,7 +169,7 @@ lemma basis_on_δ₁_other {k j : Fin n.succ} (h : k ≠ j) :
· rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp only [Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd] at h2
simp only [Fin.coe_cast, Fin.coe_castAdd, Fin.coe_addNat] at h2
omega
· rfl
@ -197,7 +197,7 @@ lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) :
· rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp only [Fin.coe_cast, Fin.coe_natAdd, Fin.coe_castAdd, add_right_inj] at h2
simp only [Fin.coe_cast, Fin.coe_natAdd, Fin.coe_castAdd, Fin.coe_addNat, add_right_inj] at h2
omega
· rfl
@ -254,11 +254,11 @@ lemma basis!_on_δ!₃ (j : Fin n) : basis!AsCharges j δ!₃ = 0 := by
simp only [basis!AsCharges, succ_eq_add_one, PureU1_numberCharges]
split<;> rename_i h
· simp only [δ!₃, succ_eq_add_one, Fin.isValue, δ!₁, Fin.ext_iff, Fin.coe_cast, Fin.coe_castAdd,
Fin.coe_fin_one, Fin.coe_natAdd] at h
Fin.val_eq_zero, Fin.coe_natAdd] at h
omega
· split <;> rename_i h2
· simp only [δ!₃, succ_eq_add_one, Fin.isValue, δ!₂, Fin.ext_iff, Fin.coe_cast,
Fin.coe_castAdd, Fin.coe_fin_one, Fin.coe_natAdd] at h2
Fin.coe_castAdd, Fin.val_eq_zero, Fin.coe_natAdd] at h2
omega
· rfl
@ -266,12 +266,12 @@ lemma basis!_on_δ!₄ (j : Fin n) : basis!AsCharges j δ!₄ = 0 := by
simp only [basis!AsCharges, succ_eq_add_one, PureU1_numberCharges]
split <;> rename_i h
· rw [Fin.ext_iff] at h
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ!₁, Fin.coe_castAdd, add_right_inj] at h
omega
· split <;> rename_i h2
· rw [Fin.ext_iff] at h2
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [succ_eq_add_one, δ!₄, Fin.isValue, Fin.coe_cast, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ!₂, Fin.coe_castAdd, add_right_inj] at h2
omega
· rfl
@ -650,7 +650,7 @@ lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n → ) :
exact Pa'_eq _ _
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) =
FiniteDimensional.finrank (PureU1 (2 * n.succ)).LinSols := by
Module.finrank (PureU1 (2 * n.succ)).LinSols := by
erw [BasisLinear.finrank_AnomalyFreeLinear]
simp only [Fintype.card_sum, Fintype.card_fin, mul_eq]
exact split_odd n

View file

@ -149,7 +149,7 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
· simp [Fin.ext_iff, δ!₂, δ!₁]
· simp [Fin.ext_iff, δ!₂, δ!₄]
· simp only [Nat.succ_eq_add_one, δ!₁, δ!₄, Fin.isValue, ne_eq, Fin.ext_iff, Fin.coe_cast,
Fin.coe_natAdd, Fin.coe_castAdd, Fin.val_last, Fin.coe_fin_one, add_zero, add_right_inj]
Fin.coe_natAdd, Fin.coe_castAdd, Fin.val_last, Fin.val_eq_zero, add_zero, add_right_inj]
omega
· exact fun M => lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)

View file

@ -150,15 +150,15 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ)).Sols}
lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ)).group),
SpecialCase ((FamilyPermutations (2 * n.succ)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ)).solAction.toFun _ _ S M)) :
LineInCubicPerm S.1.1 :=
fun M => special_case_lineInCubic (h M)
theorem special_case {S : (PureU1 (2 * n.succ.succ)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ.succ)).group),
SpecialCase ((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ.succ)).solAction.toFun _ _ S M)) :
∃ (M : (FamilyPermutations (2 * n.succ.succ)).group),
((FamilyPermutations (2 * n.succ.succ)).solAction.toFun S M).1.1
((FamilyPermutations (2 * n.succ.succ)).solAction.toFun _ _ S M).1.1
∈ Submodule.span (Set.range basis) :=
lineInCubicPerm_in_plane S (special_case_lineInCubic_perm h)

View file

@ -166,7 +166,7 @@ lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols}
ConstAbsProp (S.val i, S.val j) := by
refine Prop_two ConstAbsProp Fin.zero_ne_one ?_
intro M
let S' := (FamilyPermutations 4).solAction.toFun S M
let S' := (FamilyPermutations 4).solAction.toFun _ _ S M
have hS' : LineInPlaneCond S'.1.1 :=
(lineInPlaneCond_perm hS M)
exact linesInPlane_four S' hS'

View file

@ -90,7 +90,7 @@ lemma δa₂_δ!₁ (j : Fin n) : δa₂ j = δ!₁ j.castSucc := by
lemma δa₃_δ₃ : @δa₃ n = δ₃ := by
rw [Fin.ext_iff]
simp only [succ_eq_add_one, δa₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd,
Fin.coe_fin_one, add_zero, δ₃]
Fin.val_eq_zero, add_zero, δ₃]
exact Nat.add_comm 1 n
lemma δa₃_δ!₁ : δa₃ = δ!₁ (Fin.last n) := by
@ -265,12 +265,12 @@ lemma basis_on_δ₃ (j : Fin n) : basisAsCharges j δ₃ = 0 := by
simp only [basisAsCharges, PureU1_numberCharges]
split <;> rename_i h
· rw [Fin.ext_iff] at h
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ₁] at h
omega
· split <;> rename_i h2
· rw [Fin.ext_iff] at h2
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.coe_fin_one,
simp only [δ₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_natAdd, Fin.val_eq_zero,
add_zero, δ₂] at h2
omega
· rfl
@ -279,12 +279,12 @@ lemma basis!_on_δ!₃ (j : Fin n) : basis!AsCharges j δ!₃ = 0 := by
simp only [basis!AsCharges, PureU1_numberCharges]
split <;> rename_i h
· rw [Fin.ext_iff] at h
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_fin_one, δ!₁,
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.val_eq_zero, δ!₁,
Fin.coe_natAdd] at h
omega
· split <;> rename_i h2
· rw [Fin.ext_iff] at h2
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.coe_fin_one, δ!₂,
simp only [δ!₃, Fin.isValue, Fin.coe_cast, Fin.coe_castAdd, Fin.val_eq_zero, δ!₂,
Fin.coe_natAdd] at h2
omega
· rfl
@ -643,7 +643,7 @@ lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n.succ → ) :
exact Pa'_eq _ _
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n.succ)) =
FiniteDimensional.finrank (PureU1 (2 * n.succ + 1)).LinSols := by
Module.finrank (PureU1 (2 * n.succ + 1)).LinSols := by
erw [BasisLinear.finrank_AnomalyFreeLinear]
simp only [Fintype.card_sum, Fintype.card_fin]
exact Eq.symm (Nat.two_mul n.succ)

View file

@ -149,7 +149,7 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ + 1)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ + 1)).group),
SpecialCase ((FamilyPermutations (2 * n.succ + 1)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ + 1)).solAction.toFun _ _ S M)) :
LineInCubicPerm S.1.1 := by
intro M
have hM := special_case_lineInCubic (h M)
@ -157,9 +157,11 @@ lemma special_case_lineInCubic_perm {S : (PureU1 (2 * n.succ + 1)).Sols}
theorem special_case {S : (PureU1 (2 * n.succ.succ + 1)).Sols}
(h : ∀ (M : (FamilyPermutations (2 * n.succ.succ + 1)).group),
SpecialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun S M)) :
SpecialCase ((FamilyPermutations (2 * n.succ.succ + 1)).solAction.toFun _ _ S M)) :
S.1.1 = 0 := by
have ht := special_case_lineInCubic_perm h
exact lineInCubicPerm_zero ht
end Odd
end PureU1

View file

@ -38,8 +38,6 @@ def chargeMap {n : } (f : PermGroup n) :
map_add' _ _ := rfl
map_smul' _ _:= rfl
open PureU1Charges in
/-- The representation of `permGroup` acting on the vector space of charges. -/
@[simp]
def permCharges {n : } : Representation (PermGroup n) (PureU1 n).Charges where

View file

@ -227,43 +227,43 @@ lemma B₀_B₀_Bi_cubic {i : Fin 7} :
cubeTriLin (B 0) (B 0) (B i) = 0 := by
change cubeTriLin (B₀) (B₀) (B i) = 0
rw [B₀_cubic]
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
lemma B₁_B₁_Bi_cubic {i : Fin 7} :
cubeTriLin (B 1) (B 1) (B i) = 0 := by
change cubeTriLin (B₁) (B₁) (B i) = 0
rw [B₁_cubic]
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
lemma B₂_B₂_Bi_cubic {i : Fin 7} :
cubeTriLin (B 2) (B 2) (B i) = 0 := by
change cubeTriLin (B₂) (B₂) (B i) = 0
rw [B₂_cubic]
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
lemma B₃_B₃_Bi_cubic {i : Fin 7} :
cubeTriLin (B 3) (B 3) (B i) = 0 := by
change cubeTriLin (B₃) (B₃) (B i) = 0
rw [B₃_cubic]
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
lemma B₄_B₄_Bi_cubic {i : Fin 7} :
cubeTriLin (B 4) (B 4) (B i) = 0 := by
change cubeTriLin (B₄) (B₄) (B i) = 0
rw [B₄_cubic]
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
lemma B₅_B₅_Bi_cubic {i : Fin 7} :
cubeTriLin (B 5) (B 5) (B i) = 0 := by
change cubeTriLin (B₅) (B₅) (B i) = 0
rw [B₅_cubic]
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
lemma B₆_B₆_Bi_cubic {i : Fin 7} :
cubeTriLin (B 6) (B 6) (B i) = 0 := by
change cubeTriLin (B₆) (B₆) (B i) = 0
rw [B₆_cubic]
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
lemma Bi_Bi_Bj_cubic (i j : Fin 7) :
cubeTriLin (B i) (B i) (B j) = 0 := by
@ -297,21 +297,21 @@ lemma B_sum_is_sol (f : Fin 7 → ) : (SM 3).IsSolution (∑ i, f i • B i)
apply Fintype.sum_eq_zero _ fun i ↦ ?_
rw [map_smul]
have h : accGrav (B i) = 0 := by
fin_cases i <;> rfl
fin_cases i <;> with_unfolding_all rfl
rw [h]
exact DistribMulAction.smul_zero (f i))
(by
rw [map_sum]
apply Fintype.sum_eq_zero _ fun i ↦ ?_
rw [map_smul]
have h : accSU2 (B i) = 0 := by fin_cases i <;> rfl
have h : accSU2 (B i) = 0 := by fin_cases i <;> with_unfolding_all rfl
rw [h]
exact DistribMulAction.smul_zero (f i))
(by
rw [map_sum]
apply Fintype.sum_eq_zero _ fun i ↦ ?_
rw [map_smul]
have h : accSU3 (B i) = 0 := by fin_cases i <;> rfl
have h : accSU3 (B i) = 0 := by fin_cases i <;> with_unfolding_all rfl
rw [h]
exact DistribMulAction.smul_zero (f i))
(B_in_accCube f)

View file

@ -37,16 +37,16 @@ def BL₁ : (PlusU1 1).Sols where
intro i
simp only [PlusU1_numberLinear] at i
match i with
| 0 => rfl
| 1 => rfl
| 2 => rfl
| 3 => rfl
| 0 => with_unfolding_all rfl
| 1 => with_unfolding_all rfl
| 2 => with_unfolding_all rfl
| 3 => with_unfolding_all rfl
quadSol := by
intro i
simp only [PlusU1_numberQuadratic] at i
match i with
| 0 => rfl
cubicSol := by rfl
| 0 => with_unfolding_all rfl
cubicSol := by with_unfolding_all rfl
/-- $B - L$ in the $n$-family case. -/
@[simps!]
@ -67,7 +67,7 @@ lemma on_quadBiLin (S : (PlusU1 n).Charges) :
lemma on_quadBiLin_AFL (S : (PlusU1 n).LinSols) : quadBiLin (BL n).val S.val = 0 := by
rw [on_quadBiLin, YYsol S, SU2Sol S, SU3Sol S]
rfl
with_unfolding_all rfl
lemma add_AFL_quad (S : (PlusU1 n).LinSols) (a b : ) :
accQuad (a • S.val + b • (BL n).val) = a ^ 2 * accQuad S.val := by
@ -100,7 +100,7 @@ lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
cubeTriLin (BL n).val (BL n).val S.val = 0 := by
rw [on_cubeTriLin, gravSol S, SU3Sol S]
rfl
with_unfolding_all rfl
lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ) :
accCube (a • S.val + b • (BL n).val) =

View file

@ -57,8 +57,8 @@ theorem plane_exists_dim_le_7 {n : } (hn : ExistsPlane n) : n ≤ 7 := by
obtain ⟨B, hB⟩ := exists_plane_exists_basis hn
have h1 := LinearIndependent.fintype_card_le_finrank hB
simp only [Fintype.card_sum, Fintype.card_fin] at h1
rw [show FiniteDimensional.finrank (PlusU1 3).Charges = 18 from
FiniteDimensional.finrank_fin_fun ] at h1
rw [show Module.finrank (PlusU1 3).Charges = 18 from
Module.finrank_fin_fun ] at h1
exact Nat.le_of_add_le_add_left h1
end PlusU1

View file

@ -35,16 +35,16 @@ def Y₁ : (PlusU1 1).Sols where
intro i
simp only [PlusU1_numberLinear] at i
match i with
| 0 => rfl
| 1 => rfl
| 2 => rfl
| 3 => rfl
| 0 => with_unfolding_all rfl
| 1 => with_unfolding_all rfl
| 2 => with_unfolding_all rfl
| 3 => with_unfolding_all rfl
quadSol := by
intro i
simp only [PlusU1_numberQuadratic] at i
match i with
| 0 => rfl
cubicSol := by rfl
| 0 => with_unfolding_all rfl
cubicSol := by with_unfolding_all rfl
/-- The hypercharge for `n` family. -/
@[simps!]
@ -96,7 +96,7 @@ lemma on_cubeTriLin (S : (PlusU1 n).Charges) :
lemma on_cubeTriLin_AFL (S : (PlusU1 n).LinSols) :
cubeTriLin (Y n).val (Y n).val S.val = 0 := by
rw [on_cubeTriLin, YYsol S]
rfl
with_unfolding_all rfl
lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
cubeTriLin (Y n).val S S = 6 * accQuad S := by
@ -109,7 +109,7 @@ lemma on_cubeTriLin' (S : (PlusU1 n).Charges) :
lemma on_cubeTriLin'_ALQ (S : (PlusU1 n).QuadSols) :
cubeTriLin (Y n).val S.val S.val = 0 := by
rw [on_cubeTriLin', quadSol S]
rfl
with_unfolding_all rfl
lemma add_AFL_cube (S : (PlusU1 n).LinSols) (a b : ) :
accCube (a • S.val + b • (Y n).val) =