refactor: Last batch of multi-goal proofs

This commit is contained in:
jstoobysmith 2024-08-21 06:40:58 -04:00
parent b9479c904d
commit c0499483a8
43 changed files with 910 additions and 955 deletions

View file

@ -77,9 +77,8 @@ lemma accCube_explicit (n : ) (S : (PureU1Charges n).Charges) :
rw [accCubeTriLinSymm]
simp only [PureU1Charges_numberCharges, TriLinearSymm.mk₃_toFun_apply_apply]
apply Finset.sum_congr
simp only
ring_nf
simp
· rfl
· exact fun x _ => Eq.symm (pow_three' (S x))
end PureU1
@ -130,42 +129,40 @@ lemma pureU1_anomalyFree_ext {n : } {S T : (PureU1 n.succ).LinSols}
apply ACCSystemLinear.LinSols.ext
funext i
by_cases hi : i ≠ Fin.last n
have hiCast : ∃ j : Fin n, j.castSucc = i := by
exact Fin.exists_castSucc_eq.mpr hi
obtain ⟨j, hj⟩ := hiCast
rw [← hj]
exact h j
simp at hi
rw [hi]
rw [pureU1_last, pureU1_last]
simp only [neg_inj]
apply Finset.sum_congr
simp only
intro j _
exact h j
· have hiCast : ∃ j : Fin n, j.castSucc = i := by
exact Fin.exists_castSucc_eq.mpr hi
obtain ⟨j, hj⟩ := hiCast
rw [← hj]
exact h j
· simp at hi
rw [hi, pureU1_last, pureU1_last]
simp only [neg_inj]
apply Finset.sum_congr
· simp only
· exact fun j _ => h j
namespace PureU1
lemma sum_of_charges {n : } (f : Fin k → (PureU1 n).Charges) (j : Fin n) :
(∑ i : Fin k, (f i)) j = ∑ i : Fin k, (f i) j := by
induction k
simp
rfl
rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
have hlt := hl (f ∘ Fin.castSucc)
erw [← hlt]
simp
· simp only [univ_eq_empty, sum_empty]
rfl
· rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
have hlt := hl (f ∘ Fin.castSucc)
erw [← hlt]
simp
lemma sum_of_anomaly_free_linear {n : } (f : Fin k → (PureU1 n).LinSols) (j : Fin n) :
(∑ i : Fin k, (f i)).1 j = (∑ i : Fin k, (f i).1 j) := by
induction k
simp
rfl
rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
have hlt := hl (f ∘ Fin.castSucc)
erw [← hlt]
simp
· simp only [univ_eq_empty, sum_empty, ACCSystemLinear.linSolsAddCommMonoid_zero_val]
rfl
· rename_i k hl
rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc]
have hlt := hl (f ∘ Fin.castSucc)
erw [← hlt]
simp
end PureU1

View file

@ -38,16 +38,16 @@ lemma asCharges_ne_castSucc {k j : Fin n} (h : k ≠ j) :
asCharges k (Fin.castSucc j) = 0 := by
simp [asCharges]
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
have hj : j.val < n := by
exact j.prop
simp_all
rfl
· 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
have hj : j.val < n := by
exact j.prop
simp_all
· rfl
/-- The basis elements as `LinSols`. -/
@[simps!]
@ -61,16 +61,16 @@ def asLinSols (j : Fin n) : (PureU1 n.succ).LinSols :=
LinearMap.coe_mk, AddHom.coe_mk, Fin.coe_eq_castSucc]
rw [Fin.sum_univ_castSucc]
rw [Finset.sum_eq_single j]
simp only [asCharges, PureU1_numberCharges, ↓reduceIte]
have hn : ¬ (Fin.last n = Fin.castSucc j) := Fin.ne_of_gt j.prop
split
rename_i ht
exact (hn ht).elim
rfl
intro k _ hkj
exact asCharges_ne_castSucc hkj.symm
intro hk
simp at hk⟩
· simp only [asCharges, PureU1_numberCharges, ↓reduceIte]
have hn : ¬ (Fin.last n = Fin.castSucc j) := Fin.ne_of_gt j.prop
split
· rename_i ht
exact (hn ht).elim
· rfl
· intro k _ hkj
exact asCharges_ne_castSucc hkj.symm
· intro hk
simp at hk⟩
lemma sum_of_vectors {n : } (f : Fin k → (PureU1 n).LinSols) (j : Fin n) :
(∑ i : Fin k, (f i)).1 j = (∑ i : Fin k, (f i).1 j) :=
@ -94,11 +94,11 @@ def coordinateMap : (PureU1 n.succ).LinSols ≃ₗ[] Fin n →₀ where
asLinSols_val, Equiv.toFun_as_coe,
Fin.coe_eq_castSucc, mul_ite, mul_one, mul_neg, mul_zero, Equiv.invFun_as_coe]
rw [Finset.sum_eq_single j]
simp only [asCharges, PureU1_numberCharges, ↓reduceIte, mul_one]
intro k _ hkj
rw [asCharges_ne_castSucc hkj]
exact Rat.mul_zero (S.val k.castSucc)
simp
· simp only [asCharges, PureU1_numberCharges, ↓reduceIte, mul_one]
· intro k _ hkj
rw [asCharges_ne_castSucc hkj]
exact Rat.mul_zero (S.val k.castSucc)
· simp
right_inv f := by
simp only [PureU1_numberCharges, Equiv.invFun_as_coe]
ext
@ -109,11 +109,11 @@ def coordinateMap : (PureU1 n.succ).LinSols ≃ₗ[] Fin n →₀ where
asLinSols_val, Equiv.toFun_as_coe,
Fin.coe_eq_castSucc, mul_ite, mul_one, mul_neg, mul_zero, Equiv.invFun_as_coe]
rw [Finset.sum_eq_single j]
simp only [asCharges, PureU1_numberCharges, ↓reduceIte, mul_one]
intro k _ hkj
rw [asCharges_ne_castSucc hkj]
exact Rat.mul_zero (f k)
simp
· simp only [asCharges, PureU1_numberCharges, ↓reduceIte, mul_one]
· intro k _ hkj
rw [asCharges_ne_castSucc hkj]
exact Rat.mul_zero (f k)
· simp
/-- The basis of `LinSols`. -/
noncomputable

View file

@ -32,13 +32,11 @@ lemma constAbs_perm (S : (PureU1 n).Charges) (M :(FamilyPermutations n).group) :
ConstAbs ((FamilyPermutations n).rep M S) ↔ ConstAbs S := by
simp only [ConstAbs, PureU1_numberCharges, FamilyPermutations, PermGroup, permCharges,
MonoidHom.coe_mk, OneHom.coe_mk, chargeMap_apply]
apply Iff.intro
intro h i j
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
exact h2
intro h i j
exact h (M.invFun i) (M.invFun j)
lemma constAbs_sort {S : (PureU1 n).Charges} (CA : ConstAbs S) : ConstAbs (sort S) := by
rw [sort]
@ -58,8 +56,8 @@ lemma lt_eq {k i : Fin n.succ} (hk : S k ≤ 0) (hik : i ≤ k) : S i = S k := b
have ht := hS.1 i k
rw [sq_eq_sq_iff_eq_or_eq_neg] at ht
cases ht <;> rename_i h
exact h
linarith
· exact h
· linarith
lemma val_le_zero {i : Fin n.succ} (hi : S i ≤ 0) : S i = S (0 : Fin n.succ) := by
symm
@ -71,8 +69,8 @@ lemma gt_eq {k i: Fin n.succ} (hk : 0 ≤ S k) (hik : k ≤ i) : S i = S k := by
have ht := hS.1 i k
rw [sq_eq_sq_iff_eq_or_eq_neg] at ht
cases ht <;> rename_i h
exact h
linarith
· exact h
· linarith
lemma zero_gt (h0 : 0 ≤ S (0 : Fin n.succ)) (i : Fin n.succ) : S (0 : Fin n.succ) = S i := by
symm
@ -83,9 +81,9 @@ lemma opposite_signs_eq_neg {i j : Fin n.succ} (hi : S i ≤ 0) (hj : 0 ≤ S j)
have hSS := hS.1 i j
rw [sq_eq_sq_iff_eq_or_eq_neg] at hSS
cases' hSS with h h
simp_all
linarith
exact h
· simp_all
linarith
· exact h
lemma is_zero (h0 : S (0 : Fin n.succ) = 0) : S = 0 := by
funext i
@ -115,10 +113,10 @@ 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]
erw [Finset.sum_equiv (Fin.castOrderIso (boundary_split k)).toEquiv]
intro i
simp only [Fin.val_succ, mem_univ, RelIso.coe_fn_toEquiv]
intro i
simp
· intro i
simp only [Fin.val_succ, mem_univ, RelIso.coe_fn_toEquiv]
· intro i
simp
lemma boundary_accGrav'' (k : Fin n) (hk : Boundary S k) :
accGrav n.succ S = (2 * ↑↑k + 1 - ↑n) * S (0 : Fin n.succ) := by
@ -146,19 +144,19 @@ lemma not_hasBoundary_zero_le (hnot : ¬ (HasBoundary S)) (h0 : S (0 : Fin n.suc
intro ⟨i, hi⟩
simp at hnot
induction i
rfl
rename_i i hii
have hnott := hnot ⟨i, by {simp at hi; linarith}⟩
have hii := hii (by omega)
erw [← hii] at hnott
exact (val_le_zero hS (hnott h0)).symm
· rfl
· rename_i i hii
have hnott := hnot ⟨i, by {simp at hi; linarith}⟩
have hii := hii (by omega)
erw [← hii] at hnott
exact (val_le_zero hS (hnott h0)).symm
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
exact zero_gt hS hi i
· exact not_hasBoundary_zero_le hS hnot hi i
· simp at hi
exact zero_gt hS hi i
lemma not_hasBoundary_grav (hnot : ¬ (HasBoundary S)) :
accGrav n.succ S = n.succ * S (0 : Fin n.succ) := by
@ -171,8 +169,8 @@ lemma AFL_hasBoundary (h : A.val (0 : Fin n.succ) ≠ 0) : HasBoundary A.val :=
erw [pureU1_linear A] at h0
simp at h0
cases' h0
linarith
simp_all
· linarith
· simp_all
lemma AFL_odd_noBoundary {A : (PureU1 (2 * n + 1)).LinSols} (h : ConstAbsSorted A.val)
(hA : A.val (0 : Fin (2*n +1)) ≠ 0) : ¬ HasBoundary A.val := by
@ -224,9 +222,9 @@ lemma AFL_even_below (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.v
A.val (Fin.cast (split_equal n.succ) (Fin.castAdd n.succ i))
= A.val (0 : Fin (2*n.succ)) := by
by_cases hA : A.val (0 : Fin (2*n.succ)) = 0
rw [is_zero h hA]
rfl
exact AFL_even_below' h hA i
· rw [is_zero h hA]
rfl
· exact AFL_even_below' h hA i
lemma AFL_even_above' {A : (PureU1 (2 * n.succ)).LinSols} (h : ConstAbsSorted A.val)
(hA : A.val (0 : Fin (2*n.succ)) ≠ 0) (i : Fin n.succ) :
@ -245,9 +243,9 @@ lemma AFL_even_above (A : (PureU1 (2 * n.succ)).LinSols) (h : ConstAbsSorted A.v
A.val (Fin.cast (split_equal n.succ) (Fin.natAdd n.succ i)) =
- A.val (0 : Fin (2 * n.succ)) := by
by_cases hA : A.val (0 : Fin (2 * n.succ)) = 0
rw [is_zero h hA]
rfl
exact AFL_even_above' h hA i
· rw [is_zero h hA]
rfl
· exact AFL_even_above' h hA i
end charges

View file

@ -54,19 +54,19 @@ lemma ext_δ (S T : Fin (2 * n.succ) → ) (h1 : ∀ i, S (δ₁ i) = T (δ
(h2 : ∀ i, S (δ₂ i) = T (δ₂ i)) : S = T := by
funext i
by_cases hi : i.val < n.succ
let j : Fin n.succ := ⟨i, hi⟩
have h2 := h1 j
have h3 : δ₁ j = i := by
simp [δ₁, Fin.ext_iff]
rw [h3] at h2
exact h2
let j : Fin n.succ := ⟨i - n.succ, by omega⟩
have h2 := h2 j
have h3 : δ₂ j = i := by
simp [δ₂, Fin.ext_iff]
omega
rw [h3] at h2
exact h2
· let j : Fin n.succ := ⟨i, hi⟩
have h2 := h1 j
have h3 : δ₁ j = i := by
simp [δ₁, Fin.ext_iff]
rw [h3] at h2
exact h2
· let j : Fin n.succ := ⟨i - n.succ, by omega⟩
have h2 := h2 j
have h3 : δ₂ j = i := by
simp [δ₂, Fin.ext_iff]
omega
rw [h3] at h2
exact h2
lemma sum_δ₁_δ₂ (S : Fin (2 * n.succ) → ) :
∑ i, S i = ∑ i : Fin n.succ, ((S ∘ δ₁) i + (S ∘ δ₂) i) := by
@ -167,18 +167,18 @@ lemma basis_on_δ₁_other {k j : Fin n.succ} (h : k ≠ j) :
simp [basisAsCharges]
simp [δ₁, δ₂]
split
rename_i h1
rw [Fin.ext_iff] at h1
simp_all
rw [Fin.ext_iff] at h
simp_all
split
rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp at h2
omega
rfl
· rename_i h1
rw [Fin.ext_iff] at h1
simp_all
rw [Fin.ext_iff] at h
simp_all
· split
· rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp at h2
omega
· rfl
lemma basis_on_other {k : Fin n.succ} {j : Fin (2 * n.succ)} (h1 : j ≠ δ₁ k) (h2 : j ≠ δ₂ k) :
basisAsCharges k j = 0 := by
@ -195,18 +195,18 @@ lemma basis!_on_δ!₁_other {k j : Fin n} (h : k ≠ j) :
simp [basis!AsCharges]
simp [δ!₁, δ!₂]
split
rename_i h1
rw [Fin.ext_iff] at h1
simp_all
rw [Fin.ext_iff] at h
simp_all
split
rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp at h2
omega
rfl
· rename_i h1
rw [Fin.ext_iff] at h1
simp_all
rw [Fin.ext_iff] at h
simp_all
· split
· rename_i h1 h2
simp_all
rw [Fin.ext_iff] at h2
simp at h2
omega
· rfl
lemma basis_δ₂_eq_minus_δ₁ (j i : Fin n.succ) :
basisAsCharges j (δ₂ i) = - basisAsCharges j (δ₁ i) := by
@ -233,8 +233,8 @@ lemma basis!_δ!₂_eq_minus_δ!₁ (j i : Fin n) :
all_goals rename_i h1 h2
all_goals rw [Fin.ext_iff] at h1 h2
all_goals simp_all
subst h1
exact Fin.elim0 i
· subst h1
exact Fin.elim0 i
all_goals rename_i h3
all_goals rw [Fin.ext_iff] at h3
all_goals simp_all
@ -256,27 +256,27 @@ lemma basis!_on_δ!₂_other {k j : Fin n} (h : k ≠ j) : basis!AsCharges k (δ
lemma basis!_on_δ!₃ (j : Fin n) : basis!AsCharges j δ!₃ = 0 := by
simp [basis!AsCharges]
split <;> rename_i h
rw [Fin.ext_iff] at h
simp [δ!₃, δ!₁] at h
omega
split <;> rename_i h2
rw [Fin.ext_iff] at h2
simp [δ!₃, δ!₂] at h2
omega
rfl
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
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
omega
· rfl
lemma basis!_on_δ!₄ (j : Fin n) : basis!AsCharges j δ!₄ = 0 := by
simp [basis!AsCharges]
split <;> rename_i h
rw [Fin.ext_iff] at h
simp [δ!₄, δ!₁] at h
omega
split <;> rename_i h2
rw [Fin.ext_iff] at h2
simp [δ!₄, δ!₂] at h2
omega
rfl
· rw [Fin.ext_iff] at h
simp [δ!₄, δ!₁] at h
omega
· split <;> rename_i h2
· rw [Fin.ext_iff] at h2
simp [δ!₄, δ!₂] at h2
omega
· rfl
lemma basis_linearACC (j : Fin n.succ) : (accGrav (2 * n.succ)) (basisAsCharges j) = 0 := by
rw [accGrav]
@ -343,16 +343,15 @@ lemma swap!_as_add {S S' : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
funext i
rw [← hS, FamilyPermutations_anomalyFreeLinear_apply]
by_cases hi : i = δ!₁ j
subst hi
simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst]
by_cases hi2 : i = δ!₂ j
subst hi2
simp [HSMul.hSMul, basis!_on_δ!₂_self, pairSwap_inv_snd]
simp [HSMul.hSMul]
rw [basis!_on_other hi hi2]
change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_
erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)]
simp
· subst hi
simp [HSMul.hSMul, basis!_on_δ!₁_self, pairSwap_inv_fst]
· by_cases hi2 : i = δ!₂ j
· simp [HSMul.hSMul, hi2, basis!_on_δ!₂_self, pairSwap_inv_snd]
· simp [HSMul.hSMul]
rw [basis!_on_other hi hi2]
change S.val ((pairSwap (δ!₁ j) (δ!₂ j)).invFun i) =_
erw [pairSwap_inv_other (Ne.symm hi) (Ne.symm hi2)]
simp
/-- A point in the span of the first part of the basis as a charge. -/
def P (f : Fin n.succ → ) : (PureU1 (2 * n.succ)).Charges := ∑ i, f i • basisAsCharges i
@ -367,23 +366,23 @@ lemma P_δ₁ (f : Fin n.succ → ) (j : Fin n.succ) : P f (δ₁ j) = f j :=
rw [P, sum_of_charges]
simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j]
rw [basis_on_δ₁_self]
simp only [mul_one]
intro k _ hkj
rw [basis_on_δ₁_other hkj]
simp only [mul_zero]
simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
· rw [basis_on_δ₁_self]
simp only [mul_one]
· intro k _ hkj
rw [basis_on_δ₁_other hkj]
simp only [mul_zero]
· simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
lemma P!_δ!₁ (f : Fin n → ) (j : Fin n) : P! f (δ!₁ j) = f j := by
rw [P!, sum_of_charges]
simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j]
rw [basis!_on_δ!₁_self]
simp only [mul_one]
intro k _ hkj
rw [basis!_on_δ!₁_other hkj]
simp only [mul_zero]
simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
· rw [basis!_on_δ!₁_self]
simp only [mul_one]
· intro k _ hkj
rw [basis!_on_δ!₁_other hkj]
simp only [mul_zero]
· simp only [mem_univ, not_true_eq_false, _root_.mul_eq_zero, IsEmpty.forall_iff]
lemma Pa_δ!₁ (f : Fin n.succ → ) (g : Fin n → ) (j : Fin n) :
Pa f g (δ!₁ j) = f j.succ + g j := by
@ -395,23 +394,21 @@ lemma P_δ₂ (f : Fin n.succ → ) (j : Fin n.succ) : P f (δ₂ j) = - f j
rw [P, sum_of_charges]
simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j]
rw [basis_on_δ₂_self]
simp only [mul_neg, mul_one]
intro k _ hkj
rw [basis_on_δ₂_other hkj]
simp only [mul_zero]
simp
· simp only [basis_on_δ₂_self, mul_neg, mul_one]
· intro k _ hkj
simp only [basis_on_δ₂_other hkj, mul_zero]
· simp
lemma P!_δ!₂ (f : Fin n → ) (j : Fin n) : P! f (δ!₂ j) = - f j := by
rw [P!, sum_of_charges]
simp [HSMul.hSMul, SMul.smul]
rw [Finset.sum_eq_single j]
rw [basis!_on_δ!₂_self]
simp only [mul_neg, mul_one]
intro k _ hkj
rw [basis!_on_δ!₂_other hkj]
simp only [mul_zero]
simp
· rw [basis!_on_δ!₂_self]
simp only [mul_neg, mul_one]
· intro k _ hkj
rw [basis!_on_δ!₂_other hkj]
simp only [mul_zero]
· simp
lemma Pa_δ!₂ (f : Fin n.succ → ) (g : Fin n → ) (j : Fin n) :
Pa f g (δ!₂ j) = - f j.castSucc - g j := by
@ -474,13 +471,13 @@ lemma P_P_P!_accCube (g : Fin n.succ → ) (j : Fin n) :
rw [sum_δ!₁_δ!₂, basis!_on_δ!₃, basis!_on_δ!₄]
simp only [mul_zero, add_zero, Function.comp_apply, zero_add]
rw [Finset.sum_eq_single j, basis!_on_δ!₁_self, basis!_on_δ!₂_self]
simp [δ!₁_δ₁, δ!₂_δ₂]
rw [P_δ₁, P_δ₂]
ring
intro k _ hkj
erw [basis!_on_δ!₁_other hkj.symm, basis!_on_δ!₂_other hkj.symm]
simp only [mul_zero, add_zero]
simp
· simp [δ!₁_δ₁, δ!₂_δ₂]
rw [P_δ₁, P_δ₂]
ring
· intro k _ hkj
erw [basis!_on_δ!₁_other hkj.symm, basis!_on_δ!₂_other hkj.symm]
simp only [mul_zero, add_zero]
· simp
lemma P_P!_P!_accCube (g : Fin n → ) (j : Fin n.succ) :
accCubeTriLinSymm (P! g) (P! g) (basisAsCharges j)
@ -489,12 +486,12 @@ lemma P_P!_P!_accCube (g : Fin n → ) (j : Fin n.succ) :
rw [sum_δ₁_δ₂]
simp only [Function.comp_apply]
rw [Finset.sum_eq_single j, basis_on_δ₁_self, basis_on_δ₂_self]
simp [δ!₁_δ₁, δ!₂_δ₂]
ring
intro k _ hkj
erw [basis_on_δ₁_other hkj.symm, basis_on_δ₂_other hkj.symm]
simp only [mul_zero, add_zero]
simp
· simp [δ!₁_δ₁, δ!₂_δ₂]
ring
· intro k _ hkj
erw [basis_on_δ₁_other hkj.symm, basis_on_δ₂_other hkj.symm]
simp only [mul_zero, add_zero]
· simp
lemma P_zero (f : Fin n.succ → ) (h : P f = 0) : ∀ i, f i = 0 := by
intro i
@ -601,27 +598,24 @@ theorem basisa_linear_independent : LinearIndependent (@basisa n) := by
intro i
simp_all
cases i
simp_all
simp_all
· simp_all
· simp_all
lemma Pa'_eq (f f' : (Fin n.succ) ⊕ (Fin n) → ) : Pa' f = Pa' f' ↔ f = f' := by
apply Iff.intro
intro h
funext i
rw [Pa', Pa'] at h
have h1 : ∑ i : Fin (succ n) ⊕ Fin n, (f i + (- f' i)) • basisa i = 0 := by
simp only [add_smul, neg_smul]
rw [Finset.sum_add_distrib]
rw [h]
rw [← Finset.sum_add_distrib]
simp
have h2 : ∀ i, (f i + (- f' i)) = 0 := by
exact Fintype.linearIndependent_iff.mp (@basisa_linear_independent n)
(fun i => f i + -f' i) h1
have h2i := h2 i
linarith
intro h
rw [h]
refine Iff.intro (fun h => (funext (fun i => ?_))) (fun h => ?_)
· rw [Pa', Pa'] at h
have h1 : ∑ i : Fin (succ n) ⊕ Fin n, (f i + (- f' i)) • basisa i = 0 := by
simp only [add_smul, neg_smul]
rw [Finset.sum_add_distrib]
rw [h]
rw [← Finset.sum_add_distrib]
simp
have h2 : ∀ i, (f i + (- f' i)) = 0 := by
exact Fintype.linearIndependent_iff.mp (@basisa_linear_independent n)
(fun i => f i + -f' i) h1
have h2i := h2 i
linarith
· rw [h]
/-! TODO: Replace the definition of `join` with a Mathlib definition, most likely `Sum.elim`. -/
/-- A helper function for what follows. -/
@ -632,33 +626,24 @@ def join (g : Fin n.succ → ) (f : Fin n → ) : (Fin n.succ) ⊕ (Fin n)
lemma join_ext (g g' : Fin n.succ → ) (f f' : Fin n → ) :
join g f = join g' f' ↔ g = g' ∧ f = f' := by
apply Iff.intro
intro h
apply And.intro
funext i
exact congr_fun h (Sum.inl i)
funext i
exact congr_fun h (Sum.inr i)
intro h
rw [h.left, h.right]
refine Iff.intro (fun h => ?_) (fun h => ?_)
· apply And.intro (funext (fun i => congr_fun h (Sum.inl i)))
(funext (fun i => congr_fun h (Sum.inr i)))
· rw [h.left, h.right]
lemma join_Pa (g g' : Fin n.succ → ) (f f' : Fin n → ) :
Pa' (join g f) = Pa' (join g' f') ↔ Pa g f = Pa g' f' := by
apply Iff.intro
intro h
rw [Pa'_eq] at h
rw [join_ext] at h
rw [h.left, h.right]
intro h
apply ACCSystemLinear.LinSols.ext
rw [Pa'_P'_P!', Pa'_P'_P!']
simp [P'_val, P!'_val]
exact h
refine Iff.intro (fun h => ?_) (fun h => ?_)
· rw [Pa'_eq, join_ext] at h
rw [h.left, h.right]
· apply ACCSystemLinear.LinSols.ext
rw [Pa'_P'_P!', Pa'_P'_P!']
simp only [succ_eq_add_one, ACCSystemLinear.linSolsAddCommMonoid_add_val, P'_val, P!'_val]
exact h
lemma Pa_eq (g g' : Fin n.succ → ) (f f' : Fin n → ) :
Pa g f = Pa g' f' ↔ g = g' ∧ f = f' := by
rw [← join_Pa]
rw [← join_ext]
rw [← join_Pa, ← join_ext]
exact Pa'_eq _ _
lemma basisa_card : Fintype.card ((Fin n.succ) ⊕ (Fin n)) =
@ -694,8 +679,8 @@ lemma smul_basis!AsCharges_in_span (S : (PureU1 (2 * n.succ)).LinSols) (j : Fin
Submodule.span (Set.range basis!AsCharges) := by
apply Submodule.smul_mem
apply SetLike.mem_of_subset
apply Submodule.subset_span
simp_all only [Set.mem_range, exists_apply_eq_apply]
· exact Submodule.subset_span
· simp_all only [Set.mem_range, exists_apply_eq_apply]
lemma span_basis_swap! {S : (PureU1 (2 * n.succ)).LinSols} (j : Fin n)
(hS : ((FamilyPermutations (2 * n.succ)).linSolRep
@ -732,19 +717,19 @@ lemma vectorLikeEven_in_span (S : (PureU1 (2 * n.succ)).LinSols)
rw [sortAFL_val]
erw [P'_val]
apply ext_δ
intro i
rw [P_δ₁]
rfl
intro i
rw [P_δ₂]
have ht := hS i
change sort S.val (δ₁ i) = - sort S.val (δ₂ i) at ht
have h : sort S.val (δ₂ i) = - sort S.val (δ₁ i) := by
rw [ht]
ring
rw [h]
simp
rfl
· intro i
rw [P_δ₁]
rfl
· intro i
rw [P_δ₂]
have ht := hS i
change sort S.val (δ₁ i) = - sort S.val (δ₂ i) at ht
have h : sort S.val (δ₂ i) = - sort S.val (δ₁ i) := by
rw [ht]
ring
rw [h]
simp
rfl
end VectorLikeEvenPlane

View file

@ -135,26 +135,22 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ)).LinSols}
rw [P_P_P!_accCube' g f hfg] at h1
simp at h1
cases h1 <;> rename_i h1
apply Or.inl
linear_combination h1
cases h1 <;> rename_i h1
apply Or.inr
apply Or.inl
linear_combination -(1 * h1)
apply Or.inr
apply Or.inr
exact h1
· apply Or.inl
linear_combination h1
· cases h1 <;> rename_i h1
· refine Or.inr (Or.inl ?_)
linear_combination -(1 * h1)
· exact Or.inr (Or.inr h1)
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ)).LinSols}
(LIC : LineInCubicPerm S) : LineInPlaneCond S := by
refine @Prop_three (2 * n.succ.succ) LineInPlaneProp S (δ!₂ (Fin.last n)) (δ!₁ (Fin.last n))
δ!₄ ?_ ?_ ?_ ?_
simp [Fin.ext_iff, δ!₂, δ!₁]
simp [Fin.ext_iff, δ!₂, δ!₄]
simp [Fin.ext_iff, δ!₁, δ!₄]
omega
intro M
exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
· simp [Fin.ext_iff, δ!₂, δ!₁]
· simp [Fin.ext_iff, δ!₂, δ!₄]
· simp [Fin.ext_iff, δ!₁, δ!₄]
omega
· exact fun M => lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ)).Sols}
(LIC : LineInCubicPerm S.1.1) : ConstAbs S.val :=

View file

@ -72,9 +72,7 @@ lemma anomalyFree_param {S : (PureU1 (2 * n.succ)).Sols}
have hC := S.cubicSol
rw [hS] at hC
change (accCube (2 * n.succ)) (P g + P! f) = 0 at hC
erw [TriLinearSymm.toCubic_add] at hC
erw [P_accCube] at hC
erw [P!_accCube] at hC
erw [TriLinearSymm.toCubic_add, P_accCube, P!_accCube] at hC
linear_combination hC / 3
/-- A proposition on a solution which is true if `accCubeTriLinSymm (P g, P g, P! f) ≠ 0`.
@ -114,9 +112,9 @@ lemma generic_or_special (S : (PureU1 (2 * n.succ)).Sols) :
have h1 : accCubeTriLinSymm (P g) (P g) (P! f) ≠ 0
accCubeTriLinSymm (P g) (P g) (P! f) = 0 := by
exact ne_or_eq _ _
cases h1 <;> rename_i h1
exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩)
exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩)
rcases h1 with h1 | h1
· exact Or.inl (genericCase_exists S ⟨g, f, h, h1⟩)
· exact Or.inr (specialCase_exists S ⟨g, f, h, h1⟩)
theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : GenericCase S) :
∃ g f a, S = parameterization g f a := by
@ -128,11 +126,11 @@ theorem generic_case {S : (PureU1 (2 * n.succ)).Sols} (h : GenericCase S) :
change S.val = _ • (_ • P g + _• P! f)
rw [anomalyFree_param _ _ hS]
rw [neg_neg, ← smul_add, smul_smul, inv_mul_cancel, one_smul]
exact hS
have h := h g f hS
rw [anomalyFree_param _ _ hS] at h
simp at h
exact h
· exact hS
· have h := h g f hS
rw [anomalyFree_param _ _ hS] at h
simp at h
exact h
lemma special_case_lineInCubic {S : (PureU1 (2 * n.succ)).Sols}
(h : SpecialCase S) : LineInCubic S.1.1 := by
@ -154,9 +152,8 @@ 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)) :
LineInCubicPerm S.1.1 := by
intro M
exact special_case_lineInCubic (h 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),

View file

@ -110,9 +110,9 @@ theorem linesInPlane_constAbs {S : (PureU1 (n.succ.succ.succ.succ.succ)).LinSols
(hS : LineInPlaneCond S) : ConstAbs S.val := by
intro i j
by_cases hij : i ≠ j
exact linesInPlane_eq_sq hS i j hij
simp at hij
rw [hij]
· exact linesInPlane_eq_sq hS i j hij
· simp 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
@ -131,29 +131,29 @@ lemma linesInPlane_four (S : (PureU1 4).Sols) (hS : LineInPlaneCond S.1.1) :
have h1 : S.val (2 : Fin 4) = S.val (3 : Fin 4) := by
linear_combination l012 / 2 + -1 * l013 / 2
by_cases h2 : S.val (0 : Fin 4) = S.val (2 : Fin 4)
simp_all
have h3 : S.val (1 : Fin 4) = - 3 * S.val (2 : Fin 4) := by
linear_combination l012 + 3 * h1
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
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
linear_combination l012 + h1
simp_all
simp_all
have h4 : S.val (0 : Fin 4) = - 3 * S.val (3 : Fin 4) := by
linear_combination l023
have h5 : S.val (1 : Fin 4) = S.val (3 : Fin 4) := by
linear_combination l013 - 1 * h4
rw [h4, h5] at hcube
have h6 : S.val (3 : Fin 4) ^ 3 = 0 := by
linear_combination -1 * hcube / 24
simp at h6
simp_all
· simp_all
have h3 : S.val (1 : Fin 4) = - 3 * S.val (2 : Fin 4) := by
linear_combination l012 + 3 * h1
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
· 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
linear_combination l012 + h1
simp_all
· simp_all
have h4 : S.val (0 : Fin 4) = - 3 * S.val (3 : Fin 4) := by
linear_combination l023
have h5 : S.val (1 : Fin 4) = S.val (3 : Fin 4) := by
linear_combination l013 - 1 * h4
rw [h4, h5] at hcube
have h6 : S.val (3 : Fin 4) ^ 3 = 0 := by
linear_combination -1 * hcube / 24
simp at h6
simp_all
lemma linesInPlane_eq_sq_four {S : (PureU1 4).Sols}
(hS : LineInPlaneCond S.1.1) : ∀ (i j : Fin 4) (_ : i ≠ j),
@ -169,14 +169,14 @@ lemma linesInPlane_constAbs_four (S : (PureU1 4).Sols)
(hS : LineInPlaneCond S.1.1) : ConstAbs S.val := by
intro i j
by_cases hij : i ≠ j
exact linesInPlane_eq_sq_four hS i j hij
simp at hij
rw [hij]
· exact linesInPlane_eq_sq_four hS i j hij
· simp at hij
rw [hij]
theorem linesInPlane_constAbs_AF (S : (PureU1 (n.succ.succ.succ.succ)).Sols)
(hS : LineInPlaneCond S.1.1) : ConstAbs S.val := by
induction n
exact linesInPlane_constAbs_four S hS
exact linesInPlane_constAbs hS
· exact linesInPlane_constAbs_four S hS
· exact linesInPlane_constAbs hS
end PureU1

View file

@ -137,25 +137,22 @@ lemma lineInCubicPerm_last_cond {S : (PureU1 (2 * n.succ.succ+1)).LinSols}
rw [P_P_P!_accCube' g f hfg] at h1
simp at h1
cases h1 <;> rename_i h1
apply Or.inl
linear_combination h1
cases h1 <;> rename_i h1
apply Or.inr
apply Or.inl
linear_combination h1
apply Or.inr
apply Or.inr
linear_combination h1
· apply Or.inl
linear_combination h1
· cases h1 <;> rename_i h1
· refine Or.inr (Or.inl ?_)
linear_combination h1
· refine Or.inr (Or.inr ?_)
linear_combination h1
lemma lineInCubicPerm_last_perm {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
(LIC : LineInCubicPerm S) : LineInPlaneCond S := by
refine @Prop_three (2 * n.succ.succ + 1) LineInPlaneProp S (δ!₂ 0) (δ!₁ 0)
δ!₃ ?_ ?_ ?_ ?_
simp [Fin.ext_iff, δ!₂, δ!₁]
simp [Fin.ext_iff, δ!₂, δ!₃]
simp [Fin.ext_iff, δ!₁, δ!₃]
intro M
exact lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
· simp [Fin.ext_iff, δ!₂, δ!₁]
· simp [Fin.ext_iff, δ!₂, δ!₃]
· simp [Fin.ext_iff, δ!₁, δ!₃]
· exact fun M => lineInCubicPerm_last_cond (lineInCubicPerm_permute LIC M)
lemma lineInCubicPerm_constAbs {S : (PureU1 (2 * n.succ.succ + 1)).LinSols}
(LIC : LineInCubicPerm S) : ConstAbs S.val :=

View file

@ -149,19 +149,23 @@ lemma pairSwap_other {n : } (i j k : Fin n) (hik : i ≠ k) (hjk : j ≠ k) :
(pairSwap i j).toFun k = k := by
simp [pairSwap]
split
simp_all
split
simp_all
rfl
· rename_i h
exact False.elim (hik (id (Eq.symm h)))
· split
· rename_i h
exact False.elim (hjk (id (Eq.symm h)))
· rfl
lemma pairSwap_inv_other {n : } {i j k : Fin n} (hik : i ≠ k) (hjk : j ≠ k) :
(pairSwap i j).invFun k = k := by
simp [pairSwap]
split
simp_all
split
simp_all
rfl
· rename_i h
exact False.elim (hik (id (Eq.symm h)))
· split
· rename_i h
exact False.elim (hjk (id (Eq.symm h)))
· rfl
/-- A permutation of fermions which takes one ordered subset into another. -/
noncomputable def permOfInjection (f g : Fin m ↪ Fin n) : (FamilyPermutations n).group :=