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) =

View file

@ -124,8 +124,11 @@ lemma left_eq_right : P.toFun Φ1 Φ1 =
neg_zero, zero_div, zero_mul, Complex.mul_im, add_zero, mul_neg, neg_mul, sub_neg_eq_add,
sub_add_add_cancel, zero_add, HiggsField.Potential.toFun, neg_add_rev, neg_sub]
ring_nf
erw [show ((Complex.ofReal ‖Φ1 x‖) ^ 4).re = ‖Φ1 x‖ ^ 4 by
erw [show ((Complex.ofRealHom ‖Φ1 x‖) ^ 4).re = ‖Φ1 x‖ ^ 4 by
erw [← Complex.ofReal_pow]; rfl]
erw [show ((Complex.ofRealHom ‖Φ1 x‖) ^ 2).re = ‖Φ1 x‖ ^ 2 by
erw [← Complex.ofReal_pow]; rfl]
erw [show (Complex.ofRealHom ‖Φ1 x‖ ^ 2).im = 0 by exact normSq_apply_im_zero Φ1 x]
ring
lemma left_eq_neg_right : P.toFun Φ1 (- Φ1) =

View file

@ -84,7 +84,7 @@ lemma predAboveI_lt {i x : Fin n.succ.succ} (h : x.val < i.val) :
lemma predAboveI_ge {i x : Fin n.succ.succ} (h : i.val < x.val) :
predAboveI i x = ⟨x.val - 1, by omega⟩ := by
simp only [Nat.succ_eq_add_one, predAboveI, Fin.val_fin_lt, dite_eq_else, Fin.mk.injEq]
simp only [Nat.succ_eq_add_one, predAboveI, Fin.val_fin_lt, dite_eq_right_iff, Fin.mk.injEq]
omega
lemma succAbove_succAbove_predAboveI (i : Fin n.succ.succ) (j : Fin n.succ) (x : Fin n) :

View file

@ -94,25 +94,25 @@ def complexCoBasisFin4 : Basis (Fin 4) complexCo :=
/-- The semilinear map including real Lorentz vectors into complex contravariant
lorentz vectors. -/
def inclCongrRealLorentz : LorentzVector 3 →ₛₗ[Complex.ofReal] complexContr where
def inclCongrRealLorentz : LorentzVector 3 →ₛₗ[Complex.ofRealHom] complexContr where
toFun v := {val := ofReal ∘ v}
map_add' x y := by
apply Lorentz.ContrModule.ext
rw [Lorentz.ContrModule.val_add]
funext i
simp only [Function.comp_apply, ofReal_eq_coe, Pi.add_apply]
simp only [Function.comp_apply, ofRealHom_eq_coe, Pi.add_apply]
change ofReal (x i + y i) = _
simp only [ofReal_eq_coe, ofReal_add]
simp only [ofRealHom_eq_coe, ofReal_add]
map_smul' c x := by
apply Lorentz.ContrModule.ext
rw [Lorentz.ContrModule.val_smul]
funext i
simp only [Function.comp_apply, ofReal_eq_coe, Pi.smul_apply]
simp only [Function.comp_apply, ofRealHom_eq_coe, Pi.smul_apply]
change ofReal (c • x i) = _
simp only [smul_eq_mul, ofReal_eq_coe, ofReal_mul]
simp only [smul_eq_mul, ofRealHom_eq_coe, ofReal_mul]
lemma inclCongrRealLorentz_val (v : LorentzVector 3) :
(inclCongrRealLorentz v).val = ofReal ∘ v := rfl
(inclCongrRealLorentz v).val = ofRealHom ∘ v := rfl
lemma complexContrBasis_of_real (i : Fin 1 ⊕ Fin 3) :
(complexContrBasis i) = inclCongrRealLorentz (LorentzVector.stdBasis i) := by
@ -120,10 +120,10 @@ lemma complexContrBasis_of_real (i : Fin 1 ⊕ Fin 3) :
simp only [complexContrBasis, Basis.coe_ofEquivFun, inclCongrRealLorentz, LorentzVector.stdBasis,
LinearMap.coe_mk, AddHom.coe_mk]
ext j
simp only [Function.comp_apply, ofReal_eq_coe]
simp only [Function.comp_apply, ofRealHom_eq_coe]
erw [Pi.basisFun_apply]
change (Pi.single i 1) j = _
exact Eq.symm (Pi.apply_single (fun _ => ofReal') (congrFun rfl) i 1 j)
exact Eq.symm (Pi.apply_single (fun _ => ofRealHom) (congrFun rfl) i 1 j)
lemma inclCongrRealLorentz_ρ (M : SL(2, )) (v : LorentzVector 3) :
(complexContr.ρ M) (inclCongrRealLorentz v) =
@ -143,7 +143,7 @@ lemma SL2CRep_ρ_basis (M : SL(2, )) (i : Fin 1 ⊕ Fin 3) :
rw [complexContrBasis_of_real, inclCongrRealLorentz_ρ, SL2C.repLorentzVector_stdBasis, map_sum]
apply congrArg
funext j
simp only [LinearMap.map_smulₛₗ, ofReal_eq_coe, coe_smul]
simp only [LinearMap.map_smulₛₗ, ofRealHom_eq_coe, coe_smul]
rw [complexContrBasis_of_real]
end Lorentz

View file

@ -24,7 +24,7 @@ namespace Lorentz
/-- The metric `ηᵃᵃ` as an element of `(complexContr ⊗ complexContr).V`. -/
def contrMetricVal : (complexContr ⊗ complexContr).V :=
contrContrToMatrix.symm ((@minkowskiMatrix 3).map ofReal)
contrContrToMatrix.symm ((@minkowskiMatrix 3).map ofRealHom)
/-- The expansion of `contrMetricVal` into basis vectors. -/
lemma contrMetricVal_expand_tmul : contrMetricVal =
@ -34,7 +34,7 @@ lemma contrMetricVal_expand_tmul : contrMetricVal =
- complexContrBasis (Sum.inr 2) ⊗ₜ[] complexContrBasis (Sum.inr 2) := by
simp only [Action.instMonoidalCategory_tensorObj_V, contrMetricVal, Fin.isValue]
erw [contrContrToMatrix_symm_expand_tmul]
simp only [map_apply, ofReal_eq_coe, coe_smul, Fintype.sum_sum_type, Finset.univ_unique,
simp only [map_apply, ofRealHom_eq_coe, coe_smul, Fintype.sum_sum_type, Finset.univ_unique,
Fin.default_eq_zero, Fin.isValue, Finset.sum_singleton, Fin.sum_univ_three, ne_eq, reduceCtorEq,
not_false_eq_true, minkowskiMatrix.off_diag_zero, zero_smul, add_zero, zero_add, Sum.inr.injEq,
zero_ne_one, Fin.reduceEq, one_ne_zero]
@ -77,7 +77,7 @@ lemma contrMetric_apply_one : contrMetric.hom (1 : ) = contrMetricVal := by
/-- The metric `ηᵢᵢ` as an element of `(complexCo ⊗ complexCo).V`. -/
def coMetricVal : (complexCo ⊗ complexCo).V :=
coCoToMatrix.symm ((@minkowskiMatrix 3).map ofReal)
coCoToMatrix.symm ((@minkowskiMatrix 3).map ofRealHom)
/-- The expansion of `coMetricVal` into basis vectors. -/
lemma coMetricVal_expand_tmul : coMetricVal =
@ -87,7 +87,7 @@ lemma coMetricVal_expand_tmul : coMetricVal =
- complexCoBasis (Sum.inr 2) ⊗ₜ[] complexCoBasis (Sum.inr 2) := by
simp only [Action.instMonoidalCategory_tensorObj_V, coMetricVal, Fin.isValue]
erw [coCoToMatrix_symm_expand_tmul]
simp only [map_apply, ofReal_eq_coe, coe_smul, Fintype.sum_sum_type, Finset.univ_unique,
simp only [map_apply, ofRealHom_eq_coe, coe_smul, Fintype.sum_sum_type, Finset.univ_unique,
Fin.default_eq_zero, Fin.isValue, Finset.sum_singleton, Fin.sum_univ_three, ne_eq, reduceCtorEq,
not_false_eq_true, minkowskiMatrix.off_diag_zero, zero_smul, add_zero, zero_add, Sum.inr.injEq,
zero_ne_one, Fin.reduceEq, one_ne_zero]

View file

@ -137,11 +137,11 @@ lemma selfAdjoint_ext {A B : selfAdjoint (Matrix (Fin 2) (Fin 2) )}
(h2 : ((Matrix.trace (PauliMatrix.σ2 * A.1))).re = ((Matrix.trace (PauliMatrix.σ2 * B.1))).re)
(h3 : ((Matrix.trace (PauliMatrix.σ3 * A.1))).re = ((Matrix.trace (PauliMatrix.σ3 * B.1))).re) :
A = B := by
have h0' := congrArg ofReal h0
have h1' := congrArg ofReal h1
have h2' := congrArg ofReal h2
have h3' := congrArg ofReal h3
rw [ofReal_eq_coe, ofReal_eq_coe] at h0' h1' h2' h3'
have h0' := congrArg ofRealHom h0
have h1' := congrArg ofRealHom h1
have h2' := congrArg ofRealHom h2
have h3' := congrArg ofRealHom h3
rw [ofRealHom_eq_coe, ofRealHom_eq_coe] at h0' h1' h2' h3'
rw [selfAdjoint_trace_σ0_real A, selfAdjoint_trace_σ0_real B] at h0'
rw [selfAdjoint_trace_σ1_real A, selfAdjoint_trace_σ1_real B] at h1'
rw [selfAdjoint_trace_σ2_real A, selfAdjoint_trace_σ2_real B] at h2'

View file

@ -88,7 +88,7 @@ We also define the Higgs bundle.
abbrev HiggsBundle := Bundle.Trivial SpaceTime HiggsVec
instance : SmoothVectorBundle HiggsVec HiggsBundle SpaceTime.asSmoothManifold :=
Bundle.Trivial.smoothVectorBundle HiggsVec 𝓘(, SpaceTime)
Bundle.Trivial.smoothVectorBundle HiggsVec
/-- A Higgs field is a smooth section of the Higgs bundle. -/
abbrev HiggsField : Type := SmoothSection SpaceTime.asSmoothManifold HiggsVec HiggsBundle

View file

@ -135,10 +135,10 @@ def rotateGuageGroup {φ : HiggsVec} (hφ : φ ≠ 0) : GaugeGroupI :=
⟨1, ⟨(rotateMatrix φ), rotateMatrix_specialUnitary hφ⟩, 1⟩
lemma rotateGuageGroup_apply {φ : HiggsVec} (hφ : φ ≠ 0) :
rep (rotateGuageGroup hφ) φ = ![0, Complex.ofReal ‖φ‖] := by
rep (rotateGuageGroup hφ) φ = ![0, Complex.ofRealHom ‖φ‖] := by
rw [rep_apply]
simp only [rotateGuageGroup, rotateMatrix, one_pow, one_smul,
Nat.succ_eq_add_one, Nat.reduceAdd, ofReal_eq_coe]
Nat.succ_eq_add_one, Nat.reduceAdd, ofRealHom_eq_coe]
ext i
fin_cases i
· simp only [mulVec, Fin.zero_eta, Fin.isValue, cons_val', empty_val', cons_val_fin_one,
@ -181,7 +181,7 @@ theorem rotate_fst_real_snd_zero (φ : HiggsVec) :
rw [rep.map_mul]
change rep P (rep g φ) = _
rw [h, rep_apply]
simp only [one_pow, Nat.succ_eq_add_one, Nat.reduceAdd, ofReal_eq_coe, mulVec_cons, zero_smul,
simp only [one_pow, Nat.succ_eq_add_one, Nat.reduceAdd, ofRealHom_eq_coe, mulVec_cons, zero_smul,
coe_smul, mulVec_empty, add_zero, zero_add, one_smul]
funext i
fin_cases i

View file

@ -87,8 +87,7 @@ lemma innerProd_expand (φ1 φ2 : HiggsField) :
nth_rewrite 1 [← RCLike.re_add_im (φ2 x 0)]
nth_rewrite 1 [← RCLike.re_add_im (φ2 x 1)]
ring_nf
repeat rw [show RCLike.ofReal _ = ofReal' _ by rfl]
simp only [Algebra.id.map_eq_id, RCLike.re_to_complex, RingHom.id_apply, RCLike.I_to_complex,
simp only [Fin.isValue, RCLike.re_to_complex, coe_algebraMap, RCLike.I_to_complex,
RCLike.im_to_complex, I_sq, mul_neg, mul_one, neg_mul, sub_neg_eq_add, one_mul]
ring
@ -144,16 +143,15 @@ lemma normSq_eq_innerProd_self (φ : HiggsField) (x : SpaceTime) :
@[simp]
lemma normSq_apply_im_zero (φ : HiggsField) (x : SpaceTime) :
((Complex.ofReal' ‖φ x‖) ^ 2).im = 0 := by
((Complex.ofRealHom ‖φ x‖) ^ 2).im = 0 := by
rw [sq]
simp only [Complex.ofReal_eq_coe, Complex.mul_im, Complex.ofReal_re, Complex.ofReal_im,
mul_zero, zero_mul, add_zero]
simp only [ofRealHom_eq_coe, mul_im, ofReal_re, ofReal_im, mul_zero, zero_mul, add_zero]
@[simp]
lemma normSq_apply_re_self (φ : HiggsField) (x : SpaceTime) :
((Complex.ofReal' ‖φ x‖) ^ 2).re = φ.normSq x := by
((Complex.ofRealHom ‖φ x‖) ^ 2).re = φ.normSq x := by
rw [sq]
simp only [mul_re, ofReal_re, ofReal_im, mul_zero, sub_zero, normSq]
simp only [ofRealHom_eq_coe, mul_re, ofReal_re, ofReal_im, mul_zero, sub_zero, normSq]
exact Eq.symm (pow_two ‖φ x‖)
lemma toHiggsVec_norm (φ : HiggsField) (x : SpaceTime) :

View file

@ -126,7 +126,7 @@ def quadDiscrim (φ : HiggsField) (x : SpaceTime) : := discrim P.𝓵 (- P.
lemma quadDiscrim_nonneg (h : P.𝓵 ≠ 0) (φ : HiggsField) (x : SpaceTime) :
0 ≤ P.quadDiscrim φ x := by
have h1 := P.as_quad φ x
rw [quadratic_eq_zero_iff_discrim_eq_sq] at h1
rw [mul_assoc, quadratic_eq_zero_iff_discrim_eq_sq] at h1
· simp only [h1, ne_eq, quadDiscrim, div_eq_zero_iff, OfNat.ofNat_ne_zero, or_false]
exact sq_nonneg (2 * P.𝓵 * ‖φ‖_H ^ 2 x + - P.μ2)
· exact h
@ -148,7 +148,7 @@ lemma quadDiscrim_eq_zero_iff_normSq (h : P.𝓵 ≠ 0) (φ : HiggsField) (x : S
rw [P.quadDiscrim_eq_zero_iff h]
refine Iff.intro (fun hV => ?_) (fun hF => ?_)
· have h1 := P.as_quad φ x
rw [quadratic_eq_zero_iff_of_discrim_eq_zero h
rw [mul_assoc, quadratic_eq_zero_iff_of_discrim_eq_zero h
((P.quadDiscrim_eq_zero_iff h φ x).mpr hV)] at h1
simp_rw [h1, neg_neg]
· rw [toFun, hF]
@ -241,6 +241,7 @@ lemma neg_𝓵_sol_exists_iff (h𝓵 : P.𝓵 < 0) (c : ) : (∃ φ x, P.toFu
have hdd : discrim P.𝓵 (- P.μ2) (-c) = Real.sqrt (discrim P.𝓵 (- P.μ2) (-c))
* Real.sqrt (discrim P.𝓵 (- P.μ2) (-c)) := by
exact (Real.mul_self_sqrt hd).symm
rw [mul_assoc]
refine (quadratic_eq_zero_iff (ne_of_gt h𝓵).symm hdd _).mpr ?_
simp only [neg_neg, or_true, a]

View file

@ -98,7 +98,7 @@ lemma contrBasisVectorMul_neg {n : } {c : Fin n.succ.succ → complexLorentzT
(h : ¬ (b i).val = (b (i.succAbove j)).val) :
contrBasisVectorMul i j b = 0 := by
rw [contrBasisVectorMul]
simp only [ite_eq_else, one_ne_zero, imp_false]
simp only [ite_eq_right_iff, one_ne_zero, imp_false]
exact h
lemma contrBasisVectorMul_pos {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
@ -106,7 +106,7 @@ lemma contrBasisVectorMul_pos {n : } {c : Fin n.succ.succ → complexLorentzT
(h : (b i).val = (b (i.succAbove j)).val) :
contrBasisVectorMul i j b = 1 := by
rw [contrBasisVectorMul]
simp only [ite_eq_then, zero_ne_one, imp_false, Decidable.not_not]
simp only [ite_eq_left_iff, zero_ne_one, imp_false, Decidable.not_not]
exact h
lemma contr_basisVector {n : } {c : Fin n.succ.succ → complexLorentzTensor.C}
@ -217,8 +217,8 @@ lemma eval_basisVector {n : } {c : Fin n.succ → complexLorentzTensor.C}
Function.comp_apply, one_smul, _root_.zero_smul]
erw [TensorSpecies.evalMap_tprod]
congr 1
have h1 : Fin.ofNat' ↑j (@Fin.size_pos' (complexLorentzTensor.repDim (c i)) _) = j := by
simpa [Fin.ext_iff] using Nat.mod_eq_of_lt j.prop
have h1 : Fin.ofNat' _ ↑j = j := by
simp [Fin.ext_iff]
rw [Basis.repr_self, Finsupp.single_apply, h1]
exact ite_congr (Eq.propIntro (fun a => id (Eq.symm a)) fun a => id (Eq.symm a))
(congrFun rfl) (congrFun rfl)

View file

@ -80,7 +80,8 @@ lemma pairIsoSep_tmul {c1 c2 : C} (x : F.obj (Discrete.mk c1)) (y : F.obj (Discr
LaxMonoidalFunctor.μ_natural_assoc, Action.comp_hom, Action.instMonoidalCategory_tensorHom_hom,
Action.mkIso_inv_hom, LinearEquiv.toModuleIso_inv, Equivalence.symm_inverse,
Action.functorCategoryEquivalence_functor, Action.FunctorCategoryEquivalence.functor_obj_obj,
ModuleCat.coe_comp, Function.comp_apply, ModuleCat.MonoidalCategory.hom_apply, Functor.id_obj]
ModuleCat.coe_comp, Function.comp_apply, ModuleCat.MonoidalCategory.tensorHom_tmul,
Functor.id_obj]
erw [forgetLiftAppV_symm_apply F c1, forgetLiftAppV_symm_apply F c2]
change ((lift.obj F).map fin2Iso.inv).hom
(((lift.obj F).map ((mkIso _).hom ⊗ (mkIso _).hom)).hom

View file

@ -682,7 +682,7 @@ def tensor : ∀ {n : } {c : Fin n → S.C}, TensorTree S c → S.F.obj (Over
| prod t1 t2 => (S.F.map (OverColor.equivToIso finSumFinEquiv).hom).hom
((S.F.μ _ _).hom (t1.tensor ⊗ₜ t2.tensor))
| contr i j h t => (S.contrMap _ i j h).hom t.tensor
| eval i e t => (S.evalMap i (Fin.ofNat' e Fin.size_pos')) t.tensor
| eval i e t => (S.evalMap i (Fin.ofNat' _ e)) t.tensor
| action g t => (S.F.obj (OverColor.mk _)).ρ g t.tensor
/-- Takes a tensor tree based on `Fin 0`, into the field `S.k`. -/
@ -730,7 +730,7 @@ lemma contr_tensor {n : } {c : Fin n.succ.succ → S.C} {i : Fin n.succ.succ}
lemma neg_tensor (t : TensorTree S c) : (neg t).tensor = - t.tensor := rfl
lemma eval_tensor {n : } {c : Fin n.succ → S.C} (i : Fin n.succ) (e : ) (t : TensorTree S c) :
(eval i e t).tensor = (S.evalMap i (Fin.ofNat' e Fin.size_pos')) t.tensor := rfl
(eval i e t).tensor = (S.evalMap i (Fin.ofNat' (S.repDim (c i)) e)) t.tensor := rfl
lemma smul_tensor {c : Fin n → S.C} (a : S.k) (T : TensorTree S c) :
(smul a T).tensor = a • T.tensor:= rfl