Merge pull request #189 from HEPLean/IndexNotation
refactor: Simp to simp only
This commit is contained in:
commit
c783b7a641
39 changed files with 393 additions and 233 deletions
|
@ -70,10 +70,12 @@ lemma doublePoint_B₃_B₃ (R : MSSMACC.LinSols) : cubeTriLin B₃.val B₃.val
|
|||
simp only [mul_one, Fin.isValue, toSMSpecies_apply, one_mul, mul_neg, neg_neg, neg_mul, 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 h0 := hLin 0
|
||||
have h2 := hLin 2
|
||||
simp [Fin.sum_univ_three] at h0 h2
|
||||
simp only [Fin.isValue, Fin.sum_univ_three, Prod.mk_zero_zero, Prod.mk_one_one, LinearMap.coe_mk,
|
||||
AddHom.coe_mk] at h0 h2
|
||||
linear_combination (norm := ring_nf) 9 * (h0) - 24 * (h2)
|
||||
simp only [Fin.isValue, Prod.mk_zero_zero, Prod.mk_one_one, add_add_sub_cancel, add_neg_cancel]
|
||||
|
||||
|
|
|
@ -142,14 +142,16 @@ def accGrav : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [mul_add]
|
||||
simp only [MSSMSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add,
|
||||
toSMSpecies_apply, reduceMul, Fin.isValue, mul_add, Hd_apply, Fin.reduceFinMk, Hu_apply]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat rw [(toSMSpecies _).map_smul]
|
||||
erw [Hd.map_smul, Hu.map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [MSSMSpecies_numberCharges, HSMul.hSMul, SMul.smul, Fin.isValue, toSMSpecies_apply,
|
||||
reduceMul, Hd_apply, Fin.reduceFinMk, Hu_apply, eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
--rw [show Rat.cast a = a from rfl]
|
||||
|
@ -175,14 +177,16 @@ def accSU2 : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [mul_add]
|
||||
simp only [MSSMSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add,
|
||||
toSMSpecies_apply, reduceMul, Fin.isValue, mul_add, Hd_apply, Fin.reduceFinMk, Hu_apply]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat rw [(toSMSpecies _).map_smul]
|
||||
erw [Hd.map_smul, Hu.map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [MSSMSpecies_numberCharges, HSMul.hSMul, SMul.smul, Fin.isValue, toSMSpecies_apply,
|
||||
reduceMul, Hd_apply, Fin.reduceFinMk, Hu_apply, eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
--rw [show Rat.cast a = a from rfl]
|
||||
|
@ -208,13 +212,15 @@ def accSU3 : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [mul_add]
|
||||
simp only [MSSMSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add,
|
||||
toSMSpecies_apply, reduceMul, Fin.isValue, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat rw [(toSMSpecies _).map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [MSSMSpecies_numberCharges, HSMul.hSMul, SMul.smul, Fin.isValue, toSMSpecies_apply,
|
||||
reduceMul, eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
--rw [show Rat.cast a = a from rfl]
|
||||
|
@ -239,14 +245,16 @@ def accYY : MSSMCharges.Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [mul_add]
|
||||
simp only [MSSMSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add,
|
||||
toSMSpecies_apply, reduceMul, Fin.isValue, mul_add, Hd_apply, Fin.reduceFinMk, Hu_apply]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat rw [(toSMSpecies _).map_smul]
|
||||
erw [Hd.map_smul, Hu.map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [MSSMSpecies_numberCharges, HSMul.hSMul, SMul.smul, Fin.isValue, toSMSpecies_apply,
|
||||
reduceMul, Hd_apply, Fin.reduceFinMk, Hu_apply, eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
-- rw [show Rat.cast a = a from rfl]
|
||||
|
@ -462,14 +470,14 @@ def AnomalyFreeMk (S : MSSMACC.Charges) (hg : accGrav S = 0)
|
|||
(hquad : accQuad S = 0) (hcube : accCube S = 0) : MSSMACC.Sols :=
|
||||
⟨⟨⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [MSSMACC_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact hg
|
||||
| 1 => exact hsu2
|
||||
| 2 => exact hsu3
|
||||
| 3 => exact hyy⟩, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [MSSMACC_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => exact hquad⟩, hcube⟩
|
||||
|
||||
|
@ -485,7 +493,7 @@ def AnomalyFreeQuadMk' (S : MSSMACC.LinSols) (hquad : accQuad S.val = 0) :
|
|||
MSSMACC.QuadSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [MSSMACC_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => exact hquad⟩
|
||||
|
||||
|
@ -495,7 +503,7 @@ def AnomalyFreeMk' (S : MSSMACC.LinSols) (hquad : accQuad S.val = 0)
|
|||
(hcube : accCube S.val = 0) : MSSMACC.Sols :=
|
||||
⟨⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [MSSMACC_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => exact hquad⟩, hcube⟩
|
||||
|
||||
|
|
|
@ -81,7 +81,7 @@ lemma planeY₃B₃_val_eq' (R : MSSMACC.AnomalyFreePerp) (a b c : ℚ) (hR' : R
|
|||
obtain ⟨i, hi⟩ := hR
|
||||
have h2 := congrArg (fun S => S i) h1i
|
||||
change _ = 0 at h2
|
||||
simp [HSMul.hSMul] at h2
|
||||
simp only [HSMul.hSMul, ACCSystemCharges.chargesModule_smul, mul_eq_zero] at h2
|
||||
have hc : c + -c' = 0 := by
|
||||
cases h2 <;> rename_i h2
|
||||
exact h2
|
||||
|
|
|
@ -426,13 +426,13 @@ theorem toSol_surjective : Function.Surjective toSol := by
|
|||
intro T
|
||||
by_cases h₁ : ¬ LineEqPropSol T
|
||||
· exact toSol_toSolNSProj ⟨T, h₁⟩
|
||||
· simp at h₁
|
||||
· simp only [not_not] at h₁
|
||||
by_cases h₂ : ¬ InQuadSolProp T
|
||||
· exact toSol_inLineEq ⟨T, And.intro h₁ h₂⟩
|
||||
· simp at h₂
|
||||
· simp only [not_not] at h₂
|
||||
by_cases h₃ : ¬ InCubeSolProp T
|
||||
· exact toSol_inQuad ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩
|
||||
· simp at h₃
|
||||
· simp only [not_not] at h₃
|
||||
exact toSol_inQuadCube ⟨T, And.intro h₁ (And.intro h₂ h₃)⟩
|
||||
|
||||
end AnomalyFreePerp
|
||||
|
|
|
@ -31,7 +31,7 @@ def accGrav (n : ℕ) : ((PureU1Charges n).Charges →ₗ[ℚ] ℚ) where
|
|||
toFun S := ∑ i : Fin n, S i
|
||||
map_add' S T := Finset.sum_add_distrib
|
||||
map_smul' a S := by
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [HSMul.hSMul, SMul.smul, eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
rw [← Finset.mul_sum]
|
||||
|
||||
/-- The symmetric trilinear form used to define the cubic anomaly. -/
|
||||
|
@ -40,7 +40,7 @@ def accCubeTriLinSymm {n : ℕ} : TriLinearSymm (PureU1Charges n).Charges := Tri
|
|||
(fun S => ∑ i, S.1 i * S.2.1 i * S.2.2 i)
|
||||
(by
|
||||
intro a S L T
|
||||
simp [HSMul.hSMul]
|
||||
simp only [PureU1Charges_numberCharges, HSMul.hSMul, ACCSystemCharges.chargesModule_smul]
|
||||
rw [Finset.mul_sum]
|
||||
apply Fintype.sum_congr
|
||||
intro i
|
||||
|
@ -108,7 +108,7 @@ open BigOperators
|
|||
lemma pureU1_linear {n : ℕ} (S : (PureU1 n.succ).LinSols) :
|
||||
∑ i, S.val i = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [succ_eq_add_one, PureU1_numberLinear, PureU1_linearACCs] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma pureU1_cube {n : ℕ} (S : (PureU1 n.succ).Sols) :
|
||||
|
@ -120,7 +120,7 @@ lemma pureU1_cube {n : ℕ} (S : (PureU1 n.succ).Sols) :
|
|||
lemma pureU1_last {n : ℕ} (S : (PureU1 n.succ).LinSols) :
|
||||
S.val (Fin.last n) = - ∑ i : Fin n, S.val i.castSucc := by
|
||||
have hS := pureU1_linear S
|
||||
simp at hS
|
||||
simp only [succ_eq_add_one, PureU1_numberCharges] at hS
|
||||
rw [Fin.sum_univ_castSucc] at hS
|
||||
linear_combination hS
|
||||
|
||||
|
@ -134,7 +134,7 @@ lemma pureU1_anomalyFree_ext {n : ℕ} {S T : (PureU1 n.succ).LinSols}
|
|||
obtain ⟨j, hj⟩ := hiCast
|
||||
rw [← hj]
|
||||
exact h j
|
||||
· simp at hi
|
||||
· simp only [succ_eq_add_one, PureU1_numberCharges, ne_eq, Decidable.not_not] at hi
|
||||
rw [hi, pureU1_last, pureU1_last]
|
||||
simp only [neg_inj]
|
||||
apply Finset.sum_congr
|
||||
|
|
|
@ -36,14 +36,14 @@ lemma asCharges_eq_castSucc (j : Fin n) :
|
|||
|
||||
lemma asCharges_ne_castSucc {k j : Fin n} (h : k ≠ j) :
|
||||
asCharges k (Fin.castSucc j) = 0 := by
|
||||
simp [asCharges]
|
||||
simp only [asCharges, Nat.succ_eq_add_one, PureU1_numberCharges, Fin.castSucc_inj]
|
||||
split
|
||||
· rename_i h1
|
||||
exact False.elim (h (id (Eq.symm h1)))
|
||||
· split
|
||||
· rename_i h1 h2
|
||||
rw [Fin.ext_iff] at h1 h2
|
||||
simp at h1 h2
|
||||
simp only [Fin.coe_castSucc, Fin.val_last] at h1 h2
|
||||
have hj : j.val < n := by
|
||||
exact j.prop
|
||||
simp_all
|
||||
|
@ -54,7 +54,7 @@ lemma asCharges_ne_castSucc {k j : Fin n} (h : k ≠ j) :
|
|||
def asLinSols (j : Fin n) : (PureU1 n.succ).LinSols :=
|
||||
⟨asCharges j, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [Nat.succ_eq_add_one, PureU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 =>
|
||||
simp only [Fin.isValue, PureU1_linearACCs, accGrav,
|
||||
|
|
|
@ -34,7 +34,7 @@ lemma constAbs_perm (S : (PureU1 n).Charges) (M :(FamilyPermutations n).group) :
|
|||
MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply]
|
||||
refine Iff.intro (fun h i j => ?_) (fun h i j => h (M.invFun i) (M.invFun j))
|
||||
have h2 := h (M.toFun i) (M.toFun j)
|
||||
simp at h2
|
||||
simp only [Equiv.toFun_as_coe, Equiv.Perm.inv_apply_self] at h2
|
||||
exact h2
|
||||
|
||||
lemma constAbs_sort {S : (PureU1 n).Charges} (CA : ConstAbs S) : ConstAbs (sort S) := by
|
||||
|
@ -93,7 +93,7 @@ lemma is_zero (h0 : S (0 : Fin n.succ) = 0) : S = 0 := by
|
|||
funext i
|
||||
have ht := hS.1 i (0 : Fin n.succ)
|
||||
rw [h0] at ht
|
||||
simp at ht
|
||||
simp only [ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, pow_eq_zero_iff] at ht
|
||||
exact ht
|
||||
|
||||
/-- A boundary of `S : (PureU1 n.succ).charges` (assumed sorted, constAbs and non-zero)
|
||||
|
@ -117,7 +117,8 @@ lemma boundary_split (k : Fin n) : k.succ.val + (n.succ - k.succ.val) = n.succ :
|
|||
|
||||
lemma boundary_accGrav' (k : Fin n) : accGrav n.succ S =
|
||||
∑ i : Fin (k.succ.val + (n.succ - k.succ.val)), S (Fin.cast (boundary_split k) i) := by
|
||||
simp [accGrav]
|
||||
simp only [succ_eq_add_one, accGrav, LinearMap.coe_mk, AddHom.coe_mk, Fin.val_succ,
|
||||
PureU1_numberCharges]
|
||||
erw [Finset.sum_equiv (Fin.castOrderIso (boundary_split k)).toEquiv]
|
||||
· intro i
|
||||
simp only [Fin.val_succ, mem_univ, RelIso.coe_fn_toEquiv]
|
||||
|
@ -149,7 +150,7 @@ 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⟩
|
||||
simp at hnot
|
||||
simp only [HasBoundary, Boundary, not_exists, not_and, not_lt] at hnot
|
||||
induction i
|
||||
· rfl
|
||||
· rename_i i hii
|
||||
|
@ -163,7 +164,7 @@ 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
|
||||
· exact not_hasBoundary_zero_le hS hnot hi i
|
||||
· simp at hi
|
||||
· simp only [not_lt] at hi
|
||||
exact zero_gt hS hi i
|
||||
|
||||
include hS in
|
||||
|
@ -175,9 +176,9 @@ 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
|
||||
simp [accGrav] at h0
|
||||
simp only [succ_eq_add_one, accGrav, LinearMap.coe_mk, AddHom.coe_mk, cast_add, cast_one] at h0
|
||||
erw [pureU1_linear A] at h0
|
||||
simp at h0
|
||||
simp only [zero_eq_mul] at h0
|
||||
cases' h0
|
||||
· linarith
|
||||
· simp_all
|
||||
|
@ -187,9 +188,9 @@ lemma AFL_odd_noBoundary {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted
|
|||
by_contra hn
|
||||
obtain ⟨k, hk⟩ := hn
|
||||
have h0 := boundary_accGrav'' h k hk
|
||||
simp [accGrav] at h0
|
||||
simp only [succ_eq_add_one, accGrav, LinearMap.coe_mk, AddHom.coe_mk, cast_mul, cast_ofNat] at h0
|
||||
erw [pureU1_linear A] at h0
|
||||
simp [hA] at h0
|
||||
simp only [zero_eq_mul, hA, or_false] at h0
|
||||
have h1 : 2 * n = 2 * k.val + 1 := by
|
||||
rw [← @Nat.cast_inj ℚ]
|
||||
simp only [cast_mul, cast_ofNat, cast_add, cast_one]
|
||||
|
@ -212,7 +213,8 @@ lemma AFL_even_Boundary {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted
|
|||
have h0 := boundary_accGrav'' h k hk
|
||||
change ∑ i : Fin (succ (Nat.mul 2 n + 1)), A.val i = _ at h0
|
||||
erw [pureU1_linear A] at h0
|
||||
simp [hA] at h0
|
||||
simp only [mul_eq, cast_add, cast_mul, cast_ofNat, cast_one, add_sub_add_right_eq_sub,
|
||||
succ_eq_add_one, zero_eq_mul, hA, or_false] at h0
|
||||
rw [← @Nat.cast_inj ℚ]
|
||||
linear_combination h0 / 2
|
||||
|
||||
|
|
|
@ -158,8 +158,7 @@ lemma basis!_on_δ!₁_self (j : Fin n) : basis!AsCharges j (δ!₁ j) = 1 := by
|
|||
|
||||
lemma basis_on_δ₁_other {k j : Fin n.succ} (h : k ≠ j) :
|
||||
basisAsCharges k (δ₁ j) = 0 := by
|
||||
simp [basisAsCharges]
|
||||
simp [δ₁, δ₂]
|
||||
simp only [basisAsCharges, succ_eq_add_one, PureU1_numberCharges, δ₁, succ_eq_add_one, δ₂]
|
||||
split
|
||||
· rename_i h1
|
||||
rw [Fin.ext_iff] at h1
|
||||
|
|
|
@ -113,7 +113,7 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ)).LinSols}
|
|||
have h1 := Pa_δ!₄ f g
|
||||
have h2 := Pa_δ!₁ f g (Fin.last n)
|
||||
have h3 := Pa_δ!₂ f g (Fin.last n)
|
||||
simp at h1 h2 h3
|
||||
simp only [Fin.succ_last, Nat.succ_eq_add_one] at h1 h2 h3
|
||||
have hl : f (Fin.succ (Fin.last n)) = - Pa f g δ!₄ := by
|
||||
simp_all only [Fin.succ_last, neg_neg]
|
||||
erw [hl] at h2
|
||||
|
@ -133,7 +133,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols}
|
|||
obtain ⟨g, f, hfg⟩ := span_basis S
|
||||
have h1 := lineInCubicPerm_swap LIC (Fin.last n) g f hfg
|
||||
rw [P_P_P!_accCube' g f hfg] at h1
|
||||
simp at h1
|
||||
simp only [Nat.succ_eq_add_one, neg_add_rev, mul_eq_zero] at h1
|
||||
cases h1 <;> rename_i h1
|
||||
· apply Or.inl
|
||||
linear_combination h1
|
||||
|
@ -148,7 +148,8 @@ lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
|
|||
δ!₄ ?_ ?_ ?_ ?_
|
||||
· simp [Fin.ext_iff, δ!₂, δ!₁]
|
||||
· 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]
|
||||
omega
|
||||
· exact fun M => lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
|
||||
|
||||
|
|
|
@ -63,7 +63,7 @@ lemma parameterizationCharge_cube (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (
|
|||
/-- The construction of a `Sol` from a `Fin n.succ → ℚ`, a `Fin n → ℚ` and a `ℚ`. -/
|
||||
def parameterization (g : Fin n.succ → ℚ) (f : Fin n → ℚ) (a : ℚ) :
|
||||
(PureU1 (2 * n.succ)).Sols :=
|
||||
⟨⟨parameterizationAsLinear g f a, by intro i; simp at i; exact Fin.elim0 i⟩,
|
||||
⟨⟨parameterizationAsLinear g f a, fun i => Fin.elim0 i⟩,
|
||||
parameterizationCharge_cube g f a⟩
|
||||
|
||||
lemma anomalyFree_param {S : (PureU1 (2 * n.succ)).Sols}
|
||||
|
|
|
@ -58,19 +58,24 @@ lemma lineInPlaneCond_eq_last' {S : (PureU1 (n.succ.succ)).LinSols} (hS : LineIn
|
|||
erw [sq_eq_sq_iff_eq_or_eq_neg] at h
|
||||
rw [LineInPlaneCond] at hS
|
||||
simp only [LineInPlaneProp] at hS
|
||||
simp [not_or] at h
|
||||
simp only [Nat.succ_eq_add_one, Fin.succ_last, not_or] at h
|
||||
have h1 (i : Fin n) : S.val i.castSucc.castSucc =
|
||||
- (S.val ((Fin.last n).castSucc) + (S.val ((Fin.last n).succ))) / 2 := by
|
||||
have h1S := hS (Fin.last n).castSucc ((Fin.last n).succ) i.castSucc.castSucc
|
||||
(by simp; rw [Fin.ext_iff]; exact Nat.ne_add_one ↑(Fin.last n).castSucc)
|
||||
(by simp; rw [Fin.ext_iff]; simp; omega)
|
||||
(by simp; rw [Fin.ext_iff]; simp; omega)
|
||||
simp_all
|
||||
(by simp only [Fin.ext_iff, Nat.succ_eq_add_one, Fin.succ_last, ne_eq]
|
||||
exact Nat.ne_add_one ↑(Fin.last n).castSucc)
|
||||
(by simp only [Nat.succ_eq_add_one, Fin.succ_last, ne_eq, Fin.ext_iff, Fin.val_last,
|
||||
Fin.coe_castSucc]
|
||||
omega)
|
||||
(by simp only [Nat.succ_eq_add_one, ne_eq, Fin.ext_iff, Fin.coe_castSucc, Fin.val_last]
|
||||
omega)
|
||||
simp_all only [Nat.succ_eq_add_one, ne_eq, Fin.succ_last, false_or, neg_add_rev]
|
||||
field_simp
|
||||
linear_combination h1S
|
||||
have h2 := pureU1_last S
|
||||
rw [Fin.sum_univ_castSucc] at h2
|
||||
simp [h1] at h2
|
||||
simp only [Nat.succ_eq_add_one, h1, Fin.succ_last, neg_add_rev, Finset.sum_const,
|
||||
Finset.card_univ, Fintype.card_fin, nsmul_eq_mul] at h2
|
||||
field_simp at h2
|
||||
linear_combination h2
|
||||
|
||||
|
@ -81,7 +86,7 @@ lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
|
|||
by_contra hn
|
||||
have h := lineInPlaneCond_eq_last' hS hn
|
||||
rw [sq_eq_sq_iff_eq_or_eq_neg] at hn
|
||||
simp [or_not] at hn
|
||||
simp only [Nat.succ_eq_add_one, Fin.succ_last, not_or] at hn
|
||||
have hx : ((2 : ℚ) - ↑(n + 3)) ≠ 0 := by
|
||||
rw [Nat.cast_add]
|
||||
simp only [Nat.cast_ofNat, ne_eq]
|
||||
|
@ -91,8 +96,8 @@ lemma lineInPlaneCond_eq_last {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
|
|||
have ht : S.val ((Fin.last n.succ.succ.succ).succ) =
|
||||
- S.val ((Fin.last n.succ.succ.succ).castSucc) := by
|
||||
rw [← mul_right_inj' hx]
|
||||
simp [Nat.succ_eq_add_one]
|
||||
simp at h
|
||||
simp only [Nat.cast_add, Nat.cast_ofNat, Nat.succ_eq_add_one, Fin.succ_last, mul_neg]
|
||||
simp only [Nat.cast_add, Nat.cast_ofNat, Nat.succ_eq_add_one, neg_sub] at h
|
||||
rw [h]
|
||||
ring
|
||||
simp_all
|
||||
|
@ -111,19 +116,19 @@ theorem linesInPlane_constAbs {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
|
|||
intro i j
|
||||
by_cases hij : i ≠ j
|
||||
· exact linesInPlane_eq_sq hS i j hij
|
||||
· simp at hij
|
||||
· simp only [Nat.succ_eq_add_one, PureU1_numberCharges, ne_eq, Decidable.not_not] at hij
|
||||
rw [hij]
|
||||
|
||||
lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
|
||||
ConstAbsProp (S.val (0 : Fin 4), S.val (1 : Fin 4)) := by
|
||||
simp [ConstAbsProp]
|
||||
simp only [ConstAbsProp, Fin.isValue]
|
||||
by_contra hn
|
||||
have hLin := pureU1_linear S.1.1
|
||||
have hcube := pureU1_cube S
|
||||
simp at hLin hcube
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, PureU1_numberCharges] at hLin hcube
|
||||
rw [Fin.sum_univ_four] at hLin hcube
|
||||
rw [sq_eq_sq_iff_eq_or_eq_neg] at hn
|
||||
simp [not_or] at hn
|
||||
simp only [Fin.isValue, not_or] at hn
|
||||
have l012 := hS 0 1 2 (ne_of_beq_false rfl) (ne_of_beq_false rfl) (ne_of_beq_false rfl)
|
||||
have l013 := hS 0 1 3 (ne_of_beq_false rfl) (ne_of_beq_false rfl) (ne_of_beq_false rfl)
|
||||
have l023 := hS 0 2 3 (ne_of_beq_false rfl) (ne_of_beq_false rfl) (ne_of_beq_false rfl)
|
||||
|
@ -137,8 +142,9 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
|
|||
rw [← h1, h3] at hcube
|
||||
have h4 : S.val (2 : Fin 4) ^ 3 = 0 := by
|
||||
linear_combination -1 * hcube / 24
|
||||
simp at h4
|
||||
simp_all
|
||||
simp only [Fin.isValue, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff] at h4
|
||||
simp_all only [neg_mul, neg_neg, add_left_eq_self, add_right_eq_self, not_true_eq_false,
|
||||
false_and]
|
||||
· by_cases h3 : S.val (0 : Fin 4) = - S.val (2 : Fin 4)
|
||||
· simp_all
|
||||
have h4 : S.val (1 : Fin 4) = - S.val (2 : Fin 4) := by
|
||||
|
@ -152,7 +158,7 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
|
|||
rw [h4, h5] at hcube
|
||||
have h6 : S.val (3 : Fin 4) ^ 3 = 0 := by
|
||||
linear_combination -1 * hcube / 24
|
||||
simp at h6
|
||||
simp only [Fin.isValue, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff] at h6
|
||||
simp_all
|
||||
|
||||
lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols}
|
||||
|
@ -170,7 +176,7 @@ lemma linesInPlane_constAbs_four (S : (PureU1 4).Sols)
|
|||
intro i j
|
||||
by_cases hij : i ≠ j
|
||||
· exact linesInPlane_eq_sq_four hS i j hij
|
||||
· simp at hij
|
||||
· simp only [PureU1_numberCharges, ne_eq, Decidable.not_not] at hij
|
||||
rw [hij]
|
||||
|
||||
theorem linesInPlane_constAbs_AF (S : (PureU1 (n.succ.succ.succ.succ)).Sols)
|
||||
|
|
|
@ -24,9 +24,10 @@ namespace One
|
|||
theorem solEqZero (S : (PureU1 1).LinSols) : S = 0 := by
|
||||
apply ACCSystemLinear.LinSols.ext
|
||||
have hLin := pureU1_linear S
|
||||
simp at hLin
|
||||
simp only [succ_eq_add_one, reduceAdd, PureU1_numberCharges, univ_unique, Fin.default_eq_zero,
|
||||
Fin.isValue, sum_singleton] at hLin
|
||||
funext i
|
||||
simp at i
|
||||
simp only [PureU1_numberCharges] at i
|
||||
rw [show i = (0 : Fin 1) from Fin.fin_one_eq_zero i]
|
||||
exact hLin
|
||||
|
||||
|
|
|
@ -25,7 +25,7 @@ lemma cube_for_linSol' (S : (PureU1 3).LinSols) :
|
|||
3 * S.val (0 : Fin 3) * S.val (1 : Fin 3) * S.val (2 : Fin 3) = 0 ↔
|
||||
(PureU1 3).cubicACC S.val = 0 := by
|
||||
have hL := pureU1_linear S
|
||||
simp at hL
|
||||
simp only [succ_eq_add_one, reduceAdd, PureU1_numberCharges] at hL
|
||||
rw [Fin.sum_univ_three] at hL
|
||||
change _ ↔ accCube _ _ = _
|
||||
rw [accCube_explicit, Fin.sum_univ_three]
|
||||
|
@ -47,7 +47,7 @@ lemma three_sol_zero (S : (PureU1 3).Sols) : S.val (0 : Fin 3) = 0 ∨ S.val (1
|
|||
def solOfLinear (S : (PureU1 3).LinSols)
|
||||
(hS : S.val (0 : Fin 3) = 0 ∨ S.val (1 : Fin 3) = 0 ∨ S.val (2 : Fin 3) = 0) :
|
||||
(PureU1 3).Sols :=
|
||||
⟨⟨S, by intro i; simp at i; exact Fin.elim0 i⟩,
|
||||
⟨⟨S, fun i => Fin.elim0 i⟩,
|
||||
(cube_for_linSol S).mp hS⟩
|
||||
|
||||
theorem solOfLinear_surjects (S : (PureU1 3).Sols) :
|
||||
|
|
|
@ -23,9 +23,10 @@ namespace Two
|
|||
|
||||
/-- An equivalence between `LinSols` and `Sols`. -/
|
||||
def equiv : (PureU1 2).LinSols ≃ (PureU1 2).Sols where
|
||||
toFun S := ⟨⟨S, by intro i; simp at i; exact Fin.elim0 i⟩, by
|
||||
toFun S := ⟨⟨S, fun i => Fin.elim0 i⟩, by
|
||||
have hLin := pureU1_linear S
|
||||
simp at hLin
|
||||
simp only [succ_eq_add_one, reduceAdd, PureU1_numberCharges, Fin.sum_univ_two,
|
||||
Fin.isValue] at hLin
|
||||
erw [accCube_explicit]
|
||||
simp only [Fin.sum_univ_two, Fin.isValue]
|
||||
rw [show S.val (0 : Fin 2) = - S.val (1 : Fin 2) from eq_neg_of_add_eq_zero_left hLin]
|
||||
|
|
|
@ -459,7 +459,7 @@ lemma P_accCube (f : Fin n → ℚ) : accCube (2 * n +1) (P f) = 0 := by
|
|||
simp only [ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, Function.comp_apply, zero_add]
|
||||
apply Finset.sum_eq_zero
|
||||
intro i _
|
||||
simp [P_δ₁, P_δ₂]
|
||||
simp only [P_δ₁, P_δ₂]
|
||||
ring
|
||||
|
||||
lemma P!_accCube (f : Fin n → ℚ) : accCube (2 * n +1) (P! f) = 0 := by
|
||||
|
@ -514,7 +514,7 @@ lemma Pa_zero (f g : Fin n.succ → ℚ) (h : Pa f g = 0) :
|
|||
change 0 = _ at h1
|
||||
simp only [neg_zero, succ_eq_add_one, zero_sub, zero_eq_neg] at h1
|
||||
have h2 := Pa_δa₂ f g ⟨iv, succ_lt_succ_iff.mp hiv⟩
|
||||
simp [h, h1] at h2
|
||||
simp only [succ_eq_add_one, h, Fin.succ_mk, Fin.castSucc_mk, h1, add_zero] at h2
|
||||
exact h2.symm
|
||||
exact hinduc i.val i.prop
|
||||
|
||||
|
|
|
@ -108,7 +108,7 @@ lemma P_P_P!_accCube' {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
|
|||
have h4 := Pa_δa₄ f g 0
|
||||
have h2 := Pa_δa₂ f g 0
|
||||
rw [← hS] at h1 h2 h4
|
||||
simp at h2
|
||||
simp only [Nat.succ_eq_add_one, Fin.succ_zero_eq_one, Fin.castSucc_zero] at h2
|
||||
have h5 : f 1 = S.val (δa₂ 0) + S.val δa₁ + S.val (δa₄ 0) := by
|
||||
linear_combination -(1 * h1) - 1 * h4 - 1 * h2
|
||||
rw [h5, δa₄_δ!₂, show (δa₂ (0 : Fin n.succ)) = δ!₁ 0 from rfl, δa₁_δ!₃]
|
||||
|
@ -120,7 +120,7 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols}
|
|||
obtain ⟨g, f, hfg⟩ := span_basis S
|
||||
have h1 := lineInCubicPerm_swap LIC 0 g f hfg
|
||||
rw [P_P_P!_accCube' g f hfg] at h1
|
||||
simp at h1
|
||||
simp only [Nat.succ_eq_add_one, mul_eq_zero] at h1
|
||||
cases h1 <;> rename_i h1
|
||||
· apply Or.inl
|
||||
linear_combination h1
|
||||
|
|
|
@ -60,7 +60,7 @@ lemma parameterizationCharge_cube (g f : Fin n → ℚ) (a : ℚ) :
|
|||
/-- Given a `g f : Fin n → ℚ` and a `a : ℚ` we form a solution. -/
|
||||
def parameterization (g f : Fin n → ℚ) (a : ℚ) :
|
||||
(PureU1 (2 * n + 1)).Sols :=
|
||||
⟨⟨parameterizationAsLinear g f a, by intro i; simp at i; exact Fin.elim0 i⟩,
|
||||
⟨⟨parameterizationAsLinear g f a, fun i => Fin.elim0 i⟩,
|
||||
parameterizationCharge_cube g f a⟩
|
||||
|
||||
lemma anomalyFree_param {S : (PureU1 (2 * n + 1)).Sols}
|
||||
|
@ -125,13 +125,10 @@ theorem generic_case {S : (PureU1 (2 * n.succ + 1)).Sols} (h : GenericCase S) :
|
|||
apply ACCSystem.Sols.ext
|
||||
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 [anomalyFree_param _ _ hS, 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
|
||||
simp at h
|
||||
exact h
|
||||
· simpa only [Nat.succ_eq_add_one, accCubeTriLinSymm_toFun_apply_apply, ne_eq,
|
||||
anomalyFree_param _ _ hS, neg_eq_zero] using h g f hS
|
||||
|
||||
lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
|
||||
(h : SpecialCase S) :
|
||||
|
@ -139,15 +136,13 @@ lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ + 1)).Sols}
|
|||
intro g f hS a b
|
||||
erw [TriLinearSymm.toCubic_add]
|
||||
rw [HomogeneousCubic.map_smul, HomogeneousCubic.map_smul]
|
||||
erw [P_accCube]
|
||||
erw [P!_accCube]
|
||||
erw [P_accCube, P!_accCube]
|
||||
have h := h g f hS
|
||||
rw [accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
|
||||
accCubeTriLinSymm.map_smul₃, accCubeTriLinSymm.map_smul₁, accCubeTriLinSymm.map_smul₂,
|
||||
accCubeTriLinSymm.map_smul₃]
|
||||
rw [h]
|
||||
accCubeTriLinSymm.map_smul₃, h]
|
||||
rw [anomalyFree_param _ _ hS] at h
|
||||
simp at h
|
||||
simp only [Nat.succ_eq_add_one, accCubeTriLinSymm_toFun_apply_apply, neg_eq_zero] at h
|
||||
change accCubeTriLinSymm (P! f) (P! f) (P g) = 0 at h
|
||||
erw [h]
|
||||
simp
|
||||
|
|
|
@ -93,13 +93,15 @@ def accGrav : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply,
|
||||
Fin.isValue, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
-- rw [show Rat.cast a = a from rfl]
|
||||
|
@ -127,13 +129,15 @@ def accSU2 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply,
|
||||
Fin.isValue, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
-- rw [show Rat.cast a = a from rfl]
|
||||
|
@ -160,13 +164,15 @@ def accSU3 : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_add' S T := by
|
||||
simp only
|
||||
repeat rw [map_add]
|
||||
simp [Pi.add_apply, mul_add]
|
||||
simp only [SMνSpecies_numberCharges, ACCSystemCharges.chargesAddCommMonoid_add, toSpecies_apply,
|
||||
Fin.isValue, mul_add]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
ring
|
||||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
-- rw [show Rat.cast a = a from rfl]
|
||||
|
@ -201,7 +207,8 @@ def accYY : (SMνCharges n).Charges →ₗ[ℚ] ℚ where
|
|||
map_smul' a S := by
|
||||
simp only
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue,
|
||||
eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
repeat erw [Finset.sum_add_distrib]
|
||||
repeat erw [← Finset.mul_sum]
|
||||
-- rw [show Rat.cast a = a from rfl]
|
||||
|
@ -235,7 +242,7 @@ def quadBiLin : BiLinearSymm (SMνCharges n).Charges := BiLinearSymm.mk₂
|
|||
rw [Finset.mul_sum]
|
||||
refine Fintype.sum_congr _ _ fun i ↦ ?_
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue, neg_mul, one_mul]
|
||||
ring)
|
||||
(by
|
||||
intro S T R
|
||||
|
@ -298,7 +305,7 @@ def cubeTriLin : TriLinearSymm (SMνCharges n).Charges := TriLinearSymm.mk₃
|
|||
rw [Finset.mul_sum]
|
||||
refine Fintype.sum_congr _ _ fun i ↦ ?_
|
||||
repeat erw [map_smul]
|
||||
simp [HSMul.hSMul, SMul.smul]
|
||||
simp only [HSMul.hSMul, SMul.smul, toSpecies_apply, Fin.isValue]
|
||||
ring)
|
||||
(by
|
||||
intro S T R L
|
||||
|
|
|
@ -69,7 +69,8 @@ def speciesEmbed (m n : ℕ) :
|
|||
rfl
|
||||
map_smul' a S := by
|
||||
funext i
|
||||
simp [HSMul.hSMul]
|
||||
simp only [SMνSpecies_numberCharges, HSMul.hSMul, ACCSystemCharges.chargesModule_smul,
|
||||
eq_ratCast, Rat.cast_eq_id, id_eq]
|
||||
by_cases hi : i.val < m
|
||||
· erw [dif_pos hi, dif_pos hi]
|
||||
· erw [dif_neg hi, dif_neg hi]
|
||||
|
|
|
@ -38,12 +38,12 @@ variable {n : ℕ}
|
|||
|
||||
lemma SU2Sol (S : (SMNoGrav n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU3Sol (S : (SMNoGrav n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [SMNoGrav_numberLinear, SMNoGrav_linearACCs, Fin.isValue] at hS
|
||||
exact hS 1
|
||||
|
||||
lemma cubeSol (S : (SMNoGrav n).Sols) : accCube S.val = 0 := S.cubicSol
|
||||
|
@ -54,7 +54,7 @@ def chargeToLinear (S : (SMNoGrav n).Charges) (hSU2 : accSU2 S = 0) (hSU3 : accS
|
|||
(SMNoGrav n).LinSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [SMNoGrav_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact hSU2
|
||||
| 1 => exact hSU3⟩
|
||||
|
|
|
@ -31,7 +31,11 @@ def B₀ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₀_cubic (S T : (SM 3).Charges) : cubeTriLin B₀ S T =
|
||||
6 * (S (0 : Fin 18) * T (0 : Fin 18) - S (1 : Fin 18) * T (1 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₀, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₀, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
one_mul, zero_add, Fin.reduceFinMk, zero_mul, Fin.mk_one, neg_mul, mul_neg, Nat.reduceAdd]
|
||||
ring
|
||||
|
||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||
|
@ -43,7 +47,11 @@ def B₁ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₁_cubic (S T : (SM 3).Charges) : cubeTriLin B₁ S T =
|
||||
3 * (S (3 : Fin 18) * T (3 : Fin 18) - S (4 : Fin 18) * T (4 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₁, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₁, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
zero_mul, zero_add, Fin.reduceFinMk, one_mul, Fin.mk_one, Nat.reduceAdd, neg_mul, mul_neg]
|
||||
ring
|
||||
|
||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||
|
@ -55,7 +63,11 @@ def B₂ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₂_cubic (S T : (SM 3).Charges) : cubeTriLin B₂ S T =
|
||||
3 * (S (6 : Fin 18) * T (6 : Fin 18) - S (7 : Fin 18) * T (7 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₂, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₂, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
zero_mul, zero_add, Fin.reduceFinMk, one_mul, Fin.mk_one, Nat.reduceAdd, neg_mul, mul_neg]
|
||||
ring
|
||||
|
||||
/-- A charge assignment forming one of the basis elements of the plane. -/
|
||||
|
@ -67,7 +79,11 @@ def B₃ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₃_cubic (S T : (SM 3).Charges) : cubeTriLin B₃ S T =
|
||||
2 * (S (9 : Fin 18) * T (9 : Fin 18) - S (10 : Fin 18) * T (10 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₃, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₃, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
zero_mul, zero_add, Fin.reduceFinMk, one_mul, Fin.mk_one, Nat.reduceAdd, neg_mul, mul_neg]
|
||||
ring_nf
|
||||
rfl
|
||||
|
||||
|
@ -80,7 +96,11 @@ def B₄ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₄_cubic (S T : (SM 3).Charges) : cubeTriLin B₄ S T =
|
||||
(S (12 : Fin 18) * T (12 : Fin 18) - S (13 : Fin 18) * T (13 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₄, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₄, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
zero_mul, zero_add, Fin.reduceFinMk, one_mul, Fin.mk_one, Nat.reduceAdd, neg_mul]
|
||||
ring_nf
|
||||
rfl
|
||||
|
||||
|
@ -93,7 +113,11 @@ def B₅ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₅_cubic (S T : (SM 3).Charges) : cubeTriLin B₅ S T =
|
||||
(S (15 : Fin 18) * T (15 : Fin 18) - S (16 : Fin 18) * T (16 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₅, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₅, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
zero_mul, zero_add, Fin.reduceFinMk, one_mul, Fin.mk_one, Nat.reduceAdd, neg_mul]
|
||||
ring_nf
|
||||
rfl
|
||||
|
||||
|
@ -106,7 +130,11 @@ def B₆ : (SM 3).Charges := toSpeciesEquiv.invFun (fun s => fun i =>
|
|||
|
||||
lemma B₆_cubic (S T : (SM 3).Charges) : cubeTriLin B₆ S T =
|
||||
3 * (S (5 : Fin 18) * T (5 : Fin 18) - S (8 : Fin 18) * T (8 : Fin 18)) := by
|
||||
simp [Fin.sum_univ_three, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
simp only [B₆, Equiv.invFun_as_coe, cubeTriLin_toFun_apply_apply, Nat.reduceMul, finProdFinEquiv,
|
||||
Fin.divNat, Fin.modNat, Fin.isValue, Equiv.coe_fn_mk, Fin.val_zero, mul_zero, add_zero,
|
||||
toSpeciesEquiv_symm_apply, Fin.val_one, mul_one, Nat.ofNat_pos, Nat.add_div_right,
|
||||
Nat.add_mod_right, Fin.val_two, Nat.add_mul_mod_self_left, Fin.sum_univ_three, Fin.zero_eta,
|
||||
zero_mul, zero_add, Fin.reduceFinMk, Fin.mk_one, Nat.reduceAdd, one_mul, neg_mul, mul_neg]
|
||||
ring_nf
|
||||
|
||||
/-- The charge assignments forming a basis of the plane. -/
|
||||
|
@ -126,7 +154,8 @@ lemma B₀_Bi_cubic {i : Fin 7} (hi : 0 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin B₀ (B i) S = 0
|
||||
rw [B₀_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -134,7 +163,8 @@ lemma B₁_Bi_cubic {i : Fin 7} (hi : 1 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin B₁ (B i) S = 0
|
||||
rw [B₁_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -142,7 +172,8 @@ lemma B₂_Bi_cubic {i : Fin 7} (hi : 2 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin B₂ (B i) S = 0
|
||||
rw [B₂_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -150,7 +181,8 @@ lemma B₃_Bi_cubic {i : Fin 7} (hi : 3 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₃) (B i) S = 0
|
||||
rw [B₃_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -158,7 +190,8 @@ lemma B₄_Bi_cubic {i : Fin 7} (hi : 4 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₄) (B i) S = 0
|
||||
rw [B₄_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -166,7 +199,8 @@ lemma B₅_Bi_cubic {i : Fin 7} (hi : 5 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₅) (B i) S = 0
|
||||
rw [B₅_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).Charges) :
|
||||
|
@ -174,7 +208,8 @@ lemma B₆_Bi_cubic {i : Fin 7} (hi : 6 ≠ i) (S : (SM 3).Charges) :
|
|||
change cubeTriLin (B₆) (B i) S = 0
|
||||
rw [B₆_cubic]
|
||||
fin_cases i <;>
|
||||
simp at hi <;>
|
||||
simp only [Fin.isValue, Fin.zero_eta, ne_eq, Fin.reduceEq, not_false_eq_true, Fin.mk_one,
|
||||
Fin.reduceFinMk, not_true_eq_false] at hi <;>
|
||||
simp [B₀, B₁, B₂, B₃, B₄, B₅, B₆, Fin.divNat, Fin.modNat, finProdFinEquiv]
|
||||
|
||||
lemma Bi_Bj_ne_cubic {i j : Fin 7} (h : i ≠ j) (S : (SM 3).Charges) :
|
||||
|
@ -245,7 +280,7 @@ lemma Bi_Bj_Bk_cubic (i j k : Fin 7) :
|
|||
cubeTriLin (B i) (B j) (B k) = 0 := by
|
||||
by_cases hij : i ≠ j
|
||||
· exact Bi_Bj_ne_cubic hij (B k)
|
||||
· simp at hij
|
||||
· simp only [ne_eq, Decidable.not_not] at hij
|
||||
rw [hij]
|
||||
exact Bi_Bi_Bj_cubic j k
|
||||
|
||||
|
@ -293,9 +328,12 @@ theorem basis_linear_independent : LinearIndependent ℚ B := by
|
|||
have h5 := congrFun h (15 : Fin 18)
|
||||
have h6 := congrFun h (5 : Fin 18)
|
||||
rw [@Fin.sum_univ_seven] at h0 h1 h2 h3 h4 h5 h6
|
||||
simp [HSMul.hSMul] at h0 h1 h2 h3 h4 h5 h6
|
||||
simp only [HSMul.hSMul, Fin.isValue, B, ACCSystemCharges.chargesAddCommMonoid_add,
|
||||
ACCSystemCharges.chargesModule_smul] at h0 h1 h2 h3 h4 h5 h6
|
||||
rw [B₀, B₁, B₂, B₃, B₄, B₅, B₆] at h0 h1 h2 h3 h4 h5 h6
|
||||
simp [Fin.divNat, Fin.modNat] at h0 h1 h2 h3 h4 h5 h6
|
||||
simp only [Fin.isValue, Equiv.invFun_as_coe, toSpeciesEquiv_symm_apply, Fin.divNat, Nat.reduceMul,
|
||||
Fin.val_zero, Nat.zero_div, Fin.zero_eta, Fin.modNat, Nat.zero_mod, mul_one, mul_zero, add_zero,
|
||||
zero_add] at h0 h1 h2 h3 h4 h5 h6
|
||||
intro i
|
||||
match i with
|
||||
| 0 => exact h0
|
||||
|
|
|
@ -35,7 +35,7 @@ def BL₁ : (PlusU1 1).Sols where
|
|||
| (5 : Fin 6) => 3
|
||||
linearSol := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
| 1 => rfl
|
||||
|
@ -43,7 +43,7 @@ def BL₁ : (PlusU1 1).Sols where
|
|||
| 3 => rfl
|
||||
quadSol := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
cubicSol := by rfl
|
||||
|
|
|
@ -40,27 +40,27 @@ variable {n : ℕ}
|
|||
|
||||
lemma gravSol (S : (PlusU1 n).LinSols) : accGrav S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma SU2Sol (S : (PlusU1 n).LinSols) : accSU2 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 1
|
||||
|
||||
lemma SU3Sol (S : (PlusU1 n).LinSols) : accSU3 S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 2
|
||||
|
||||
lemma YYsol (S : (PlusU1 n).LinSols) : accYY S.val = 0 := by
|
||||
have hS := S.linearSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberLinear, PlusU1_linearACCs, Fin.isValue] at hS
|
||||
exact hS 3
|
||||
|
||||
lemma quadSol (S : (PlusU1 n).QuadSols) : accQuad S.val = 0 := by
|
||||
have hS := S.quadSol
|
||||
simp at hS
|
||||
simp only [PlusU1_numberQuadratic, HomogeneousQuadratic.eq_1, PlusU1_quadraticACCs] at hS
|
||||
exact hS 0
|
||||
|
||||
lemma cubeSol (S : (PlusU1 n).Sols) : accCube S.val = 0 := by
|
||||
|
@ -73,7 +73,7 @@ def chargeToLinear (S : (PlusU1 n).Charges) (hGrav : accGrav S = 0)
|
|||
(PlusU1 n).LinSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact hGrav
|
||||
| 1 => exact hSU2
|
||||
|
@ -86,7 +86,7 @@ def linearToQuad (S : (PlusU1 n).LinSols) (hQ : accQuad S.val = 0) :
|
|||
(PlusU1 n).QuadSols :=
|
||||
⟨S, by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => exact hQ⟩
|
||||
|
||||
|
@ -122,7 +122,7 @@ def perm (n : ℕ) : ACCSystemGroupAction (PlusU1 n) where
|
|||
rep := repCharges
|
||||
linearInvariant := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => exact accGrav_invariant
|
||||
| 1 => exact accSU2_invariant
|
||||
|
@ -130,7 +130,7 @@ def perm (n : ℕ) : ACCSystemGroupAction (PlusU1 n) where
|
|||
| 3 => exact accYY_invariant
|
||||
quadInvariant := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => exact accQuad_invariant
|
||||
cubicInvariant := accCube_invariant
|
||||
|
|
|
@ -33,7 +33,7 @@ def Y₁ : (PlusU1 1).Sols where
|
|||
| (5 : Fin 6) => 0
|
||||
linearSol := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberLinear] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
| 1 => rfl
|
||||
|
@ -41,7 +41,7 @@ def Y₁ : (PlusU1 1).Sols where
|
|||
| 3 => rfl
|
||||
quadSol := by
|
||||
intro i
|
||||
simp at i
|
||||
simp only [PlusU1_numberQuadratic] at i
|
||||
match i with
|
||||
| 0 => rfl
|
||||
cubicSol := by rfl
|
||||
|
|
|
@ -162,7 +162,7 @@ lemma isSolution_sum_part (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑
|
|||
rw [isSolution_f0 f hS, isSolution_f1 f hS, isSolution_f2 f hS, isSolution_f3 f hS,
|
||||
isSolution_f4 f hS, isSolution_f5 f hS,
|
||||
isSolution_f6 f hS, isSolution_f7 f hS, isSolution_f8 f hS]
|
||||
simp
|
||||
simp only [Fin.isValue, zero_smul, add_zero, zero_add]
|
||||
rfl
|
||||
|
||||
lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
||||
|
@ -172,7 +172,7 @@ lemma isSolution_grav (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f
|
|||
have hg := gravSol S.toLinSols
|
||||
rw [hS', hx, accGrav.map_add, accGrav.map_smul, accGrav.map_smul, show accGrav B₉ = 3 by rfl,
|
||||
show accGrav B₁₀ = 1 by rfl] at hg
|
||||
simp at hg
|
||||
simp only [Fin.isValue, smul_eq_mul, mul_one] at hg
|
||||
linear_combination hg
|
||||
|
||||
lemma isSolution_sum_part' (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i • B i)) :
|
||||
|
@ -193,7 +193,7 @@ lemma isSolution_f9 (f : Fin 11 → ℚ) (hS : (PlusU1 3).IsSolution (∑ i, f i
|
|||
cubeTriLin.map_smul₃, cubeTriLin.map_smul₃] at hc
|
||||
rw [show accCube B₉ = 9 by rfl, show accCube B₁₀ = 1 by rfl, show cubeTriLin B₉ B₉ B₁₀ = 0 by rfl,
|
||||
show cubeTriLin B₁₀ B₁₀ B₉ = 0 by rfl] at hc
|
||||
simp at hc
|
||||
simp only [Fin.isValue, neg_mul, mul_one, mul_zero, add_zero] at hc
|
||||
have h1 : f 9 ^ 3 * 9 + (-(3 * f 9)) ^ 3 = - 18 * f 9 ^ 3 := by ring
|
||||
rw [h1] at hc
|
||||
simpa using hc
|
||||
|
|
|
@ -55,7 +55,9 @@ lemma accQuad_α₁_α₂ (S : (PlusU1 n).LinSols) :
|
|||
lemma accQuad_α₁_α₂_zero (S : (PlusU1 n).LinSols) (h1 : α₁ C S = 0)
|
||||
(h2 : α₂ S = 0) (a b : ℚ) : accQuad (a • S + b • C.1).val = 0 := by
|
||||
erw [add_AFL_quad]
|
||||
simp [α₁, α₂] at h1 h2
|
||||
simp only [α₁, quadBiLin_toFun_apply, Fin.isValue, neg_mul, neg_eq_zero, mul_eq_zero,
|
||||
OfNat.ofNat_ne_zero, false_or, α₂, HomogeneousQuadratic.eq_1, accQuad,
|
||||
BiLinearSymm.toHomogeneousQuad_apply] at h1 h2
|
||||
field_simp [h1, h2]
|
||||
|
||||
/-- The construction of a `QuadSol` from a `LinSols` in the generic case. -/
|
||||
|
|
|
@ -118,14 +118,14 @@ lemma quadSolToSolInv_α₁_α₂_neq_zero (S : (PlusU1 n).Sols) (h : α₁ S.1
|
|||
lemma quadSolToSolInv_special (S : (PlusU1 n).Sols) (h : α₁ S.1 = 0) :
|
||||
special (quadSolToSolInv S).1 (quadSolToSolInv S).2.1 (quadSolToSolInv S).2.2
|
||||
(quadSolToSolInv_α₁_α₂_zero S h).1 (quadSolToSolInv_α₁_α₂_zero S h).2 = S := by
|
||||
simp [quadSolToSolInv_1]
|
||||
simp only [quadSolToSolInv_1]
|
||||
rw [show (quadSolToSolInv S).2.1 = 1 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [show (quadSolToSolInv S).2.2 = 0 by rw [quadSolToSolInv, if_pos h]]
|
||||
rw [special_on_AF]
|
||||
|
||||
lemma quadSolToSolInv_generic (S : (PlusU1 n).Sols) (h : α₁ S.1 ≠ 0) :
|
||||
(quadSolToSolInv S).2.1 • generic (quadSolToSolInv S).1 = S := by
|
||||
simp [quadSolToSolInv_1]
|
||||
simp only [quadSolToSolInv_1]
|
||||
rw [show (quadSolToSolInv S).2.1 = (α₁ S.1)⁻¹ by rw [quadSolToSolInv, if_neg h]]
|
||||
rw [generic_on_AF_α₁_ne_zero S h]
|
||||
|
||||
|
|
|
@ -45,7 +45,11 @@ lemma phaseShiftMatrix_star (a b c : ℝ) :
|
|||
(phaseShiftMatrix a b c)ᴴ = phaseShiftMatrix (- a) (- b) (- c) := by
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [← exp_conj, conj_ofReal]
|
||||
simp only [phaseShiftMatrix, Fin.zero_eta, Fin.isValue, conjTranspose_apply, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, vecCons_const, star_def, ← exp_conj,
|
||||
_root_.map_mul, conj_I, conj_ofReal, neg_mul, ofReal_neg, mul_neg, Fin.mk_one, cons_val_one,
|
||||
head_fin_const, star_zero, head_cons, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd,
|
||||
tail_val', head_val', tail_cons, Fin.reduceFinMk, map_eq_zero]
|
||||
· rfl
|
||||
· rfl
|
||||
|
||||
|
@ -53,7 +57,10 @@ lemma phaseShiftMatrix_mul (a b c d e f : ℝ) :
|
|||
phaseShiftMatrix a b c * phaseShiftMatrix d e f = phaseShiftMatrix (a + d) (b + e) (c + f) := by
|
||||
ext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three]
|
||||
simp only [phaseShiftMatrix, Fin.zero_eta, Fin.isValue, mul_apply, cons_val', empty_val',
|
||||
cons_val_fin_one, cons_val_zero, vecCons_const, Fin.sum_univ_three, cons_val_one, head_cons,
|
||||
head_fin_const, mul_zero, add_zero, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons,
|
||||
tail_val', head_val', zero_mul, ofReal_add, Fin.mk_one, Fin.reduceFinMk, zero_add]
|
||||
any_goals rw [mul_add, exp_add]
|
||||
change cexp (I * ↑c) * 0 = 0
|
||||
simp
|
||||
|
@ -116,8 +123,8 @@ lemma phaseShiftRelation_equiv : Equivalence PhaseShiftRelation where
|
|||
def CKMMatrix : Type := unitaryGroup (Fin 3) ℂ
|
||||
|
||||
lemma CKMMatrix_ext {U V : CKMMatrix} (h : U.val = V.val) : U = V := by
|
||||
cases U; cases V
|
||||
simp at h
|
||||
cases U
|
||||
cases V
|
||||
subst h
|
||||
rfl
|
||||
|
||||
|
@ -286,7 +293,11 @@ lemma VAbs'_equiv (i j : Fin 3) (V U : CKMMatrix) (h : V ≈ U) :
|
|||
vecCons_const, cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons,
|
||||
head_fin_const, mul_zero]
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [Complex.abs_exp]
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, zero_mul, add_zero, mul_zero,
|
||||
_root_.map_mul, abs_exp, mul_re, I_re, ofReal_re, I_im, ofReal_im, sub_self, Real.exp_zero,
|
||||
one_mul, mul_one, Fin.mk_one, cons_val_one, head_cons, zero_add, head_fin_const,
|
||||
Fin.reduceFinMk, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons,
|
||||
tail_val', head_val']
|
||||
all_goals change Complex.abs (0 * _ + _) = _
|
||||
all_goals simp [Complex.abs_exp]
|
||||
|
||||
|
|
|
@ -62,7 +62,8 @@ lemma shift_ub_phase_zero {V : CKMMatrix} (h1 : u + b = - arg [V]ub) :
|
|||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VubAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cs_phase_zero {V : CKMMatrix} (h1 : c + s = - arg [V]cs) :
|
||||
|
@ -71,10 +72,11 @@ lemma shift_cs_phase_zero {V : CKMMatrix} (h1 : c + s = - arg [V]cs) :
|
|||
rw [← abs_mul_exp_arg_mul_I [V]cs]
|
||||
rw [mul_comm, mul_assoc, ← exp_add]
|
||||
have h2 : ↑(arg [V]cs) * I + (↑c * I + ↑s * I) = ↑(arg [V]cs + (c + s)) * I := by
|
||||
simp [add_assoc]
|
||||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VcsAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cb_phase_zero {V : CKMMatrix} (h1 : c + b = - arg [V]cb) :
|
||||
|
@ -83,10 +85,11 @@ lemma shift_cb_phase_zero {V : CKMMatrix} (h1 : c + b = - arg [V]cb) :
|
|||
rw [← abs_mul_exp_arg_mul_I [V]cb]
|
||||
rw [mul_comm, mul_assoc, ← exp_add]
|
||||
have h2 : ↑(arg [V]cb) * I + (↑c * I + ↑b * I) = ↑(arg [V]cb + (c + b)) * I := by
|
||||
simp [add_assoc]
|
||||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VcbAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_tb_phase_zero {V : CKMMatrix} (h1 : t + b = - arg [V]tb) :
|
||||
|
@ -95,10 +98,11 @@ lemma shift_tb_phase_zero {V : CKMMatrix} (h1 : t + b = - arg [V]tb) :
|
|||
rw [← abs_mul_exp_arg_mul_I [V]tb]
|
||||
rw [mul_comm, mul_assoc, ← exp_add]
|
||||
have h2 : ↑(arg [V]tb) * I + (↑t * I + ↑b * I) = ↑(arg [V]tb + (t + b)) * I := by
|
||||
simp [add_assoc]
|
||||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_neg_cancel, ofReal_zero, zero_mul, exp_zero, mul_one, VtbAbs,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cd_phase_pi {V : CKMMatrix} (h1 : c + d = Real.pi - arg [V]cd) :
|
||||
|
@ -107,10 +111,11 @@ lemma shift_cd_phase_pi {V : CKMMatrix} (h1 : c + d = Real.pi - arg [V]cd) :
|
|||
rw [← abs_mul_exp_arg_mul_I [V]cd]
|
||||
rw [mul_comm, mul_assoc, ← exp_add]
|
||||
have h2 : ↑(arg [V]cd) * I + (↑c * I + ↑d * I) = ↑(arg [V]cd + (c + d)) * I := by
|
||||
simp [add_assoc]
|
||||
simp only [Fin.isValue, ofReal_add]
|
||||
ring
|
||||
rw [h2, h1]
|
||||
simp
|
||||
simp only [Fin.isValue, add_sub_cancel, exp_pi_mul_I, mul_neg, mul_one, VcdAbs, neg_inj,
|
||||
ofReal_inj]
|
||||
rfl
|
||||
|
||||
lemma shift_cross_product_phase_zero {V : CKMMatrix} {τ : ℝ}
|
||||
|
@ -120,33 +125,35 @@ lemma shift_cross_product_phase_zero {V : CKMMatrix} {τ : ℝ}
|
|||
change _ = phaseShiftApply.ucCross _ _ _ _ _ _ _
|
||||
funext i
|
||||
fin_cases i
|
||||
· simp
|
||||
· simp only [Fin.zero_eta, Fin.isValue]
|
||||
rw [phaseShiftApply.ucCross_fst]
|
||||
simp [tRow, phaseShiftApply.td]
|
||||
simp only [tRow, Fin.isValue, phaseShiftApply.td, phaseShiftApply_coe, cons_val_zero, neg_mul]
|
||||
have hτ0 := congrFun hτ 0
|
||||
simp [tRow] at hτ0
|
||||
simp only [Fin.isValue, Pi.smul_apply, smul_eq_mul, tRow, cons_val_zero] at hτ0
|
||||
rw [← hτ0]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
congr 2
|
||||
simp only [ofReal_sub, ofReal_neg]
|
||||
ring
|
||||
· simp
|
||||
· simp only [Fin.mk_one, Fin.isValue]
|
||||
rw [phaseShiftApply.ucCross_snd]
|
||||
simp [tRow, phaseShiftApply.ts]
|
||||
simp only [tRow, Fin.isValue, phaseShiftApply_coe, phaseShiftApply.ts, cons_val_one, head_cons,
|
||||
neg_mul]
|
||||
have hτ0 := congrFun hτ 1
|
||||
simp [tRow] at hτ0
|
||||
simp only [Fin.isValue, Pi.smul_apply, smul_eq_mul, tRow, cons_val_one, head_cons] at hτ0
|
||||
rw [← hτ0]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
congr 2
|
||||
simp only [ofReal_sub, ofReal_neg]
|
||||
ring
|
||||
· simp
|
||||
· simp only [Fin.reduceFinMk, Fin.isValue]
|
||||
rw [phaseShiftApply.ucCross_thd]
|
||||
simp [tRow, phaseShiftApply.tb]
|
||||
simp only [tRow, Fin.isValue, phaseShiftApply_coe, phaseShiftApply.tb, cons_val_two,
|
||||
Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, head_cons, neg_mul]
|
||||
have hτ0 := congrFun hτ 2
|
||||
simp [tRow] at hτ0
|
||||
rw [← hτ0]
|
||||
rw [← mul_assoc, ← exp_add, h1]
|
||||
simp only [Fin.isValue, Pi.smul_apply, smul_eq_mul, tRow, cons_val_two, Nat.succ_eq_add_one,
|
||||
Nat.reduceAdd, tail_cons, head_cons] at hτ0
|
||||
rw [← hτ0, ← mul_assoc, ← exp_add, h1]
|
||||
congr 2
|
||||
simp only [ofReal_sub, ofReal_neg]
|
||||
ring
|
||||
|
@ -271,26 +278,26 @@ lemma ubOnePhaseCond_hold_up_to_equiv_of_ub_one {V : CKMMatrix} (hb : ¬ ([V]ud
|
|||
apply And.intro
|
||||
· exact phaseShiftApply.equiv _ _ _ _ _ _ _
|
||||
· apply And.intro
|
||||
· simp [not_or] at hb
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] at hb
|
||||
have h1 : VudAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, hb]
|
||||
simp [VAbs] at h1
|
||||
simp only [VudAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk, map_eq_zero] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] at hb
|
||||
have h1 : VusAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, hb]
|
||||
simp [VAbs] at h1
|
||||
simp only [VusAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk, map_eq_zero] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· simp [not_or] at hb
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] at hb
|
||||
have h3 := cb_eq_zero_of_ud_us_zero hb
|
||||
have h1 : VcbAbs ⟦U⟧ = 0 := by
|
||||
rw [hUV]
|
||||
simp [VAbs, h3]
|
||||
simp [VAbs] at h1
|
||||
simp only [VcbAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk, map_eq_zero] at h1
|
||||
exact h1
|
||||
apply And.intro
|
||||
· have hU1 : [U]ub = VubAbs ⟦V⟧ := by
|
||||
|
@ -325,7 +332,8 @@ lemma cd_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠
|
|||
exact hV.2.2.2.2
|
||||
rw [cd_of_ud_us_ub_cb_tb hb hτ]
|
||||
rw [hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
|
||||
simp [sq, conj_ofReal]
|
||||
simp only [Fin.isValue, VudAbs, VcbAbs, ofReal_zero, zero_mul, exp_zero, VtbAbs, conj_ofReal,
|
||||
one_mul, VusAbs, neg_add_rev, normSq_ofReal, ofReal_mul, neg_mul, sq, VubAbs]
|
||||
have hx := Vabs_sq_add_neq_zero hb
|
||||
field_simp
|
||||
have h1 : conj [V]ub = VubAbs ⟦V⟧ * cexp (- arg [V]ub * I) := by
|
||||
|
@ -343,15 +351,16 @@ lemma cs_of_fstRowThdColRealCond {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠
|
|||
have hτ : [V]t = cexp ((0 : ℝ) * I) • (conj ([V]u) ×₃ conj ([V]c)) := by
|
||||
simp only [ofReal_zero, zero_mul, exp_zero, one_smul]
|
||||
exact hV.2.2.2.2
|
||||
rw [cs_of_ud_us_ub_cb_tb hb hτ]
|
||||
rw [hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
|
||||
simp [sq, conj_ofReal]
|
||||
rw [cs_of_ud_us_ub_cb_tb hb hτ, hV.1, hV.2.1, hV.2.2.1, hV.2.2.2.1]
|
||||
simp only [Fin.isValue, VusAbs, neg_mul, VcbAbs, ofReal_zero, zero_mul, exp_zero, VtbAbs,
|
||||
conj_ofReal, one_mul, VudAbs, normSq_ofReal, ofReal_mul, sq, VubAbs]
|
||||
have hx := Vabs_sq_add_neq_zero hb
|
||||
field_simp
|
||||
have h1 : conj [V]ub = VubAbs ⟦V⟧ * cexp (- arg [V]ub * I) := by
|
||||
nth_rewrite 1 [← abs_mul_exp_arg_mul_I [V]ub]
|
||||
rw [@RingHom.map_mul]
|
||||
simp [conj_ofReal, ← exp_conj, VAbs]
|
||||
simp only [Fin.isValue, conj_ofReal, ← exp_conj, _root_.map_mul, conj_I, mul_neg, VubAbs, VAbs,
|
||||
VAbs', Quotient.lift_mk, neg_mul]
|
||||
rw [h1]
|
||||
ring_nf
|
||||
|
||||
|
|
|
@ -30,7 +30,8 @@ lemma VAbs_sum_sq_row_eq_one (V : Quotient CKMMatrixSetoid) (i : Fin 3) :
|
|||
have hV := V.prop
|
||||
rw [mem_unitaryGroup_iff] at hV
|
||||
have ht := congrFun (congrFun hV i) i
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
|
||||
simp only [mul_apply, star_apply, RCLike.star_def, Fin.sum_univ_three, Fin.isValue,
|
||||
one_apply_eq] at ht
|
||||
rw [mul_conj, mul_conj, mul_conj] at ht
|
||||
repeat rw [← Complex.sq_abs] at ht
|
||||
rw [← ofReal_inj]
|
||||
|
@ -72,7 +73,7 @@ lemma ud_us_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
|
|||
have h2 := V.fst_row_normalized_abs
|
||||
refine Iff.intro (fun h h1 => ?_) (fun h => ?_)
|
||||
· rw [h1] at h2
|
||||
simp at h2
|
||||
simp only [Fin.isValue, one_pow, add_left_eq_self] at h2
|
||||
rw [add_eq_zero_iff_of_nonneg (sq_nonneg _) (sq_nonneg _)] at h2
|
||||
simp_all
|
||||
· by_contra hn
|
||||
|
@ -92,7 +93,7 @@ lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℝ {V : CKMMatrix} (hb : [V]ud ≠ 0
|
|||
rw [← Complex.sq_abs] at hn
|
||||
have h2 : Complex.abs (V.1 0 2) ^2 = 1 := by
|
||||
linear_combination -(1 * hn)
|
||||
simp at h2
|
||||
simp only [Fin.isValue, sq_eq_one_iff] at h2
|
||||
cases' h2 with h2 h2
|
||||
· exact hb h2
|
||||
· have h3 := Complex.abs.nonneg [V]ub
|
||||
|
@ -120,7 +121,7 @@ lemma VAbsub_neq_zero_sqrt_Vud_Vus_neq_zero {V : Quotient CKMMatrixSetoid}
|
|||
lemma normSq_Vud_plus_normSq_Vus_neq_zero_ℂ {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) :
|
||||
(normSq [V]ud : ℂ) + normSq [V]us ≠ 0 := by
|
||||
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℝ hb
|
||||
simp at h1
|
||||
simp only [Fin.isValue, ne_eq] at h1
|
||||
rw [← ofReal_inj] at h1
|
||||
simp_all
|
||||
|
||||
|
@ -128,7 +129,7 @@ lemma Vabs_sq_add_neq_zero {V : CKMMatrix} (hb : [V]ud ≠ 0 ∨ [V]us ≠ 0) :
|
|||
((VudAbs ⟦V⟧ : ℂ) * ↑(VudAbs ⟦V⟧) + ↑(VusAbs ⟦V⟧) * ↑(VusAbs ⟦V⟧)) ≠ 0 := by
|
||||
have h1 := normSq_Vud_plus_normSq_Vus_neq_zero_ℂ hb
|
||||
rw [← Complex.sq_abs, ← Complex.sq_abs] at h1
|
||||
simp [sq] at h1
|
||||
simp only [Fin.isValue, sq, ofReal_mul, ne_eq] at h1
|
||||
exact h1
|
||||
|
||||
section orthogonal
|
||||
|
@ -138,7 +139,8 @@ lemma fst_row_orthog_snd_row (V : CKMMatrix) :
|
|||
have hV := V.prop
|
||||
rw [mem_unitaryGroup_iff] at hV
|
||||
have ht := congrFun (congrFun hV 1) 0
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
|
||||
simp only [Fin.isValue, mul_apply, star_apply, RCLike.star_def, Fin.sum_univ_three, ne_eq,
|
||||
one_ne_zero, not_false_eq_true, one_apply_ne] at ht
|
||||
exact ht
|
||||
|
||||
lemma fst_row_orthog_thd_row (V : CKMMatrix) :
|
||||
|
@ -146,7 +148,8 @@ lemma fst_row_orthog_thd_row (V : CKMMatrix) :
|
|||
have hV := V.prop
|
||||
rw [mem_unitaryGroup_iff] at hV
|
||||
have ht := congrFun (congrFun hV 2) 0
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
|
||||
simp only [Fin.isValue, mul_apply, star_apply, RCLike.star_def, Fin.sum_univ_three, ne_eq,
|
||||
Fin.reduceEq, not_false_eq_true, one_apply_ne] at ht
|
||||
exact ht
|
||||
|
||||
lemma Vcd_mul_conj_Vud (V : CKMMatrix) :
|
||||
|
@ -163,14 +166,14 @@ lemma VAbs_thd_eq_one_fst_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3} (hV
|
|||
VAbs i 0 V = 0 := by
|
||||
have h := VAbs_sum_sq_row_eq_one V i
|
||||
rw [hV] at h
|
||||
simp at h
|
||||
simp only [Fin.isValue, one_pow, add_left_eq_self] at h
|
||||
nlinarith
|
||||
|
||||
lemma VAbs_thd_eq_one_snd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3} (hV : VAbs i 2 V = 1) :
|
||||
VAbs i 1 V = 0 := by
|
||||
have h := VAbs_sum_sq_row_eq_one V i
|
||||
rw [hV] at h
|
||||
simp at h
|
||||
simp only [Fin.isValue, one_pow, add_left_eq_self] at h
|
||||
nlinarith
|
||||
|
||||
section crossProduct
|
||||
|
@ -179,9 +182,12 @@ lemma conj_Vtb_cross_product {V : CKMMatrix} {τ : ℝ}
|
|||
(hτ : [V]t = cexp (τ * I) • (conj [V]u ×₃ conj [V]c)) :
|
||||
conj [V]tb = cexp (- τ * I) * ([V]cs * [V]ud - [V]us * [V]cd) := by
|
||||
have h1 := congrFun hτ 2
|
||||
simp [crossProduct, tRow, uRow, cRow] at h1
|
||||
simp only [tRow, Fin.isValue, cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons,
|
||||
head_cons, crossProduct, uRow, cRow, LinearMap.mk₂_apply, Pi.conj_apply, cons_val_one,
|
||||
cons_val_zero, Pi.smul_apply, smul_eq_mul] at h1
|
||||
apply congrArg conj at h1
|
||||
simp at h1
|
||||
simp only [Fin.isValue, _root_.map_mul, map_sub, RingHomCompTriple.comp_apply,
|
||||
RingHom.id_apply] at h1
|
||||
rw [h1]
|
||||
simp only [← exp_conj, _root_.map_mul, conj_ofReal, conj_I, mul_neg, Fin.isValue, neg_mul]
|
||||
field_simp
|
||||
|
@ -194,7 +200,8 @@ lemma conj_Vtb_mul_Vud {V : CKMMatrix} {τ : ℝ}
|
|||
cexp (τ * I) * conj [V]tb * conj [V]ud =
|
||||
[V]cs * (normSq [V]ud + normSq [V]us) + [V]cb * conj [V]ub * [V]us := by
|
||||
rw [conj_Vtb_cross_product hτ]
|
||||
simp [exp_neg]
|
||||
simp only [neg_mul, exp_neg, Fin.isValue, ne_eq, exp_ne_zero, not_false_eq_true,
|
||||
mul_inv_cancel_left₀]
|
||||
have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]ud = [V]cs
|
||||
* [V]ud * conj [V]ud - [V]us * ([V]cd * conj [V]ud) := by
|
||||
ring
|
||||
|
@ -208,7 +215,8 @@ lemma conj_Vtb_mul_Vus {V : CKMMatrix} {τ : ℝ}
|
|||
cexp (τ * I) * conj [V]tb * conj [V]us =
|
||||
- ([V]cd * (normSq [V]ud + normSq [V]us) + [V]cb * conj ([V]ub) * [V]ud) := by
|
||||
rw [conj_Vtb_cross_product hτ]
|
||||
simp [exp_neg]
|
||||
simp only [neg_mul, exp_neg, Fin.isValue, ne_eq, exp_ne_zero, not_false_eq_true,
|
||||
mul_inv_cancel_left₀, neg_add_rev]
|
||||
have h2 : ([V]cs * [V]ud - [V]us * [V]cd) * conj [V]us = ([V]cs
|
||||
* conj [V]us) * [V]ud - [V]us * [V]cd * conj [V]us := by
|
||||
ring
|
||||
|
@ -265,7 +273,8 @@ lemma VAbs_sum_sq_col_eq_one (V : Quotient CKMMatrixSetoid) (i : Fin 3) :
|
|||
have hV := V.prop
|
||||
rw [mem_unitaryGroup_iff'] at hV
|
||||
have ht := congrFun (congrFun hV i) i
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
|
||||
simp only [mul_apply, star_apply, RCLike.star_def, Fin.sum_univ_three, Fin.isValue,
|
||||
one_apply_eq] at ht
|
||||
rw [mul_comm, mul_conj, mul_comm, mul_conj, mul_comm, mul_conj] at ht
|
||||
repeat rw [← Complex.sq_abs] at ht
|
||||
rw [← ofReal_inj]
|
||||
|
@ -274,7 +283,7 @@ lemma VAbs_sum_sq_col_eq_one (V : Quotient CKMMatrixSetoid) (i : Fin 3) :
|
|||
lemma thd_col_normalized_abs (V : CKMMatrix) :
|
||||
abs [V]ub ^ 2 + abs [V]cb ^ 2 + abs [V]tb ^ 2 = 1 := by
|
||||
have h1 := VAbs_sum_sq_col_eq_one ⟦V⟧ 2
|
||||
simp [VAbs] at h1
|
||||
simp only [VAbs, VAbs', Fin.isValue, Quotient.lift_mk] at h1
|
||||
exact h1
|
||||
|
||||
lemma thd_col_normalized_normSq (V : CKMMatrix) :
|
||||
|
@ -287,11 +296,13 @@ lemma cb_eq_zero_of_ud_us_zero {V : CKMMatrix} (h : [V]ud = 0 ∧ [V]us = 0) :
|
|||
[V]cb = 0 := by
|
||||
have h1 := fst_row_normalized_abs V
|
||||
rw [← thd_col_normalized_abs V] at h1
|
||||
simp [h] at h1
|
||||
simp only [Fin.isValue, h, map_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow,
|
||||
add_zero, zero_add] at h1
|
||||
rw [add_assoc] at h1
|
||||
simp at h1
|
||||
simp only [Fin.isValue, self_eq_add_right] at h1
|
||||
rw [add_eq_zero_iff_of_nonneg (sq_nonneg _) (sq_nonneg _)] at h1
|
||||
simp at h1
|
||||
simp only [Fin.isValue, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, pow_eq_zero_iff,
|
||||
map_eq_zero] at h1
|
||||
exact h1.1
|
||||
|
||||
lemma cs_of_ud_us_zero {V : CKMMatrix} (ha : ¬ ([V]ud ≠ 0 ∨ [V]us ≠ 0)) :
|
||||
|
@ -299,9 +310,10 @@ lemma cs_of_ud_us_zero {V : CKMMatrix} (ha : ¬ ([V]ud ≠ 0 ∨ [V]us ≠ 0)) :
|
|||
have h1 := snd_row_normalized_abs V
|
||||
symm
|
||||
rw [Real.sqrt_eq_iff_eq_sq]
|
||||
· simp [not_or] at ha
|
||||
· simp only [Fin.isValue, ne_eq, not_or, Decidable.not_not] at ha
|
||||
rw [cb_eq_zero_of_ud_us_zero ha] at h1
|
||||
simp at h1
|
||||
simp only [Fin.isValue, map_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow,
|
||||
add_zero] at h1
|
||||
simp only [VcdAbs, VAbs, VAbs', Fin.isValue, Quotient.lift_mk, VcsAbs]
|
||||
rw [← h1]
|
||||
ring
|
||||
|
@ -323,14 +335,14 @@ lemma cb_tb_neq_zero_iff_ub_neq_one (V : CKMMatrix) :
|
|||
have h2 := V.thd_col_normalized_abs
|
||||
refine Iff.intro (fun h h1 => ?_) (fun h => ?_)
|
||||
· rw [h1] at h2
|
||||
simp at h2
|
||||
simp only [one_pow, Fin.isValue] at h2
|
||||
have h2 : Complex.abs (V.1 1 2) ^ 2 + Complex.abs (V.1 2 2) ^ 2 = 0 := by
|
||||
linear_combination h2
|
||||
rw [add_eq_zero_iff_of_nonneg (sq_nonneg _) (sq_nonneg _)] at h2
|
||||
simp_all
|
||||
· by_contra hn
|
||||
rw [not_or] at hn
|
||||
simp at hn
|
||||
simp only [Fin.isValue, ne_eq, Decidable.not_not] at hn
|
||||
simp_all only [map_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, add_zero,
|
||||
sq_eq_one_iff, false_or, one_ne_zero]
|
||||
have h1 := Complex.abs.nonneg [V]ub
|
||||
|
@ -342,14 +354,14 @@ lemma VAbs_fst_col_eq_one_snd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3}
|
|||
(hV : VAbs 0 i V = 1) : VAbs 1 i V = 0 := by
|
||||
have h := VAbs_sum_sq_col_eq_one V i
|
||||
rw [hV] at h
|
||||
simp at h
|
||||
simp only [one_pow, Fin.isValue] at h
|
||||
nlinarith
|
||||
|
||||
lemma VAbs_fst_col_eq_one_thd_eq_zero {V : Quotient CKMMatrixSetoid} {i : Fin 3}
|
||||
(hV : VAbs 0 i V = 1) : VAbs 2 i V = 0 := by
|
||||
have h := VAbs_sum_sq_col_eq_one V i
|
||||
rw [hV] at h
|
||||
simp at h
|
||||
simp only [one_pow, Fin.isValue] at h
|
||||
nlinarith
|
||||
|
||||
end columns
|
||||
|
|
|
@ -47,7 +47,8 @@ lemma uRow_normalized (V : CKMMatrix) : conj [V]u ⬝ᵥ [V]u = 1 := by
|
|||
have hV := V.prop
|
||||
rw [mem_unitaryGroup_iff] at hV
|
||||
have ht := congrFun (congrFun hV 0) 0
|
||||
simp [Matrix.mul_apply, Fin.sum_univ_three] at ht
|
||||
simp only [Fin.isValue, mul_apply, star_apply, RCLike.star_def, Fin.sum_univ_three,
|
||||
one_apply_eq] at ht
|
||||
rw [mul_comm (V.1 0 0) _, mul_comm (V.1 0 1) _, mul_comm (V.1 0 2) _] at ht
|
||||
exact ht
|
||||
|
||||
|
@ -215,7 +216,7 @@ lemma cRow_cross_tRow_eq_uRow (V : CKMMatrix) :
|
|||
fin_cases i <;> simp
|
||||
simp only [h3, Fin.isValue, smul_dotProduct, Pi.conj_apply, smul_eq_mul,
|
||||
uRow_normalized, Fin.isValue, mul_one, mul_conj, ← Complex.sq_abs] at h2
|
||||
simp at h2
|
||||
simp only [Fin.isValue, ofReal_pow, sq_eq_one_iff, ofReal_eq_one] at h2
|
||||
cases' h2 with h2 h2
|
||||
· have hx : [V]u = (g 0)⁻¹ • (conj ([V]c) ×₃ conj ([V]t)) := by
|
||||
rw [← hg, smul_smul, inv_mul_cancel₀, one_smul]
|
||||
|
@ -242,7 +243,7 @@ lemma uRow_cross_cRow_eq_tRow (V : CKMMatrix) :
|
|||
obtain ⟨g, hg⟩ := (mem_span_range_iff_exists_fun ℂ).mp (Basis.mem_span (rowBasis V)
|
||||
(conj ([V]u) ×₃ conj ([V]c)))
|
||||
rw [Fin.sum_univ_three, rowBasis] at hg
|
||||
simp at hg
|
||||
simp only [Fin.isValue, coe_basisOfLinearIndependentOfCardEqFinrank, rows] at hg
|
||||
have h0 := congrArg (fun X => conj [V]u ⬝ᵥ X) hg
|
||||
have h1 := congrArg (fun X => conj [V]c ⬝ᵥ X) hg
|
||||
simp only [Fin.isValue, dotProduct_add, dotProduct_smul, Pi.conj_apply,
|
||||
|
@ -308,21 +309,26 @@ def ucCross : Fin 3 → ℂ :=
|
|||
|
||||
lemma ucCross_fst (V : CKMMatrix) : (ucCross V a b c d e f) 0 =
|
||||
cexp ((- a * I) + (- b * I) + (- e * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 0 := by
|
||||
simp [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue,
|
||||
LinearMap.mk₂_apply, Pi.conj_apply, cons_val_zero, neg_mul, uRow, us, ub, cRow, cs, cb,
|
||||
exp_add, exp_sub, ← exp_conj, conj_ofReal]
|
||||
simp only [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, uRow,
|
||||
phaseShiftApply_coe, us, exp_add, ub, cRow, cs, cb, LinearMap.mk₂_apply, Pi.conj_apply,
|
||||
cons_val_one, head_cons, _root_.map_mul, ← exp_conj, conj_ofReal, conj_I, mul_neg, cons_val_two,
|
||||
tail_cons, cons_val_zero, neg_mul]
|
||||
ring
|
||||
|
||||
lemma ucCross_snd (V : CKMMatrix) : (ucCross V a b c d e f) 1 =
|
||||
cexp ((- a * I) + (- b * I) + (- d * I) + (- f * I)) * (conj [V]u ×₃ conj [V]c) 1 := by
|
||||
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add,
|
||||
exp_sub, ← exp_conj, conj_ofReal]
|
||||
simp only [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, uRow, ud,
|
||||
exp_add, us, ub, cRow, cd, cs, cb, LinearMap.mk₂_apply, Pi.conj_apply, cons_val_one, head_cons,
|
||||
_root_.map_mul, ← exp_conj, conj_ofReal, conj_I, mul_neg, cons_val_two, tail_cons,
|
||||
cons_val_zero, neg_mul]
|
||||
ring
|
||||
|
||||
lemma ucCross_thd (V : CKMMatrix) : (ucCross V a b c d e f) 2 =
|
||||
cexp ((- a * I) + (- b * I) + (- d * I) + (- e * I)) * (conj [V]u ×₃ conj [V]c) 2 := by
|
||||
simp [ucCross, crossProduct, uRow, us, ub, cRow, cs, cb, ud, cd, exp_add, exp_sub,
|
||||
← exp_conj, conj_ofReal]
|
||||
simp only [ucCross, crossProduct, Nat.succ_eq_add_one, Nat.reduceAdd, Fin.isValue, uRow, ud,
|
||||
exp_add, us, ub, cRow, cd, cs, cb, LinearMap.mk₂_apply, Pi.conj_apply, cons_val_one, head_cons,
|
||||
_root_.map_mul, ← exp_conj, conj_ofReal, conj_I, mul_neg, cons_val_two, tail_cons,
|
||||
cons_val_zero, neg_mul]
|
||||
ring
|
||||
|
||||
lemma uRow_mul (V : CKMMatrix) (a b c : ℝ) :
|
||||
|
|
|
@ -47,9 +47,12 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
empty_val', cons_val_fin_one, star_mul', RCLike.star_def, conj_ofReal, cons_val_one, head_cons,
|
||||
star_sub, star_neg, ← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons,
|
||||
head_fin_const]
|
||||
simp [conj_ofReal]
|
||||
simp only [ofReal_cos, ofReal_sin]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
fin_cases i <;>
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, one_apply_eq, Fin.mk_one, cons_val_one,
|
||||
head_cons, ne_eq, zero_ne_one, not_false_eq_true, one_apply_ne, Fin.reduceFinMk,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, Fin.reduceEq]
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq, sin_sq]
|
||||
|
@ -65,9 +68,12 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
· simp only [Fin.mk_one, Fin.isValue, conjTranspose_apply, cons_val', cons_val_one, head_cons,
|
||||
empty_val', cons_val_fin_one, cons_val_zero, star_mul', RCLike.star_def, conj_ofReal, star_sub,
|
||||
← exp_conj, _root_.map_mul, conj_I, neg_mul, cons_val_two, tail_cons, head_fin_const, star_neg]
|
||||
simp [conj_ofReal]
|
||||
simp only [ofReal_sin, ofReal_cos]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
fin_cases i <;>
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, Fin.mk_one, ne_eq, one_ne_zero,
|
||||
not_false_eq_true, one_apply_ne, cons_val_one, head_cons, one_apply_eq, Fin.reduceFinMk,
|
||||
cons_val_two, Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, Fin.reduceEq]
|
||||
· ring_nf
|
||||
field_simp
|
||||
rw [sin_sq, sin_sq]
|
||||
|
@ -86,7 +92,10 @@ lemma standParamAsMatrix_unitary (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) :
|
|||
head_fin_const]
|
||||
simp only [ofReal_sin, ofReal_cos]
|
||||
rw [exp_neg]
|
||||
fin_cases i <;> simp
|
||||
fin_cases i <;>
|
||||
simp only [Fin.zero_eta, Fin.isValue, cons_val_zero, Fin.reduceFinMk, ne_eq, Fin.reduceEq,
|
||||
not_false_eq_true, one_apply_ne, Fin.mk_one, cons_val_one, head_cons, cons_val_two,
|
||||
Nat.succ_eq_add_one, Nat.reduceAdd, tail_cons, one_apply_eq]
|
||||
· ring_nf
|
||||
rw [sin_sq]
|
||||
ring
|
||||
|
@ -148,7 +157,7 @@ lemma eq_rows (U : CKMMatrix) {θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ} (hu :
|
|||
|
||||
lemma eq_exp_of_phases (θ₁₂ θ₁₃ θ₂₃ δ₁₃ δ₁₃' : ℝ) (h : cexp (δ₁₃ * I) = cexp (δ₁₃' * I)) :
|
||||
standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃ = standParam θ₁₂ θ₁₃ θ₂₃ δ₁₃' := by
|
||||
simp [standParam, standParamAsMatrix]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, ofReal_sin, neg_mul]
|
||||
apply CKMMatrix_ext
|
||||
simp only
|
||||
rw [show exp (I * δ₁₃) = exp (I * δ₁₃') by rw [mul_comm, h, mul_comm]]
|
||||
|
@ -169,15 +178,16 @@ lemma VusVubVcdSq_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Rea
|
|||
neg_zero, Real.exp_zero, mul_one, _root_.sq_abs]
|
||||
rw [_root_.abs_of_nonneg h1, _root_.abs_of_nonneg h3, _root_.abs_of_nonneg h2,
|
||||
_root_.abs_of_nonneg h4]
|
||||
simp [sq]
|
||||
simp only [sq]
|
||||
ring_nf
|
||||
nth_rewrite 2 [Real.sin_sq θ₁₂]
|
||||
ring_nf
|
||||
field_simp
|
||||
ring
|
||||
· simp at hx
|
||||
· simp only [ne_eq, Decidable.not_not] at hx
|
||||
rw [hx]
|
||||
simp
|
||||
simp only [abs_zero, mul_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow,
|
||||
zero_mul, add_zero, div_zero]
|
||||
|
||||
open Invariant in
|
||||
lemma mulExpδ₁₃_eq (θ₁₂ θ₁₃ θ₂₃ δ₁₃ : ℝ) (h1 : 0 ≤ Real.sin θ₁₂)
|
||||
|
|
|
@ -397,10 +397,16 @@ lemma on_param_cos_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
have hS12 := congrArg ofReal (S₁₂_of_Vub_one (VubAbs_of_cos_θ₁₃_zero h))
|
||||
simp only [ofReal_eq_coe, ← S₁₂_eq_ℂsin_θ₁₂, map_zero] at hS12
|
||||
use 0, 0, 0, δ₁₃, 0, -δ₁₃
|
||||
simp [standParam, standParamAsMatrix, h, phaseShift, hS13, hC12, hS12]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, hC12, h, ofReal_zero, mul_zero, ofReal_sin,
|
||||
hS12, hS13, neg_mul, one_mul, neg_zero, zero_mul, mul_one, zero_sub, sub_zero, phaseShift,
|
||||
phaseShiftMatrix, exp_zero, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_zero, tail_val', head_val', zero_add, Fin.mk_one, Fin.reduceFinMk,
|
||||
neg_mul, mul_one]
|
||||
· rfl
|
||||
· rfl
|
||||
|
||||
|
@ -413,7 +419,11 @@ lemma on_param_cos_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
exp_zero, mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_zero, tail_val', head_val', zero_add, zero_eq_mul, mul_neg, neg_mul,
|
||||
mul_one, neg_inj, Fin.mk_one, Fin.reduceFinMk]
|
||||
· exact Or.inr rfl
|
||||
· change _ = _ + _ * 0
|
||||
field_simp
|
||||
|
@ -429,10 +439,16 @@ lemma on_param_cos_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.c
|
|||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, δ₁₃, 0, 0, 0, - δ₁₃
|
||||
have hb := exp_ne_zero (I * δ₁₃)
|
||||
simp [standParam, standParamAsMatrix, h, phaseShift, exp_neg]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, ofReal_sin, neg_mul, exp_neg, h,
|
||||
ofReal_zero, mul_zero, zero_sub, zero_mul, sub_zero, phaseShift, phaseShiftMatrix, exp_zero,
|
||||
mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', self_eq_add_right, mul_eq_zero,
|
||||
zero_add, mul_neg, neg_inj, Fin.mk_one, Fin.reduceFinMk]
|
||||
· exact Or.inr rfl
|
||||
· ring_nf
|
||||
change _ = _ + _ * 0
|
||||
|
@ -448,7 +464,11 @@ lemma on_param_sin_θ₁₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
Submonoid.mk_mul_mk, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', zero_add, self_eq_add_right,
|
||||
mul_eq_zero, Fin.mk_one, Fin.reduceFinMk]
|
||||
· exact Or.inr rfl
|
||||
· exact Or.inr rfl
|
||||
|
||||
|
@ -456,10 +476,16 @@ lemma on_param_sin_θ₁₂_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) δ₁₃ ≈ standParam (θ₁₂ ⟦V⟧) (θ₁₃ ⟦V⟧) (θ₂₃ ⟦V⟧) 0 := by
|
||||
use 0, δ₁₃, δ₁₃, 0, -δ₁₃, - δ₁₃
|
||||
have hb := exp_ne_zero (I * δ₁₃)
|
||||
simp [standParam, standParamAsMatrix, h, phaseShift, exp_neg]
|
||||
simp only [standParam, standParamAsMatrix, ofReal_cos, h, ofReal_zero, zero_mul, ofReal_sin,
|
||||
neg_mul, exp_neg, neg_zero, zero_sub, sub_zero, phaseShift, phaseShiftMatrix, mul_zero,
|
||||
exp_zero, mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', self_eq_add_right, mul_eq_zero,
|
||||
Fin.mk_one, zero_add, Fin.reduceFinMk, mul_neg, neg_mul, neg_inj]
|
||||
· apply Or.inr rfl
|
||||
· change _ = _ + _ * 0
|
||||
simp only [mul_zero, add_zero, neg_inj]
|
||||
|
@ -484,7 +510,11 @@ lemma on_param_sin_θ₂₃_eq_zero {V : CKMMatrix} (δ₁₃ : ℝ) (h : Real.s
|
|||
mul_one, Submonoid.mk_mul_mk, ofReal_neg, mul_neg, Subtype.mk.injEq]
|
||||
funext i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [mul_apply, Fin.sum_univ_three, mul_apply, Fin.sum_univ_three]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, Fin.zero_eta, Fin.isValue, cons_val',
|
||||
cons_val_zero, empty_val', cons_val_fin_one, mul_apply, Fin.sum_univ_three, one_mul,
|
||||
cons_val_one, head_cons, zero_mul, add_zero, cons_val_two, tail_cons, head_fin_const,
|
||||
vecCons_const, mul_one, mul_zero, tail_val', head_val', self_eq_add_right, mul_eq_zero,
|
||||
Fin.mk_one, zero_add, Fin.reduceFinMk, mul_neg, neg_inj]
|
||||
· exact Or.inr rfl
|
||||
· change _ = _ + _ * 0
|
||||
simp only [mul_zero, add_zero, neg_inj]
|
||||
|
|
|
@ -31,7 +31,10 @@ lemma toMatrix_isSelfAdjoint (x : LorentzVector 3) : IsSelfAdjoint (toMatrix x)
|
|||
rw [isSelfAdjoint_iff, star_eq_conjTranspose, ← Matrix.ext_iff]
|
||||
intro i j
|
||||
fin_cases i <;> fin_cases j <;>
|
||||
simp [toMatrix, conj_ofReal]
|
||||
simp only [toMatrix, LorentzVector.time, Fin.isValue, LorentzVector.space, Function.comp_apply,
|
||||
Fin.zero_eta, conjTranspose_apply, of_apply, cons_val', cons_val_zero, empty_val',
|
||||
cons_val_fin_one, star_add, RCLike.star_def, conj_ofReal, Fin.mk_one, cons_val_one,
|
||||
head_fin_const, star_mul', conj_I, mul_neg, head_cons, star_sub, sub_neg_eq_add]
|
||||
rfl
|
||||
|
||||
/-- A self-adjoint matrix formed from a Lorentz vector point. -/
|
||||
|
|
|
@ -62,7 +62,7 @@ def neg : Potential where
|
|||
|
||||
@[simp]
|
||||
lemma toFun_neg (φ : HiggsField) (x : SpaceTime) : P.neg.toFun φ x = - P.toFun φ x := by
|
||||
simp [neg, toFun]
|
||||
simp only [toFun, neg, neg_neg, normSq, neg_mul, neg_add_rev]
|
||||
ring
|
||||
|
||||
@[simp]
|
||||
|
@ -247,7 +247,8 @@ lemma neg_𝓵_sol_exists_iff (h𝓵 : P.𝓵 < 0) (c : ℝ) : (∃ φ x, P.toFu
|
|||
lemma pos_𝓵_sol_exists_iff (h𝓵 : 0 < P.𝓵) (c : ℝ) : (∃ φ x, P.toFun φ x = c) ↔ (P.μ2 < 0 ∧ 0 ≤ c) ∨
|
||||
(0 ≤ P.μ2 ∧ - P.μ2 ^ 2 / (4 * P.𝓵) ≤ c) := by
|
||||
have h1 := P.neg.neg_𝓵_sol_exists_iff (by simpa using h𝓵) (- c)
|
||||
simp at h1
|
||||
simp only [toFun_neg, neg_inj, μ2_neg, Left.neg_pos_iff, Left.neg_nonpos_iff, even_two,
|
||||
Even.neg_pow, 𝓵_neg, mul_neg, neg_div_neg_eq] at h1
|
||||
rw [neg_le, neg_div'] at h1
|
||||
exact h1
|
||||
|
||||
|
@ -310,7 +311,7 @@ lemma isBounded_of_𝓵_pos (h : 0 < P.𝓵) : P.IsBounded := by
|
|||
simp only [IsBounded]
|
||||
have h2 := P.pos_𝓵_quadDiscrim_zero_bound h
|
||||
by_contra hn
|
||||
simp at hn
|
||||
simp only [not_exists, not_forall, not_le] at hn
|
||||
obtain ⟨φ, x, hx⟩ := hn (-P.μ2 ^ 2 / (4 * P.𝓵))
|
||||
have h2' := h2 φ x
|
||||
linarith
|
||||
|
@ -342,7 +343,6 @@ lemma isMinOn_iff_of_μSq_nonpos_𝓵_pos (h𝓵 : 0 < P.𝓵) (hμ2 : P.μ2 ≤
|
|||
(x : SpaceTime) : IsMinOn (fun (φ, x) => P.toFun φ x) Set.univ (φ, x)
|
||||
↔ P.toFun φ x = 0 := by
|
||||
have h1 := P.pos_𝓵_sol_exists_iff h𝓵
|
||||
simp [hμ2] at h1
|
||||
rw [isMinOn_univ_iff]
|
||||
simp only [Prod.forall]
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
|
@ -375,7 +375,7 @@ lemma isMinOn_iff_of_μSq_nonneg_𝓵_pos (h𝓵 : 0 < P.𝓵) (hμ2 : 0 ≤ P.
|
|||
(x : SpaceTime) : IsMinOn (fun (φ, x) => P.toFun φ x) Set.univ (φ, x) ↔
|
||||
P.toFun φ x = - P.μ2 ^ 2 / (4 * P.𝓵) := by
|
||||
have h1 := P.pos_𝓵_sol_exists_iff h𝓵
|
||||
simp [hμ2, not_lt.mpr hμ2] at h1
|
||||
simp only [not_lt.mpr hμ2, false_and, hμ2, true_and, false_or] at h1
|
||||
rw [isMinOn_univ_iff]
|
||||
simp only [Prod.forall]
|
||||
refine Iff.intro (fun h => ?_) (fun h => ?_)
|
||||
|
|
|
@ -152,7 +152,7 @@ lemma map_tprod {X Y : OverColor Color} (p : (i : X.left) → (colorToRep (X.hom
|
|||
PiTensorProduct.tprod ℂ fun (i : Y.left) => colorToRepCongr
|
||||
(OverColor.Hom.toEquiv_comp_inv_apply f i) (p ((OverColor.Hom.toEquiv f).symm i)) := by
|
||||
change (colorFun.map' f).hom ((PiTensorProduct.tprod ℂ) p) = _
|
||||
simp [colorFun.map', mapToLinearEquiv']
|
||||
simp only [map', mapToLinearEquiv', Functor.id_obj]
|
||||
erw [LinearEquiv.trans_apply]
|
||||
change (PiTensorProduct.congr fun i => colorToRepCongr _)
|
||||
((PiTensorProduct.reindex ℂ (fun x => _) (OverColor.Hom.toEquiv f))
|
||||
|
@ -276,7 +276,9 @@ def μ (X Y : OverColor Color) : colorFun.obj X ⊗ colorFun.obj Y ≅ colorFun.
|
|||
inv := {
|
||||
hom := (μModEquiv X Y).symm.toLinearMap
|
||||
comm := fun M => by
|
||||
simp [CategoryStruct.comp]
|
||||
simp only [Action.instMonoidalCategory_tensorObj_V, CategoryStruct.comp,
|
||||
colorFun_obj_V_carrier, OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.tensor_ρ']
|
||||
erw [LinearEquiv.eq_comp_toLinearMap_symm,LinearMap.comp_assoc,
|
||||
LinearEquiv.toLinearMap_symm_comp_eq]
|
||||
refine HepLean.PiTensorProduct.induction_tmul (fun p q => ?_)
|
||||
|
@ -428,7 +430,14 @@ lemma left_unitality (X : OverColor Color) : (leftUnitor (colorFun.obj X)).hom =
|
|||
change TensorProduct.lid ℂ (colorFun.obj X) (x ⊗ₜ[ℂ] (PiTensorProduct.tprod ℂ) q) =
|
||||
(colorFun.map (λ_ X).hom).hom ((μ (𝟙_ (OverColor Color)) X).hom.hom
|
||||
((((PiTensorProduct.isEmptyEquiv Empty).symm x) ⊗ₜ[ℂ] (PiTensorProduct.tprod ℂ) q)))
|
||||
simp [PiTensorProduct.isEmptyEquiv]
|
||||
simp only [Functor.id_obj, lid_tmul, colorFun_obj_V_carrier,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_tensorObj_V,
|
||||
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, PiTensorProduct.isEmptyEquiv,
|
||||
LinearEquiv.coe_symm_mk]
|
||||
rw [TensorProduct.smul_tmul, TensorProduct.tmul_smul]
|
||||
erw [LinearMap.map_smul, LinearMap.map_smul]
|
||||
apply congrArg
|
||||
|
@ -454,7 +463,14 @@ lemma right_unitality (X : OverColor Color) : (MonoidalCategory.rightUnitor (col
|
|||
change TensorProduct.rid ℂ (colorFun.obj X) ((PiTensorProduct.tprod ℂ) p ⊗ₜ[ℂ] x) =
|
||||
(colorFun.map (ρ_ X).hom).hom ((μ X (𝟙_ (OverColor Color))).hom.hom
|
||||
((((PiTensorProduct.tprod ℂ) p ⊗ₜ[ℂ] ((PiTensorProduct.isEmptyEquiv Empty).symm x)))))
|
||||
simp [PiTensorProduct.isEmptyEquiv]
|
||||
simp only [Functor.id_obj, rid_tmul, colorFun_obj_V_carrier,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorUnit_left,
|
||||
OverColor.instMonoidalCategoryStruct_tensorObj_hom, Action.instMonoidalCategory_tensorObj_V,
|
||||
Equivalence.symm_inverse, Action.functorCategoryEquivalence_functor,
|
||||
Action.FunctorCategoryEquivalence.functor_obj_obj,
|
||||
OverColor.instMonoidalCategoryStruct_tensorUnit_hom, PiTensorProduct.isEmptyEquiv,
|
||||
LinearEquiv.coe_symm_mk, tmul_smul]
|
||||
erw [LinearMap.map_smul, LinearMap.map_smul]
|
||||
apply congrArg
|
||||
change _ = (colorFun.map (ρ_ X).hom).hom ((μ X (𝟙_ (OverColor Color))).hom.hom
|
||||
|
|
|
@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
|
|||
Authors: Joseph Tooby-Smith
|
||||
-/
|
||||
import HepLean.Tensors.OverColor.Basic
|
||||
import HepLean.Tensors.OverColor.Functors
|
||||
import HepLean.Tensors.ComplexLorentz.ColorFun
|
||||
import HepLean.Mathematics.PiTensorProduct
|
||||
/-!
|
||||
|
@ -59,7 +58,7 @@ lemma mapToLinearEquiv'_tprod {f g : OverColor Color} (m : f ⟶ g)
|
|||
PiTensorProduct.tprod ℂ fun i =>
|
||||
(TensorProduct.congr (colorToRepCongr (OverColor.Hom.toEquiv_symm_apply m i))
|
||||
(colorToRepCongr (mapToLinearEquiv'.proof_4 m i))) (x ((OverColor.Hom.toEquiv m).symm i)) := by
|
||||
simp [mapToLinearEquiv']
|
||||
simp only [mapToLinearEquiv', Functor.id_obj, LinearEquiv.trans_apply]
|
||||
change (PiTensorProduct.congr fun i => TensorProduct.congr (colorToRepCongr _)
|
||||
(colorToRepCongr _)) ((PiTensorProduct.reindex ℂ
|
||||
(fun x => ↑(colorToRep (f.hom x)).V ⊗[ℂ] ↑(colorToRep (τ (f.hom x))).V)
|
||||
|
|
|
@ -5,7 +5,6 @@ Authors: Joseph Tooby-Smith
|
|||
-/
|
||||
import HepLean.Tensors.Tree.Basic
|
||||
import HepLean.Tensors.ComplexLorentz.TensorStruct
|
||||
import HepLean.Tensors.Tree.Elab
|
||||
/-!
|
||||
|
||||
## The tensor structure for complex Lorentz tensors
|
||||
|
@ -33,6 +32,8 @@ def upDown : Fin 2 → complexLorentzTensor.C
|
|||
variable (T S : complexLorentzTensor.F.obj (OverColor.mk upDown))
|
||||
|
||||
/-
|
||||
import HepLean.Tensors.Tree.Elab
|
||||
|
||||
#check {T | i m ⊗ S | m l}ᵀ.dot
|
||||
#check {T | i m ⊗ S | l τ(m)}ᵀ.dot
|
||||
-/
|
||||
|
|
|
@ -54,13 +54,13 @@ def toEquiv (m : f ⟶ g) : f.left ≃ g.left where
|
|||
@[simp]
|
||||
lemma toEquiv_id (f : OverColor C) : toEquiv (𝟙 f) = Equiv.refl f.left := by
|
||||
ext x
|
||||
simp [toEquiv]
|
||||
simp only [toEquiv, Equiv.coe_fn_mk, Equiv.refl_apply]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma toEquiv_comp (m : f ⟶ g) (n : g ⟶ h) : toEquiv (m ≫ n) = (toEquiv m).trans (toEquiv n) := by
|
||||
ext x
|
||||
simp [toEquiv]
|
||||
simp only [toEquiv, Equiv.coe_fn_mk, Equiv.trans_apply]
|
||||
rfl
|
||||
|
||||
lemma toEquiv_symm_apply (m : f ⟶ g) (i : g.left) :
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue