refactor: Free simps

This commit is contained in:
jstoobysmith 2025-01-29 16:41:10 +00:00
parent e5c85ac109
commit 22636db606
9 changed files with 205 additions and 171 deletions

View file

@ -307,7 +307,7 @@ lemma bosonicProj_mul (a b : 𝓕.CrAnAlgebra) :
conv_lhs =>
rw [← bosonicProj_add_fermionicProj a]
rw [← bosonicProj_add_fermionicProj b]
simp [mul_add, add_mul]
simp only [mul_add, add_mul, map_add, Submodule.coe_add]
rw [bosonicProj_of_mem_bosonic]
conv_lhs =>
left
@ -317,7 +317,7 @@ lemma bosonicProj_mul (a b : 𝓕.CrAnAlgebra) :
have h1 : fermionic = fermionic + bosonic := by simp
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp)]
conv_lhs =>
right
@ -327,23 +327,27 @@ lemma bosonicProj_mul (a b : 𝓕.CrAnAlgebra) :
have h1 : fermionic = bosonic + fermionic := by simp
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp)]
conv_lhs =>
right
right
rw [bosonicProj_of_mem_bosonic _
(by
have h1 : bosonic = fermionic + fermionic := by simp; rfl
have h1 : bosonic = fermionic + fermionic := by
simp only [add_eq_mul, instCommGroup, mul_self]
rfl
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp)]
simp
· have h1 : bosonic = bosonic + bosonic := by simp; rfl
simp only [ZeroMemClass.coe_zero, add_zero, zero_add]
· have h1 : bosonic = bosonic + bosonic := by
simp only [add_eq_mul, instCommGroup, mul_self]
rfl
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp
lemma fermionicProj_mul (a b : 𝓕.CrAnAlgebra) :
@ -352,16 +356,18 @@ lemma fermionicProj_mul (a b : 𝓕.CrAnAlgebra) :
conv_lhs =>
rw [← bosonicProj_add_fermionicProj a]
rw [← bosonicProj_add_fermionicProj b]
simp [mul_add, add_mul]
simp only [mul_add, add_mul, map_add, Submodule.coe_add]
conv_lhs =>
left
left
rw [fermionicProj_of_mem_bosonic _
(by
have h1 : bosonic = bosonic + bosonic := by simp; rfl
have h1 : bosonic = bosonic + bosonic := by
simp only [add_eq_mul, instCommGroup, mul_self]
rfl
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp)]
conv_lhs =>
left
@ -371,7 +377,7 @@ lemma fermionicProj_mul (a b : 𝓕.CrAnAlgebra) :
have h1 : fermionic = fermionic + bosonic := by simp
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp)]
conv_lhs =>
right
@ -381,19 +387,21 @@ lemma fermionicProj_mul (a b : 𝓕.CrAnAlgebra) :
have h1 : fermionic = bosonic + fermionic := by simp
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp)]
conv_lhs =>
right
right
rw [fermionicProj_of_mem_bosonic _
(by
have h1 : bosonic = fermionic + fermionic := by simp; rfl
have h1 : bosonic = fermionic + fermionic := by
simp only [add_eq_mul, instCommGroup, mul_self]
rfl
conv_lhs => rw [h1]
apply crAnAlgebraGrade.mul_mem
simp
simp only [SetLike.coe_mem]
simp)]
simp
simp only [ZeroMemClass.coe_zero, zero_add, add_zero]
abel
end

View file

@ -446,32 +446,42 @@ lemma summerCommute_jacobi_ofCrAnList (φs1 φs2 φs3 : List 𝓕.CrAnStates) :
(- 𝓢(𝓕 |>ₛ φs2, 𝓕 |>ₛ φs3) • [ofCrAnList φs3, [ofCrAnList φs1, ofCrAnList φs2]ₛca]ₛca -
𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs2) • [ofCrAnList φs2, [ofCrAnList φs3, ofCrAnList φs1]ₛca]ₛca) := by
repeat rw [superCommute_ofCrAnList_ofCrAnList]
simp
simp only [instCommGroup, map_sub, map_smul, neg_smul]
repeat rw [superCommute_ofCrAnList_ofCrAnList]
simp only [instCommGroup.eq_1, ofList_append_eq_mul, List.append_assoc]
by_cases h1 : (𝓕 |>ₛ φs1) = bosonic <;>
by_cases h2 : (𝓕 |>ₛ φs2) = bosonic <;>
by_cases h3 : (𝓕 |>ₛ φs3) = bosonic
· simp [h1, h2, exchangeSign_bosonic, h3, mul_one, one_smul]
· simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub]
abel
· simp [h1, h2, exchangeSign_bosonic, bosonic_exchangeSign, mul_one, one_smul]
· simp only [h1, h2, bosonic_exchangeSign, one_smul, mul_bosonic, mul_self, map_one,
exchangeSign_bosonic, neg_sub]
abel
· simp [h1, bosonic_exchangeSign, h3, exchangeSign_bosonic, mul_one, one_smul]
· simp only [h1, h3, mul_bosonic, bosonic_exchangeSign, one_smul, exchangeSign_bosonic, neg_sub,
mul_self, map_one]
abel
· simp at h1 h2 h3
simp [h1, h2, h3]
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
simp only [h1, h2, h3, mul_self, bosonic_exchangeSign, one_smul,
fermionic_exchangeSign_fermionic, neg_smul, neg_sub, bosonic_mul_fermionic, sub_neg_eq_add,
mul_bosonic, smul_add, exchangeSign_bosonic]
abel
· simp at h1 h2 h3
simp [h1, h2, h3]
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
simp only [h1, h2, h3, mul_self, map_one, one_smul, exchangeSign_bosonic, mul_bosonic,
bosonic_exchangeSign, bosonic_mul_fermionic, neg_sub]
abel
· simp at h1 h2 h3
simp [h1, h2, h3]
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
simp only [h1, h2, h3, bosonic_mul_fermionic, fermionic_exchangeSign_fermionic, neg_smul,
one_smul, sub_neg_eq_add, bosonic_exchangeSign, mul_bosonic, smul_add, exchangeSign_bosonic,
neg_sub, mul_self]
abel
· simp at h1 h2 h3
simp [h1, h2, h3]
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
simp only [h1, h2, h3, mul_bosonic, fermionic_exchangeSign_fermionic, neg_smul, one_smul,
sub_neg_eq_add, exchangeSign_bosonic, bosonic_mul_fermionic, smul_add, mul_self,
bosonic_exchangeSign, neg_sub]
abel
· simp at h1 h2 h3
simp [h1, h2, h3]
· simp only [neq_bosonic_iff_eq_fermionic] at h1 h2 h3
simp only [h1, h2, h3, mul_self, map_one, one_smul, fermionic_exchangeSign_fermionic, neg_smul,
neg_sub]
abel
/-!
@ -488,14 +498,14 @@ lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic}
apply Submodule.span_induction (p := p)
· intro x hx
obtain ⟨φs, rfl, hφs⟩ := hx
simp [p]
simp only [add_eq_mul, instCommGroup, p]
let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f1) : Prop :=
[a2, ofCrAnList φs]ₛca ∈ statisticSubmodule (f1 + f2)
change p a ha
apply Submodule.span_induction (p := p)
· intro x hx
obtain ⟨φs', rfl, hφs'⟩ := hx
simp [p]
simp only [add_eq_mul, instCommGroup, p]
rw [superCommute_ofCrAnList_ofCrAnList]
apply Submodule.sub_mem _
· apply ofCrAnList_mem_statisticSubmodule_of
@ -506,18 +516,18 @@ lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic}
rw [mul_comm]
· simp [p]
· intro x y hx hy hp1 hp2
simp [p]
simp only [add_eq_mul, instCommGroup, map_add, LinearMap.add_apply, p]
exact Submodule.add_mem _ hp1 hp2
· intro c x hx hp1
simp [p]
simp only [add_eq_mul, instCommGroup, map_smul, LinearMap.smul_apply, p]
exact Submodule.smul_mem _ c hp1
· exact ha
· simp [p]
· intro x y hx hy hp1 hp2
simp [p]
simp only [add_eq_mul, instCommGroup, map_add, p]
exact Submodule.add_mem _ hp1 hp2
· intro c x hx hp1
simp [p]
simp only [add_eq_mul, instCommGroup, map_smul, p]
exact Submodule.smul_mem _ c hp1
· exact hb
@ -536,19 +546,19 @@ lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra}
apply Submodule.span_induction (p := p)
· intro x hx
obtain ⟨φs', rfl, hφs'⟩ := hx
simp [p]
simp only [p]
rw [superCommute_ofCrAnList_ofCrAnList]
simp [hφs, ofCrAnList_append]
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
· exact ha
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [p, map_add, mul_add, add_mul]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
@ -569,19 +579,19 @@ lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra}
apply Submodule.span_induction (p := p)
· intro x hx
obtain ⟨φs', rfl, hφs'⟩ := hx
simp [p]
simp only [p]
rw [superCommute_ofCrAnList_ofCrAnList]
simp [hφs, hφs', ofCrAnList_append]
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
· exact ha
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [p, map_add, mul_add, add_mul]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
@ -602,19 +612,19 @@ lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra}
apply Submodule.span_induction (p := p)
· intro x hx
obtain ⟨φs', rfl, hφs'⟩ := hx
simp [p]
simp only [p]
rw [superCommute_ofCrAnList_ofCrAnList]
simp [hφs, hφs', ofCrAnList_append]
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
· exact ha
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [map_add, mul_add, add_mul, p]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
@ -661,19 +671,19 @@ lemma superCommute_fermionic_fermionic {a b : 𝓕.CrAnAlgebra}
apply Submodule.span_induction (p := p)
· intro x hx
obtain ⟨φs', rfl, hφs'⟩ := hx
simp [p]
simp only [p]
rw [superCommute_ofCrAnList_ofCrAnList]
simp [hφs, hφs', ofCrAnList_append]
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [p, map_add, LinearMap.add_apply, add_mul, mul_add]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
· exact ha
· simp [p]
· intro x y hx hy hp1 hp2
simp_all [p, mul_add, add_mul]
simp_all only [map_add, mul_add, add_mul, p]
abel
· intro c x hx hp1
simp_all [p, smul_sub]
@ -692,9 +702,10 @@ lemma superCommute_expand_bosonicProj_fermionicProj (a b : 𝓕.CrAnAlgebra) :
fermionicProj a * bosonicProj b - bosonicProj b * fermionicProj a +
fermionicProj a * fermionicProj b + fermionicProj b * fermionicProj a := by
conv_lhs => rw [← bosonicProj_add_fermionicProj a, ← bosonicProj_add_fermionicProj b]
simp
rw [superCommute_bonsonic (by simp), superCommute_bosonic_fermionic (by simp) (by simp),
simp only [map_add, LinearMap.add_apply]
rw [superCommute_bonsonic (by simp),
superCommute_fermionic_bonsonic (by simp) (by simp),
superCommute_bosonic_fermionic (by simp) (by simp),
superCommute_fermionic_fermionic (by simp) (by simp)]
abel
@ -784,14 +795,14 @@ lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List
apply Submodule.span_induction (p := p)
· intro a ha
obtain ⟨φs, rfl, hφs⟩ := ha
simp [p]
simp only [List.get_eq_getElem, p]
rw [superCommute_ofCrAnList_ofCrAnList_eq_sum]
congr
funext n
simp [hφs]
· simp [p]
· intro x y hx hy hpx hpy
simp_all [p]
simp_all only [List.get_eq_getElem, map_add, LinearMap.add_apply, p]
rw [← Finset.sum_add_distrib]
congr
funext n
@ -813,20 +824,22 @@ lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : Lis
apply Submodule.span_induction (p := p)
· intro a ha
obtain ⟨φs, rfl, hφs⟩ := ha
simp [p]
simp only [instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, p]
rw [superCommute_ofCrAnList_ofCrAnList_eq_sum]
congr
funext n
simp [hφs]
· simp [p]
· intro x y hx hy hpx hpy
simp_all [p]
simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_add,
LinearMap.add_apply]
rw [← Finset.sum_add_distrib]
congr
funext n
simp [mul_add, add_mul]
· intro c x hx hpx
simp_all [p, Finset.smul_sum]
simp_all only [p, instCommGroup, List.get_eq_getElem, Algebra.smul_mul_assoc, map_smul,
LinearMap.smul_apply, Finset.smul_sum, Algebra.mul_smul_comm]
congr
funext x
simp [smul_smul, mul_comm]
@ -837,12 +850,12 @@ lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates}
(𝓕 |>ₛ φs) ≠ (𝓕 |>ₛ φs') [ofCrAnList φs, ofCrAnList φs']ₛca = 0 := by
by_cases h0 : [ofCrAnList φs, ofCrAnList φs']ₛca = 0
· simp [h0]
simp [h0]
simp only [ne_eq, h0, or_false]
by_contra hn
refine h0 (eq_zero_of_bosonic_and_fermionic ?_ h)
by_cases hc : (𝓕 |>ₛ φs) = bosonic
· have h1 : bosonic = bosonic + bosonic := by
simp
simp only [add_eq_mul, instCommGroup, mul_self]
rfl
rw [h1]
apply superCommute_grade
@ -850,7 +863,7 @@ lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates}
apply ofCrAnList_mem_statisticSubmodule_of _ _
rw [← hn, hc]
· have h1 : bosonic = fermionic + fermionic := by
simp
simp only [add_eq_mul, instCommGroup, mul_self]
rfl
rw [h1]
apply superCommute_grade

View file

@ -46,21 +46,21 @@ lemma timeOrder_timeOrder_mid (a b c : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b * c) =
apply Submodule.span_induction
· intro x hx
obtain ⟨φs, rfl⟩ := hx
simp [pc]
simp only [ofListBasis_eq_ofList, pc]
let pb (b : 𝓕.CrAnAlgebra) (hb : b ∈ Submodule.span (Set.range ofCrAnListBasis)) :
Prop := 𝓣ᶠ(a * b * ofCrAnList φs) = 𝓣ᶠ(a * 𝓣ᶠ(b) * ofCrAnList φs)
change pb b (Basis.mem_span _ b)
apply Submodule.span_induction
· intro x hx
obtain ⟨φs', rfl⟩ := hx
simp [pb]
simp only [ofListBasis_eq_ofList, pb]
let pa (a : 𝓕.CrAnAlgebra) (ha : a ∈ Submodule.span (Set.range ofCrAnListBasis)) :
Prop := 𝓣ᶠ(a * ofCrAnList φs' * ofCrAnList φs) = 𝓣ᶠ(a * 𝓣ᶠ(ofCrAnList φs') * ofCrAnList φs)
change pa a (Basis.mem_span _ a)
apply Submodule.span_induction
· intro x hx
obtain ⟨φs'', rfl⟩ := hx
simp [pa]
simp only [ofListBasis_eq_ofList, pa]
rw [timeOrder_ofCrAnList]
simp only [← ofCrAnList_append, Algebra.mul_smul_comm,
Algebra.smul_mul_assoc, map_smul]
@ -185,21 +185,21 @@ lemma timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel
{φ1 φ2 : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ1 φ2) (a : 𝓕.CrAnAlgebra) :
𝓣ᶠ([a, [ofCrAnState φ1, ofCrAnState φ2]ₛca]ₛca) = 0 := by
rw [← bosonicProj_add_fermionicProj a]
simp
simp only [map_add, LinearMap.add_apply]
rw [bosonic_superCommute (Submodule.coe_mem (bosonicProj a))]
simp
simp only [map_sub]
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
simp
simp only [sub_self, zero_add]
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton]
rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h' | h'
· rw [superCommute_bonsonic h']
simp [ofCrAnList_singleton]
simp only [ofCrAnList_singleton, map_sub]
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
simp
· rw [superCommute_fermionic_fermionic (Submodule.coe_mem (fermionicProj a)) h']
simp [ofCrAnList_singleton]
simp only [ofCrAnList_singleton, map_add]
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h]
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h]
simp
@ -210,11 +210,13 @@ lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
rw [summerCommute_jacobi_ofCrAnList]
simp [ofCrAnList_singleton]
simp only [instCommGroup.eq_1, ofList_singleton, ofCrAnList_singleton, neg_smul, map_smul,
map_sub, map_neg, smul_eq_zero]
right
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12]
rw [superCommute_ofCrAnState_ofCrAnState_symm φ3]
simp
simp only [smul_zero, neg_zero, instCommGroup.eq_1, neg_smul, map_neg, map_smul, smul_neg,
sub_neg_eq_add, zero_add, smul_eq_zero]
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13]
simp
@ -224,12 +226,13 @@ lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel'
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
rw [summerCommute_jacobi_ofCrAnList]
simp [ofCrAnList_singleton]
simp only [instCommGroup.eq_1, ofList_singleton, ofCrAnList_singleton, neg_smul, map_smul,
map_sub, map_neg, smul_eq_zero]
right
rw [superCommute_ofCrAnState_ofCrAnState_symm φ1]
simp
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, smul_neg, neg_neg]
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12]
simp
simp only [smul_zero, zero_sub, neg_eq_zero, smul_eq_zero]
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13]
simp
@ -239,35 +242,35 @@ lemma timeOrder_superCommute_ofCrAnState_superCommute_all_not_crAnTimeOrderRel
crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧
crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2)) :
𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by
simp at h
simp only [not_and] at h
by_cases h23 : ¬ crAnTimeOrderRel φ2 φ3
· simp_all
· simp_all only [IsEmpty.forall_iff, implies_true]
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h23]
simp_all
simp_all only [Decidable.not_not, forall_const]
by_cases h32 : ¬ crAnTimeOrderRel φ3 φ2
· simp_all
· simp_all only [not_false_eq_true, implies_true]
rw [superCommute_ofCrAnState_ofCrAnState_symm]
simp
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, neg_eq_zero, smul_eq_zero]
rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h32]
simp
simp_all
simp_all only [imp_false, Decidable.not_not]
by_cases h12 : ¬ crAnTimeOrderRel φ1 φ2
· have h13 : ¬ crAnTimeOrderRel φ1 φ3 := by
intro h13
apply h12
exact IsTrans.trans φ1 φ3 φ2 h13 h32
rw [timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel h12 h13]
simp_all
simp_all only [Decidable.not_not, forall_const]
have h13 : crAnTimeOrderRel φ1 φ3 := IsTrans.trans φ1 φ2 φ3 h12 h23
simp_all
simp_all only [forall_const]
by_cases h21 : ¬ crAnTimeOrderRel φ2 φ1
· simp_all
· simp_all only [IsEmpty.forall_iff]
have h31 : ¬ crAnTimeOrderRel φ3 φ1 := by
intro h31
apply h21
exact IsTrans.trans φ2 φ3 φ1 h23 h31
rw [timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel' h21 h31]
simp_all
simp_all only [Decidable.not_not, forall_const]
refine False.elim (h ?_)
exact IsTrans.trans φ3 φ2 φ1 h32 h21

View file

@ -106,7 +106,7 @@ lemma ι_superCommute_zero_of_fermionic (φ ψ : 𝓕.CrAnStates)
ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] at h ⊢
rcases statistic_neq_of_superCommute_fermionic h with h | h
· simp [ofCrAnList_singleton]
· simp only [ofCrAnList_singleton]
apply ι_superCommute_of_diff_statistic
simpa using h
· simp [h]
@ -116,7 +116,7 @@ lemma ι_superCommute_ofCrAnState_ofCrAnState_bosonic_or_zero (φ ψ : 𝓕.CrAn
ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by
rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [ψ] with h | h
· simp_all [ofCrAnList_singleton]
· simp_all [ofCrAnList_singleton]
· simp_all only [ofCrAnList_singleton]
right
exact ι_superCommute_zero_of_fermionic _ _ h
@ -187,7 +187,7 @@ lemma ι_superCommute_ofCrAnState_ofCrAnState_mem_center (φ ψ : 𝓕.CrAnState
have h0 := ι_commute_crAnAlgebra_superCommute_ofCrAnState_ofCrAnState φ ψ a
trans ι ((superCommute (ofCrAnState φ)) (ofCrAnState ψ)) * ι a + 0
swap
simp
simp only [add_zero]
rw [← h0]
abel
@ -208,7 +208,7 @@ lemma ι_eq_zero_iff_mem_ideal (x : CrAnAlgebra 𝓕) :
lemma bosonicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x ∈ 𝓕.fieldOpIdealSet) :
x.bosonicProj.1 ∈ 𝓕.fieldOpIdealSet x.bosonicProj = 0 := by
have hx' := hx
simp [fieldOpIdealSet] at hx
simp only [fieldOpIdealSet, exists_prop, Set.mem_setOf_eq] at hx
rcases hx with ⟨φ1, φ2, φ3, rfl⟩ | ⟨φc, φc', hφc, hφc', rfl⟩ | ⟨φa, φa', hφa, hφa', rfl⟩ |
⟨φ, φ', hdiff, rfl⟩
· rcases superCommute_superCommute_ofCrAnState_bosonic_or_fermionic φ1 φ2 φ3 with h | h
@ -239,7 +239,7 @@ lemma bosonicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x ∈
lemma fermionicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x ∈ 𝓕.fieldOpIdealSet) :
x.fermionicProj.1 ∈ 𝓕.fieldOpIdealSet x.fermionicProj = 0 := by
have hx' := hx
simp [fieldOpIdealSet] at hx
simp only [fieldOpIdealSet, exists_prop, Set.mem_setOf_eq] at hx
rcases hx with ⟨φ1, φ2, φ3, rfl⟩ | ⟨φc, φc', hφc, hφc', rfl⟩ | ⟨φa, φa', hφa, hφa', rfl⟩ |
⟨φ, φ', hdiff, rfl⟩
· rcases superCommute_superCommute_ofCrAnState_bosonic_or_fermionic φ1 φ2 φ3 with h | h
@ -275,11 +275,11 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
change p x hx
apply AddSubgroup.closure_induction
· intro x hx
simp [p]
simp only [p]
obtain ⟨a, ha, b, hb, rfl⟩ := Set.mem_mul.mp hx
obtain ⟨d, hd, y, hy, rfl⟩ := Set.mem_mul.mp ha
rw [bosonicProj_mul, bosonicProj_mul, fermionicProj_mul]
simp [mul_add, add_mul]
simp only [add_mul]
rcases fermionicProj_mem_fieldOpIdealSet_or_zero y hy with hfy | hfy
<;> rcases bosonicProj_mem_fieldOpIdealSet_or_zero y hy with hby | hby
· apply TwoSidedIdeal.add_mem
@ -292,7 +292,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use bosonicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (bosonicProj y).1
simp [hby]
· use ↑(bosonicProj b)
@ -305,7 +305,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use fermionicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (fermionicProj y).1
simp [hby, hfy]
· use ↑(bosonicProj b)
@ -319,7 +319,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use bosonicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (fermionicProj y).1
simp [hby, hfy]
· use ↑(fermionicProj b)
@ -332,12 +332,12 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use fermionicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (bosonicProj y).1
simp [hby, hfy]
· use ↑(fermionicProj b)
simp
· simp [hby]
· simp only [hby, ZeroMemClass.coe_zero, mul_zero, zero_mul, zero_add, add_zero]
apply TwoSidedIdeal.add_mem
· /- fermion, fermion, boson mem-/
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
@ -347,7 +347,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use fermionicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (fermionicProj y).1
simp [hby, hfy]
· use ↑(bosonicProj b)
@ -360,12 +360,12 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use bosonicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (fermionicProj y).1
simp [hby, hfy]
· use ↑(fermionicProj b)
simp
· simp [hfy]
· simp only [hfy, ZeroMemClass.coe_zero, mul_zero, zero_mul, add_zero, zero_add]
apply TwoSidedIdeal.add_mem
· /- boson, boson, boson mem-/
rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure]
@ -375,7 +375,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use bosonicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (bosonicProj y).1
simp [hby]
· use ↑(bosonicProj b)
@ -388,7 +388,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
apply And.intro
· apply Set.mem_mul.mpr
use fermionicProj d
simp
simp only [Set.mem_univ, mul_eq_mul_left_iff, ZeroMemClass.coe_eq_zero, true_and]
use (bosonicProj y).1
simp [hby, hfy]
· use ↑(fermionicProj b)
@ -396,7 +396,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa
· simp [hfy, hby]
· simp [p]
· intro x y hx hy hpx hpy
simp_all [p]
simp_all only [map_add, Submodule.coe_add, p]
apply TwoSidedIdeal.add_mem
exact hpx
exact hpy
@ -408,7 +408,7 @@ lemma fermionicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.s
have hb := bosonicProj_mem_ideal x hx
rw [← ι_eq_zero_iff_mem_ideal] at hx hb ⊢
rw [← bosonicProj_add_fermionicProj x] at hx
simp at hx
simp only [map_add] at hx
simp_all
lemma ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero (x : CrAnAlgebra 𝓕) :

View file

@ -79,10 +79,10 @@ lemma superCommuteRight_eq_of_equiv (a1 a2 : 𝓕.CrAnAlgebra) (h : a1 ≈ a2) :
rw [superCommuteRight_apply_ι]
apply ι_superCommute_eq_zero_of_ι_left_zero
exact (ι_eq_zero_iff_mem_ideal (a1 - a2)).mpr h
simp_all [superCommuteRight_apply_ι]
simp_all only [superCommuteRight_apply_ι, map_sub, LinearMap.sub_apply]
trans ι ((superCommute a2) b) + 0
rw [← ha1b1]
simp
simp only [add_sub_cancel]
simp
noncomputable def superCommute : FieldOpAlgebra 𝓕 →ₗ[]
@ -95,7 +95,7 @@ noncomputable def superCommute : FieldOpAlgebra 𝓕 →ₗ[]
obtain ⟨b, rfl⟩ := ι_surjective b
rw [← map_add, ι_apply, ι_apply, ι_apply, ι_apply]
rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk]
simp
simp only [LinearMap.add_apply]
rw [superCommuteRight_apply_quot, superCommuteRight_apply_quot, superCommuteRight_apply_quot]
simp
map_smul' c y := by
@ -103,7 +103,7 @@ noncomputable def superCommute : FieldOpAlgebra 𝓕 →ₗ[]
ext b
obtain ⟨b, rfl⟩ := ι_surjective b
rw [← map_smul, ι_apply, ι_apply, ι_apply]
simp
simp only [Quotient.lift_mk, RingHom.id_apply, LinearMap.smul_apply]
rw [superCommuteRight_apply_quot, superCommuteRight_apply_quot]
simp

View file

@ -56,7 +56,7 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 :
rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp, ← crAnTimeOrderSign]
· simp
· intro φ4 hφ4
simp at hφ4
simp only [List.mem_cons, List.mem_singleton, List.not_mem_nil, or_false] at hφ4
rcases hφ4 with hφ4 | hφ4 | hφ4
all_goals
subst hφ4
@ -78,7 +78,7 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 :
rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp231, ← crAnTimeOrderSign]
· simp
· intro φ4 hφ4
simp at hφ4
simp only [List.mem_cons, List.mem_singleton, List.not_mem_nil, or_false] at hφ4
rcases hφ4 with hφ4 | hφ4 | hφ4
all_goals
subst hφ4
@ -99,24 +99,25 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 :
rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp, ← crAnTimeOrderSign]
· simp
· intro φ4 hφ4
simp at hφ4
simp only [List.mem_cons, List.mem_singleton, List.not_mem_nil, or_false] at hφ4
rcases hφ4 with hφ4 | hφ4 | hφ4
all_goals
subst hφ4
simp_all
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
rw [superCommute_ofCrAnList_ofCrAnList]
simp
simp only [List.singleton_append, instCommGroup.eq_1, ofList_singleton, map_sub, map_smul]
rw [superCommute_ofCrAnList_ofCrAnList, superCommute_ofCrAnList_ofCrAnList]
simp [mul_sub, sub_mul, ← ofCrAnList_append]
simp only [List.cons_append, List.nil_append, instCommGroup.eq_1, ofList_singleton, mul_sub, ←
ofCrAnList_append, Algebra.mul_smul_comm, sub_mul, List.append_assoc, Algebra.smul_mul_assoc,
map_sub, map_smul]
rw [h123, h132, h231, h321]
simp [smul_smul]
simp only [smul_smul]
rw [mul_comm, ← smul_smul, mul_comm, ← smul_smul]
rw [← smul_sub, ← smul_sub, smul_smul, mul_comm, ← smul_smul, ← smul_sub]
simp
simp only [smul_eq_zero]
right
rw [← smul_mul_assoc, ← mul_smul_comm, mul_assoc]
rw [← smul_mul_assoc, ← mul_smul_comm]
rw [smul_sub]
rw [← smul_mul_assoc, ← mul_smul_comm]
@ -129,9 +130,10 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 :
congr
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton]
rw [superCommute_ofCrAnList_ofCrAnList]
simp
simp only [List.singleton_append, instCommGroup.eq_1, ofList_singleton, map_sub, map_smul]
rw [superCommute_ofCrAnList_ofCrAnList, superCommute_ofCrAnList_ofCrAnList]
simp [smul_sub]
simp only [List.cons_append, List.nil_append, instCommGroup.eq_1, ofList_singleton, map_sub,
map_smul, smul_sub]
simp_all
lemma ι_timeOrder_superCommute_superCommute_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates}
@ -156,14 +158,14 @@ lemma ι_timeOrder_superCommute_superCommute {φ1 φ2 φ3 : 𝓕.CrAnStates} (a
apply Submodule.span_induction
· intro x hx
obtain ⟨φs, rfl⟩ := hx
simp [pb]
simp only [ofListBasis_eq_ofList, pb]
let pa (a : 𝓕.CrAnAlgebra) (hc : a ∈ Submodule.span (Set.range ofCrAnListBasis)) :
Prop := ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ofCrAnList φs) = 0
change pa a (Basis.mem_span _ a)
apply Submodule.span_induction
· intro x hx
obtain ⟨φs', rfl⟩ := hx
simp [pa]
simp only [ofListBasis_eq_ofList, pa]
exact ι_timeOrder_superCommute_superCommute_ofCrAnList φs' φs
· simp [pa]
· intro x y hx hy hpx hpy
@ -187,7 +189,7 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates}
apply Submodule.span_induction
· intro x hx
obtain ⟨φs, rfl⟩ := hx
simp [pb]
simp only [ofListBasis_eq_ofList, map_mul, pb]
let pa (a : 𝓕.CrAnAlgebra) (hc : a ∈ Submodule.span (Set.range ofCrAnListBasis)) :
Prop := ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * ofCrAnList φs) =
ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a* ofCrAnList φs))
@ -195,22 +197,23 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates}
apply Submodule.span_induction
· intro x hx
obtain ⟨φs', rfl⟩ := hx
simp [pa]
simp only [ofListBasis_eq_ofList, map_mul, pa]
conv_lhs =>
rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, superCommute_ofCrAnList_ofCrAnList]
simp [mul_sub, sub_mul, ← ofCrAnList_append]
rw [timeOrder_ofCrAnList, timeOrder_ofCrAnList]
have h1 : crAnTimeOrderSign (φs' ++ φ :: ψ :: φs) = crAnTimeOrderSign (φs' ++ ψ :: φ :: φs) := by
have h1 : crAnTimeOrderSign (φs' ++ φ :: ψ :: φs) =
crAnTimeOrderSign (φs' ++ ψ :: φ :: φs) := by
trans crAnTimeOrderSign (φs' ++ [φ, ψ] ++ φs)
simp
simp only [List.append_assoc, List.cons_append, List.nil_append]
rw [crAnTimeOrderSign]
have hp : List.Perm [φ,ψ] [ψ,φ] := by exact List.Perm.swap ψ φ []
rw [Wick.koszulSign_perm_eq _ _ φ _ _ _ _ _ hp]
simp
simp only [List.append_assoc, List.cons_append, List.singleton_append]
rfl
simp_all
rw [h1]
simp
simp only [map_smul]
have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ φs' [φ, ψ] φs
(by simp_all)
rw [crAnTimeOrderList, show φs' ++ φ :: ψ :: φs = φs' ++ [φ, ψ] ++ φs by simp, h1]
@ -244,7 +247,8 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates}
rw [← map_mul, ← map_mul, ← map_mul, ← map_mul]
rw [← ofCrAnList_append, ← ofCrAnList_append, ← ofCrAnList_append, ← ofCrAnList_append]
have h1 := insertionSort_of_takeWhile_filter 𝓕.crAnTimeOrderRel φ φs' φs
simp at h1 ⊢
simp only [decide_not, Bool.decide_and, List.append_assoc, List.cons_append,
List.singleton_append, Algebra.mul_smul_comm, map_mul] at h1 ⊢
rw [← h1]
rw [← crAnTimeOrderList]
by_cases hq : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ)
@ -252,16 +256,17 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates}
simp
· rw [crAnTimeOrderSign, Wick.koszulSign_eq_rel_eq_stat _ _, ← crAnTimeOrderSign]
rw [timeOrder_ofCrAnList]
simp
simp only [map_smul, Algebra.mul_smul_comm]
simp only [List.nil_append]
exact hψφ
exact hφψ
simpa using hq
· simp [pa]
· simp only [map_mul, zero_mul, map_zero, mul_zero, pa]
· intro x y hx hy hpx hpy
simp_all [pa,mul_add, add_mul]
· intro x hx hpx
simp_all [pa, hpx]
· simp [pb]
· simp only [map_mul, mul_zero, map_zero, pb]
· intro x y hx hy hpx hpy
simp_all [pb,mul_add, add_mul]
· intro x hx hpx
@ -277,12 +282,13 @@ lemma ι_timeOrder_superCommute_neq_time {φ ψ : 𝓕.CrAnStates}
rcases hφψ with hφψ | hφψ
· rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel ]
have ht := IsTotal.total (r := crAnTimeOrderRel) φ ψ
simp_all
simp_all only [false_and, not_false_eq_true, false_or, mul_zero, zero_mul, map_zero]
simp_all
· rw [superCommute_ofCrAnState_ofCrAnState_symm]
simp
simp only [instCommGroup.eq_1, neg_smul, map_neg, map_smul, mul_neg, Algebra.mul_smul_comm,
neg_mul, Algebra.smul_mul_assoc, neg_eq_zero, smul_eq_zero]
rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel ]
simp
simp only [mul_zero, zero_mul, map_zero, or_true]
simp_all
@ -306,14 +312,14 @@ lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra)
match hc with
| Or.inl hc =>
obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc
simp
simp only [ι_timeOrder_superCommute_superCommute]
| Or.inr (Or.inl hc) =>
obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc
by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa)
· rw [ι_timeOrder_superCommute_eq_time]
simp
simp only [map_mul]
rw [ι_superCommute_of_create_create]
simp
simp only [zero_mul]
· exact hφa
· exact hφb
· exact heqt.1
@ -323,9 +329,9 @@ lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra)
obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc
by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa)
· rw [ι_timeOrder_superCommute_eq_time]
simp
simp only [map_mul]
rw [ι_superCommute_of_annihilate_annihilate]
simp
simp only [zero_mul]
· exact hφa
· exact hφb
· exact heqt.1
@ -335,9 +341,9 @@ lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra)
obtain ⟨φa, φb, hdiff, rfl⟩ := hc
by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa)
· rw [ι_timeOrder_superCommute_eq_time]
simp
simp only [map_mul]
rw [ι_superCommute_of_diff_statistic]
simp
simp only [zero_mul]
· exact hdiff
· exact heqt.1
· exact heqt.2

