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

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