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

@ -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