View file

@ -292,11 +292,12 @@ instance : AddMonoid FieldStatistic where
add a b := a * b
nsmul n a := ∏ (i : Fin n), a
zero_add a := by
cases a <;> simp <;> rfl
cases a <;> simp only [instCommGroup] <;> rfl
add_zero a := by
cases a <;> simp <;> rfl
cases a <;>
simp only [instCommGroup] <;> rfl
add_assoc a b c := by
cases a <;> cases b <;> cases c <;> simp <;> rfl
cases a <;> cases b <;> cases c <;> simp only [instCommGroup] <;> rfl
nsmul_zero a := by
simp only [Finset.univ_eq_empty, Finset.prod_const, instCommGroup, Finset.card_empty, pow_zero]
rfl

View file

@ -284,13 +284,13 @@ lemma koszulSign_eq_rel_eq_stat_append {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal
(h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) →
koszulSign q le (φ :: ψ :: φs) = koszulSign q le φs := by
intro φs
simp [koszulSign, ← mul_assoc]
simp only [koszulSign, ← mul_assoc]
trans 1 * koszulSign q le φs
swap
simp
simp only [one_mul]
congr
simp [koszulSignInsert]
simp_all
simp only [koszulSignInsert, ite_mul, neg_mul]
simp_all only [and_self, ite_true]
rw [koszulSignInsert_eq_rel_eq_stat q le h1 h2 hq]
simp
@ -298,31 +298,31 @@ lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal 𝓕 l
(h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs' φs : List 𝓕) →
koszulSign q le (φs' ++ φ :: ψ :: φs) = koszulSign q le (φs' ++ φs)
| [], φs => by
simp
simp only [List.nil_append]
exact koszulSign_eq_rel_eq_stat_append q le h1 h2 hq φs
| φ'' :: φs', φs => by
simp [koszulSign]
simp only [koszulSign, List.append_eq]
rw [koszulSign_eq_rel_eq_stat h1 h2 hq φs' φs]
simp
simp only [mul_eq_mul_right_iff]
left
trans koszulSignInsert q le φ'' (φ :: ψ :: (φs' ++ φs))
apply koszulSignInsert_eq_perm
refine List.Perm.symm (List.perm_cons_append_cons φ ?_)
exact List.Perm.symm List.perm_middle
rw [koszulSignInsert_eq_remove_same_stat_append q le]
simp_all
simp_all
simp_all
exact h1
exact h2
exact hq
lemma koszulSign_of_sorted : (φs : List 𝓕)
→ (hs : List.Sorted le φs) → koszulSign q le φs = 1
| [], _ => by
simp [koszulSign]
| φ :: φs, h => by
simp [koszulSign]
simp at h
simp only [koszulSign]
simp only [List.sorted_cons] at h
rw [koszulSign_of_sorted φs h.2]
simp
simp only [mul_one]
exact koszulSignInsert_of_le_mem _ _ _ _ h.1
@[simp]
@ -344,14 +344,15 @@ lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕
rw [insertIdx_length_fst_append]
rw [h1, h2]
rw [koszulSign_insertIdx]
simp
simp only [instCommGroup.eq_1, List.take_left', List.length_insertionSort]
rw [koszulSign_insertIdx]
simp [mul_assoc]
simp only [mul_assoc, instCommGroup.eq_1, List.length_insertionSort, List.take_left',
ofList_insertionSort, mul_eq_mul_left_iff]
left
rw [koszulSign_of_append_eq_insertionSort_left φs φs']
simp [mul_assoc]
simp only [mul_assoc, mul_eq_mul_left_iff]
left
simp [mul_comm]
simp only [mul_comm, mul_eq_mul_left_iff]
left
congr 3
· have h2 : (List.insertionSort le φs ++ φ :: φs') =
@ -359,9 +360,10 @@ lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕
rw [← insertIdx_length_fst_append]
simp
rw [insertionSortEquiv_congr _ _ h2.symm]
simp
simp only [Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_mk,
Fin.coe_cast]
rw [insertionSortEquiv_insertionSort_append]
simp
simp only [finCongr_apply, Fin.coe_cast]
rw [insertionSortEquiv_congr _ _ h1.symm]
simp
· rw [insertIdx_length_fst_append]
@ -376,7 +378,7 @@ lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le]
koszulSign q le (φs'' ++ φs ++ φs') =
koszulSign q le (φs'' ++ List.insertionSort le φs ++ φs') * koszulSign q le φs
| [], φs, φs'=> by
simp
simp only [List.nil_append]
exact koszulSign_of_append_eq_insertionSort_left q le φs φs'
| φ'' :: φs'', φs, φs' => by
simp only [koszulSign, List.append_eq]
@ -403,18 +405,19 @@ lemma koszulSign_perm_eq_append [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕)
apply List.Perm.recOn
· simp [motive]
· intro x l1 l2 h ih hxφ
simp_all [motive]
simp [koszulSign, ih]
simp_all only [List.mem_cons, or_true, and_self, implies_true, nonempty_prop, forall_const,
forall_eq_or_imp, List.cons_append, motive]
simp only [koszulSign, ih, mul_eq_mul_right_iff]
left
apply koszulSignInsert_eq_perm
exact (List.perm_append_right_iff φs2).mpr h
· intro x y l h
simp_all [motive]
simp_all only [List.mem_cons, forall_eq_or_imp, List.cons_append]
apply Wick.koszulSign_swap_eq_rel_cons
exact IsTrans.trans y φ x h.1.2 h.2.1.1
exact IsTrans.trans x φ y h.2.1.2 h.1.1
· intro l1 l2 l3 h1 h2 ih1 ih2 h
simp_all [motive]
simp_all only [and_self, implies_true, nonempty_prop, forall_const, motive]
refine (ih2 ?_)
intro φ' hφ
refine h φ' ?_
@ -424,7 +427,7 @@ lemma koszulSign_perm_eq [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) : (φs1
(h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) → (hp : φs.Perm φs') →
koszulSign q le (φs1 ++ φs ++ φs2) = koszulSign q le (φs1 ++ φs' ++ φs2)
| [], φs, φs', φs2, h, hp => by
simp
simp only [List.nil_append]
exact koszulSign_perm_eq_append q le φ φs φs' φs2 hp h
| φ1 :: φs1, φs, φs', φs2, h, hp => by
simp only [koszulSign, List.append_eq]

View file

@ -240,7 +240,7 @@ lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ b
| [], _ => by
simp [koszulSignInsert]
| φ1 :: φs, h => by
simp [koszulSignInsert]
simp only [koszulSignInsert]
rw [if_pos]
· apply koszulSignInsert_of_le_mem
· intro b hb
@ -253,26 +253,26 @@ lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTotal 𝓕 le] [IsTrans
| [] => by
simp [koszulSignInsert]
| φ' :: φs => by
simp [koszulSignInsert]
simp_all
simp only [koszulSignInsert]
simp_all only
by_cases hr : le φ φ'
· simp [hr]
· simp only [hr, ↓reduceIte]
have h1' : le ψ φ' := by
apply IsTrans.trans ψ φ φ' h2 hr
simp [h1']
simp only [h1', ↓reduceIte]
exact koszulSignInsert_eq_rel_eq_stat h1 h2 hq φs
· have hψφ' : ¬ le ψ φ' := by
intro hψφ'
apply hr
apply IsTrans.trans φ ψ φ' h1 hψφ'
simp [hr, hψφ']
simp only [hr, ↓reduceIte, hψφ']
rw [koszulSignInsert_eq_rel_eq_stat h1 h2 hq φs]
lemma koszulSignInsert_eq_remove_same_stat_append {ψ φ φ' : 𝓕} [IsTotal 𝓕 le] [IsTrans 𝓕 le]
(h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) →
koszulSignInsert q le φ' (φ :: ψ :: φs) = koszulSignInsert q le φ' φs := by
intro φs
simp_all [koszulSignInsert]
simp_all only [koszulSignInsert, and_self, ite_true, ite_false, ite_self]
by_cases hφ'φ : le φ' φ
· have hφ'ψ : le φ' ψ := by
apply IsTrans.trans φ' φ ψ hφ'φ h1