From 48b0a60f34b2ddfeb6d369397139a9c49d6c33ed Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Tue, 28 Jan 2025 11:53:24 +0000 Subject: [PATCH 1/8] feat: Redefine FieldOpAlgebra --- .../Algebras/CrAnAlgebra/Grading.lean | 25 +- .../Algebras/CrAnAlgebra/SuperCommute.lean | 344 ++++++++++++++++++ .../Algebras/FieldOpAlgebra/Basic.lean | 114 +++++- .../Algebras/FieldOpAlgebra/NormalOrder.lean | 2 +- .../FieldStatistics/Basic.lean | 3 + .../FieldStatistics/ExchangeSign.lean | 4 + 6 files changed, 469 insertions(+), 23 deletions(-) diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean index 59758d73..f5a34ef4 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean @@ -3,8 +3,7 @@ Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ -import HepLean.PerturbationTheory.FieldSpecification.NormalOrder -import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute +import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Basic import HepLean.PerturbationTheory.Koszul.KoszulSign import Mathlib.RingTheory.GradedAlgebra.Basic /-! @@ -25,6 +24,19 @@ noncomputable section def statisticSubmodule (f : FieldStatistic) : Submodule ℂ 𝓕.CrAnAlgebra := Submodule.span ℂ {a | ∃ φs, a = ofCrAnList φs ∧ (𝓕 |>ₛ φs) = f} +lemma ofCrAnList_mem_statisticSubmodule_of (φs : List 𝓕.CrAnStates) (f : FieldStatistic) + (h : (𝓕 |>ₛ φs) = f) : + ofCrAnList φs ∈ statisticSubmodule f := by + refine Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩ + +lemma ofCrAnList_bosonic_or_fermionic (φs : List 𝓕.CrAnStates) : + ofCrAnList φs ∈ statisticSubmodule bosonic ∨ ofCrAnList φs ∈ statisticSubmodule fermionic := by + by_cases h : (𝓕 |>ₛ φs) = bosonic + · left + exact ofCrAnList_mem_statisticSubmodule_of φs bosonic h + · right + exact ofCrAnList_mem_statisticSubmodule_of φs fermionic (by simpa using h) + /-- The projection of an element of `CrAnAlgebra` onto it's bosonic part. -/ def bosonicProj : 𝓕.CrAnAlgebra →ₗ[ℂ] statisticSubmodule (𝓕 := 𝓕) bosonic := Basis.constr ofCrAnListBasis ℂ fun φs => @@ -276,6 +288,15 @@ instance : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where fermionicProj_of_fermionic_part, zero_add] conv_rhs => rw [directSum_eq_bosonic_plus_fermionic a] +lemma eq_zero_of_bosonic_and_fermionic {a : 𝓕.CrAnAlgebra} + (hb : a ∈ statisticSubmodule bosonic) (hf : a ∈ statisticSubmodule fermionic) : a = 0 := by + have ha := bosonicProj_of_mem_bosonic a hb + have hb := fermionicProj_of_mem_fermionic a hf + have hc := (bosonicProj_add_fermionicProj a) + rw [ha, hb] at hc + simpa using hc + + end end CrAnAlgebra diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean index 292dc309..59a66053 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Basic +import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Grading /-! # Super Commute @@ -438,6 +439,349 @@ lemma superCommute_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) : · simp · simp [Finset.mul_sum, smul_smul, ofStateList_cons, mul_assoc, FieldStatistic.ofList_cons_eq_mul, mul_comm] +/-! + +## Interaction with grading. + +-/ + +lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic} + (ha : a ∈ statisticSubmodule f1) (hb : b ∈ statisticSubmodule f2) : + [a, b]ₛca ∈ statisticSubmodule (f1 + f2) := by + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop := + [a, a2]ₛca ∈ statisticSubmodule (f1 + f2) + change p b hb + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs, rfl, hφs⟩ := hx + simp [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] + rw [superCommute_ofCrAnList_ofCrAnList] + apply Submodule.sub_mem _ + · apply ofCrAnList_mem_statisticSubmodule_of + rw [ofList_append_eq_mul, hφs, hφs'] + · apply Submodule.smul_mem + apply ofCrAnList_mem_statisticSubmodule_of + rw [ofList_append_eq_mul, hφs, hφs'] + rw [mul_comm] + · simp [p] + · intro x y hx hy hp1 hp2 + simp [p] + exact Submodule.add_mem _ hp1 hp2 + · intro c x hx hp1 + simp [p] + exact Submodule.smul_mem _ c hp1 + · exact ha + · simp [p] + · intro x y hx hy hp1 hp2 + simp [p] + exact Submodule.add_mem _ hp1 hp2 + · intro c x hx hp1 + simp [p] + exact Submodule.smul_mem _ c hp1 + · exact hb + +lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra} + (ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule bosonic) : + [a, b]ₛca = a * b - b * a := by + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := + [a, a2]ₛca = a * a2 - a2 * a + change p b hb + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs, rfl, hφs⟩ := hx + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := + [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 + change p a ha + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs', rfl, hφs'⟩ := hx + simp [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] + 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] + abel + · intro c x hx hp1 + simp_all [p, smul_sub] + · exact hb + + +lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra} + (ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule fermionic) : + [a, b]ₛca = a * b - b * a := by + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := + [a, a2]ₛca = a * a2 - a2 * a + change p b hb + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs, rfl, hφs⟩ := hx + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := + [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 + change p a ha + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs', rfl, hφs'⟩ := hx + simp [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] + 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] + abel + · intro c x hx hp1 + simp_all [p, smul_sub] + · exact hb + + +lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra} + (ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule bosonic) : + [a, b]ₛca = a * b - b * a := by + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := + [a, a2]ₛca = a * a2 - a2 * a + change p b hb + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs, rfl, hφs⟩ := hx + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := + [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 + change p a ha + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs', rfl, hφs'⟩ := hx + simp [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] + 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] + abel + · intro c x hx hp1 + simp_all [p, smul_sub] + · exact hb + +lemma superCommute_bonsonic {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) : + [a, b]ₛca = a * b - b * a := by + rw [← bosonicProj_add_fermionicProj a] + simp only [map_add, LinearMap.add_apply] + rw [superCommute_bosonic_bosonic (by simp) hb, superCommute_fermionic_bonsonic (by simp) hb] + simp only [add_mul, mul_add] + abel + +lemma bosonic_superCommute {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) : + [a, b]ₛca = a * b - b * a := by + rw [← bosonicProj_add_fermionicProj b] + simp only [map_add, LinearMap.add_apply] + rw [superCommute_bosonic_bosonic ha (by simp), superCommute_bosonic_fermionic ha (by simp)] + simp only [add_mul, mul_add] + abel + +lemma superCommute_bonsonic_symm {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) : + [a, b]ₛca = - [b, a]ₛca := by + rw [bosonic_superCommute hb, superCommute_bonsonic hb] + simp + +lemma bonsonic_superCommute_symm {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) : + [a, b]ₛca = - [b, a]ₛca := by + rw [bosonic_superCommute ha, superCommute_bonsonic ha] + simp + +lemma superCommute_fermionic_fermionic {a b : 𝓕.CrAnAlgebra} + (ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) : + [a, b]ₛca = a * b + b * a := by + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := + [a, a2]ₛca = a * a2 + a2 * a + change p b hb + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs, rfl, hφs⟩ := hx + let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := + [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs + ofCrAnList φs * a2 + change p a ha + apply Submodule.span_induction (p := p) + · intro x hx + obtain ⟨φs', rfl, hφs'⟩ := hx + simp [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] + 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] + abel + · intro c x hx hp1 + simp_all [p, smul_sub] + · exact hb + +lemma superCommute_fermionic_fermionic_symm {a b : 𝓕.CrAnAlgebra} + (ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) : + [a, b]ₛca = [b, a]ₛca := by + rw [superCommute_fermionic_fermionic ha hb] + rw [superCommute_fermionic_fermionic hb ha] + abel + +lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List 𝓕.CrAnStates) : + [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule bosonic ∨ + [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic := by + by_cases h1 : (𝓕 |>ₛ φs) = bosonic <;> by_cases h2 : (𝓕 |>ₛ φs') = bosonic + · left + have h : bosonic = bosonic + bosonic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade + apply ofCrAnList_mem_statisticSubmodule_of _ _ h1 + apply ofCrAnList_mem_statisticSubmodule_of _ _ h2 + · right + have h : fermionic = bosonic + fermionic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade + apply ofCrAnList_mem_statisticSubmodule_of _ _ h1 + apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2) + · right + have h : fermionic = fermionic + bosonic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade + apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1) + apply ofCrAnList_mem_statisticSubmodule_of _ _ h2 + · left + have h : bosonic = fermionic + fermionic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade + apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1) + apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2) + + +lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates) + (ha : a ∈ statisticSubmodule bosonic) : + [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), + ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * + ofCrAnList (φs.drop (n + 1)) := by + let p (a : 𝓕.CrAnAlgebra) (ha : a ∈ statisticSubmodule bosonic) : Prop := + [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), + ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * + ofCrAnList (φs.drop (n + 1)) + change p a ha + apply Submodule.span_induction (p := p) + · intro a ha + obtain ⟨φs, rfl, hφs⟩ := ha + simp [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] + rw [← Finset.sum_add_distrib] + congr + funext n + simp [mul_add, add_mul] + · intro c x hx hpx + simp_all [p, Finset.smul_sum] + · exact ha + + +lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates) + (ha : a ∈ statisticSubmodule fermionic) : + [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) • + ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * + ofCrAnList (φs.drop (n + 1)) := by + let p (a : 𝓕.CrAnAlgebra) (ha : a ∈ statisticSubmodule fermionic) : Prop := + [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) • + ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * + ofCrAnList (φs.drop (n + 1)) + change p a ha + apply Submodule.span_induction (p := p) + · intro a ha + obtain ⟨φs, rfl, hφs⟩ := ha + simp [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] + rw [← Finset.sum_add_distrib] + congr + funext n + simp [mul_add, add_mul] + · intro c x hx hpx + simp_all [p, Finset.smul_sum] + congr + funext x + simp [smul_smul, mul_comm] + · exact ha + + +lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates} + (h : [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic) : + (𝓕 |>ₛ φs) ≠ (𝓕 |>ₛ φs') ∨ [ofCrAnList φs, ofCrAnList φs']ₛca = 0 := by + by_cases h0 : [ofCrAnList φs, ofCrAnList φs']ₛca = 0 + · simp [h0] + simp [h0] + by_contra hn + refine h0 (eq_zero_of_bosonic_and_fermionic ?_ h) + by_cases hc : (𝓕 |>ₛ φs) = bosonic + · have h1 : bosonic = bosonic + bosonic := by + simp + rfl + rw [h1] + apply superCommute_grade + apply ofCrAnList_mem_statisticSubmodule_of _ _ hc + apply ofCrAnList_mem_statisticSubmodule_of _ _ + rw [← hn, hc] + · have h1 : bosonic = fermionic + fermionic := by + simp + rfl + rw [h1] + apply superCommute_grade + apply ofCrAnList_mem_statisticSubmodule_of _ _ + simpa using hc + apply ofCrAnList_mem_statisticSubmodule_of _ _ + rw [← hn] + simpa using hc end CrAnAlgebra diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean index 033436f6..9fba2622 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean @@ -3,10 +3,9 @@ Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ -import HepLean.PerturbationTheory.WickContraction.TimeContract -import HepLean.Meta.Remark.Basic -import Mathlib.RingTheory.TwoSidedIdeal.Operations +import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute import Mathlib.Algebra.RingQuot +import Mathlib.RingTheory.TwoSidedIdeal.Operations /-! # Field operator algebra @@ -15,9 +14,7 @@ import Mathlib.Algebra.RingQuot namespace FieldSpecification open CrAnAlgebra -open ProtoOperatorAlgebra open HepLean.List -open WickContraction open FieldStatistic variable (𝓕 : FieldSpecification) @@ -26,8 +23,8 @@ variable (𝓕 : FieldSpecification) This contains e.g. the super-commutor of two creation operators. -/ def fieldOpIdealSet : Set (CrAnAlgebra 𝓕) := { x | - (∃ (φ ψ : 𝓕.CrAnStates) (a : CrAnAlgebra 𝓕), - x = a * [ofCrAnState φ, ofCrAnState ψ]ₛca - [ofCrAnState φ, ofCrAnState ψ]ₛca * a) + (∃ (φ1 φ2 φ3 : 𝓕.CrAnStates), + x = [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) ∨ (∃ (φc φc' : 𝓕.CrAnStates) (_ : 𝓕 |>ᶜ φc = .create) (_ : 𝓕 |>ᶜ φc' = .create), x = [ofCrAnState φc, ofCrAnState φc']ₛca) ∨ (∃ (φa φa' : 𝓕.CrAnStates) (_ : 𝓕 |>ᶜ φa = .annihilate) (_ : 𝓕 |>ᶜ φa' = .annihilate), @@ -75,19 +72,6 @@ lemma ι_of_mem_fieldOpIdealSet (x : CrAnAlgebra 𝓕) (hx : x ∈ 𝓕.fieldOpI refine RingConGen.Rel.of x 0 ?_ simpa using hx -lemma ι_superCommute_ofCrAnState_ofCrAnState_mem_center (φ ψ : 𝓕.CrAnStates) : - ι [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by - rw [Subalgebra.mem_center_iff] - intro b - obtain ⟨b, rfl⟩ := ι_surjective b - rw [← map_mul, ← map_mul] - rw [LinearMap.sub_mem_ker_iff.mp] - simp only [LinearMap.mem_ker] - apply ι_of_mem_fieldOpIdealSet - simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq] - left - use φ, ψ, b - lemma ι_superCommute_of_create_create (φc φc' : 𝓕.CrAnStates) (hφc : 𝓕 |>ᶜ φc = .create) (hφc' : 𝓕 |>ᶜ φc' = .create) : ι [ofCrAnState φc, ofCrAnState φc']ₛca = 0 := by apply ι_of_mem_fieldOpIdealSet @@ -117,5 +101,95 @@ lemma ι_superCommute_of_diff_statistic (φ ψ : 𝓕.CrAnStates) right use φ, ψ +lemma ι_superCommute_zero_of_fermionic (φ ψ : 𝓕.CrAnStates) + (h : [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ statisticSubmodule fermionic) : + ι [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] + apply ι_superCommute_of_diff_statistic _ _ + simpa using h + · simp [h] + +lemma ι_superCommute_ofCrAnState_ofCrAnState_bosonic_or_zero (φ ψ : 𝓕.CrAnStates) : + [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ statisticSubmodule bosonic ∨ + ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by + rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [ψ] with h | h + · simp_all [ofCrAnList_singleton] + · simp_all [ofCrAnList_singleton] + right + exact ι_superCommute_zero_of_fermionic _ _ h + +/-! + +## Super-commutes are in the center + +-/ + +@[simp] +lemma ι_superCommute_ofCrAnState_superCommute_ofCrAnState_ofCrAnState (φ1 φ2 φ3 : 𝓕.CrAnStates) : + ι [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca = 0 := by + apply ι_of_mem_fieldOpIdealSet + simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq] + left + use φ1, φ2, φ3 + +@[simp] +lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState (φ1 φ2 φ3 : 𝓕.CrAnStates) : + ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, ofCrAnState φ3]ₛca = 0 := by + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton] + rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h | h + · rw [bonsonic_superCommute_symm h] + simp [ofCrAnList_singleton] + · rcases ofCrAnList_bosonic_or_fermionic [φ3] with h' | h' + · rw [superCommute_bonsonic_symm h'] + simp [ofCrAnList_singleton] + · rw [superCommute_fermionic_fermionic_symm h h'] + simp [ofCrAnList_singleton] + +@[simp] +lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnList (φ1 φ2 : 𝓕.CrAnStates) + (φs : List 𝓕.CrAnStates) : + ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, ofCrAnList φs]ₛca = 0 := by + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] + rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h | h + · rw [superCommute_bosonic_ofCrAnList_eq_sum _ _ h] + simp [ofCrAnList_singleton] + · rw [superCommute_fermionic_ofCrAnList_eq_sum _ _ h] + simp [ofCrAnList_singleton] + +@[simp] +lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_crAnAlgebra (φ1 φ2 : 𝓕.CrAnStates) + (a : 𝓕.CrAnAlgebra) : ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, a]ₛca = 0 := by + change (ι.toLinearMap ∘ₗ superCommute [ofCrAnState φ1, ofCrAnState φ2]ₛca) a = _ + have h1 : (ι.toLinearMap ∘ₗ superCommute [ofCrAnState φ1, ofCrAnState φ2]ₛca) = 0 := by + apply (ofCrAnListBasis.ext fun l ↦ ?_) + simp + rw [h1] + simp + +lemma ι_commute_crAnAlgebra_superCommute_ofCrAnState_ofCrAnState (φ1 φ2 : 𝓕.CrAnStates) + (a : 𝓕.CrAnAlgebra) : ι a * ι [ofCrAnState φ1, ofCrAnState φ2]ₛca - + ι [ofCrAnState φ1, ofCrAnState φ2]ₛca * ι a = 0 := by + rcases ι_superCommute_ofCrAnState_ofCrAnState_bosonic_or_zero φ1 φ2 with h | h + swap + · simp [h] + trans - ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, a]ₛca + · rw [bosonic_superCommute h] + simp + · simp + +lemma ι_superCommute_ofCrAnState_ofCrAnState_mem_center (φ ψ : 𝓕.CrAnStates) : + ι [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by + rw [Subalgebra.mem_center_iff] + intro a + obtain ⟨a, rfl⟩ := ι_surjective a + have h0 := ι_commute_crAnAlgebra_superCommute_ofCrAnState_ofCrAnState φ ψ a + trans ι ((superCommute (ofCrAnState φ)) (ofCrAnState ψ)) * ι a + 0 + swap + simp + rw [← h0] + abel + end FieldOpAlgebra end FieldSpecification diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean index a6f0656a..a96e7d5e 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean @@ -3,6 +3,7 @@ Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ +import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.NormalOrder import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic /-! @@ -12,7 +13,6 @@ import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic namespace FieldSpecification open CrAnAlgebra -open ProtoOperatorAlgebra open HepLean.List open WickContraction open FieldStatistic diff --git a/HepLean/PerturbationTheory/FieldStatistics/Basic.lean b/HepLean/PerturbationTheory/FieldStatistics/Basic.lean index 2bd3d200..6faa8304 100644 --- a/HepLean/PerturbationTheory/FieldStatistics/Basic.lean +++ b/HepLean/PerturbationTheory/FieldStatistics/Basic.lean @@ -300,5 +300,8 @@ instance : AddMonoid FieldStatistic where simp only [instCommGroup, Finset.prod_const, Finset.card_univ, Fintype.card_fin] rfl +@[simp] +lemma add_eq_mul (a b : FieldStatistic) : a + b = a * b := rfl + end ofListTake end FieldStatistic diff --git a/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean b/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean index 115f0a9e..b83dd38c 100644 --- a/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean +++ b/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean @@ -71,6 +71,10 @@ lemma exchangeSign_bosonic (a : FieldStatistic) : 𝓢(a, bosonic) = 1 := by lemma bosonic_exchangeSign (a : FieldStatistic) : 𝓢(bosonic, a) = 1 := by rw [exchangeSign_symm, exchangeSign_bosonic] +@[simp] +lemma fermionic_exchangeSign_fermionic : 𝓢(fermionic, fermionic) = - 1 := by + rfl + lemma exchangeSign_eq_if (a b : FieldStatistic) : 𝓢(a, b) = if a = fermionic ∧ b = fermionic then - 1 else 1 := by fin_cases a <;> fin_cases b <;> rfl From a79d0f8fed26c4d3ba883ce07d6348e6756dc03c Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Tue, 28 Jan 2025 16:56:20 +0000 Subject: [PATCH 2/8] feat: KoszulSign partial sort --- HepLean/Mathematics/List/InsertIdx.lean | 7 + HepLean/Mathematics/List/InsertionSort.lean | 286 ++++++++++++++++++ .../Algebras/FieldOpAlgebra/NormalOrder.lean | 1 - .../FieldSpecification/TimeOrder.lean | 23 +- .../PerturbationTheory/Koszul/KoszulSign.lean | 92 ++++++ .../Koszul/KoszulSignInsert.lean | 13 + 6 files changed, 401 insertions(+), 21 deletions(-) diff --git a/HepLean/Mathematics/List/InsertIdx.lean b/HepLean/Mathematics/List/InsertIdx.lean index dfa843eb..99943e96 100644 --- a/HepLean/Mathematics/List/InsertIdx.lean +++ b/HepLean/Mathematics/List/InsertIdx.lean @@ -118,6 +118,13 @@ lemma insertIdx_eraseIdx_fin {I : Type} : List.insertIdx_succ_cons, List.cons.injEq, true_and] exact insertIdx_eraseIdx_fin as ⟨n, Nat.lt_of_succ_lt_succ h⟩ +lemma insertIdx_length_fst_append {I : Type} (φ : I) : (φs φs' : List I) → + List.insertIdx φs.length φ (φs ++ φs') = (φs ++ φ :: φs') + | [], φs' => by simp + | φ' :: φs, φs' => by + simp + exact insertIdx_length_fst_append φ φs φs' + lemma get_eq_insertIdx_succAbove {I : Type} (i : I) (r : List I) (k : Fin r.length.succ) : r.get = (List.insertIdx k i r).get ∘ (finCongr (insertIdx_length_fin i r k).symm) ∘ k.succAbove := by diff --git a/HepLean/Mathematics/List/InsertionSort.lean b/HepLean/Mathematics/List/InsertionSort.lean index 471e6ba2..97b840be 100644 --- a/HepLean/Mathematics/List/InsertionSort.lean +++ b/HepLean/Mathematics/List/InsertionSort.lean @@ -91,4 +91,290 @@ lemma insertionSortMin_lt_mem_insertionSortDropMinPos_of_lt {α : Type} (r : α simp only [hl, Nat.succ_eq_add_one, Fin.val_eq_val, ne_eq] exact Fin.succAbove_ne (insertionSortMinPosFin r a l) i +lemma insertionSort_insertionSort {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (l1 : List α): + List.insertionSort r (List.insertionSort r l1) = List.insertionSort r l1 := by + apply List.Sorted.insertionSort_eq + exact List.sorted_insertionSort r l1 + +lemma orderedInsert_commute {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) : (l : List α) → + List.orderedInsert r a (List.orderedInsert r b l) = List.orderedInsert r b (List.orderedInsert r a l) + | [] => by + have hrb : r b a := by + have ht := IsTotal.total (r := r) a b + simp_all + simp [hr, hrb] + | c :: l => by + have hrb : r b a := by + have ht := IsTotal.total (r := r) a b + simp_all + simp + by_cases h : r a c + · simp [h, hrb] + rw [if_pos] + simp [hrb, hr, h] + exact IsTrans.trans (r :=r) _ _ _ hrb h + · simp [h] + have hrca : r c a := by + have ht := IsTotal.total (r := r) a c + simp_all + by_cases hbc : r b c + · simp [hbc, hr, h] + · simp [hbc, h] + exact orderedInsert_commute r a b hr l + +lemma insertionSort_orderedInsert_append {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) : (l1 l2 : List α) → + List.insertionSort r (List.orderedInsert r a l1 ++ l2) = List.insertionSort r (a :: l1 ++ l2) + | [], l2 => by + simp + | b :: l1, l2 => by + conv_lhs => simp + by_cases h : r a b + · simp [h] + conv_lhs => simp [h] + rw [insertionSort_orderedInsert_append r a l1 l2] + simp + rw [orderedInsert_commute r a b h] + + +lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] : (l1 l2 : List α) → + List.insertionSort r (List.insertionSort r l1 ++ l2) = List.insertionSort r (l1 ++ l2) + | [], l2 => by + simp + | a :: l1, l2 => by + conv_lhs => simp + rw [insertionSort_orderedInsert_append] + simp + rw [insertionSort_insertionSort_append r l1 l2] + + +@[simp] +lemma orderedInsert_length {α : Type} (r : α → α → Prop) [DecidableRel r] (a : α) (l : List α) : + (List.orderedInsert r a l).length = (a :: l).length := by + apply List.Perm.length_eq + exact List.perm_orderedInsert r a l + +lemma takeWhile_orderedInsert {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] + (a b : α) (hr : ¬ r a b) : (l : List α) → + (List.takeWhile (fun c => !decide (r a c)) (List.orderedInsert r b l)).length = + (List.takeWhile (fun c => !decide (r a c)) l).length + 1 + | [] => by + simp [hr] + | c :: l => by + simp + by_cases h : r b c + · simp [h] + rw [List.takeWhile_cons_of_pos] + simp + simp [hr] + · simp [h] + have hrba : r b a:= by + have ht := IsTotal.total (r := r) a b + simp_all + have hl : ¬ r a c := by + by_contra hn + apply h + exact IsTrans.trans _ _ _ hrba hn + simp [hl] + exact takeWhile_orderedInsert r a b hr l + +lemma takeWhile_orderedInsert' {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] + (a b : α) (hr : ¬ r a b) : (l : List α) → + (List.takeWhile (fun c => !decide (r b c)) (List.orderedInsert r a l)).length = + (List.takeWhile (fun c => !decide (r b c)) l).length + | [] => by + simp + have ht := IsTotal.total (r := r) a b + simp_all + | c :: l => by + have hrba : r b a:= by + have ht := IsTotal.total (r := r) a b + simp_all + simp + by_cases h : r b c + · simp [h, hrba] + by_cases hac : r a c + · simp [hac, hrba] + · simp [hac, h] + · have hcb : r c b := by + have ht := IsTotal.total (r := r) b c + simp_all + by_cases hac : r a c + · refine False.elim (h ?_) + exact IsTrans.trans _ _ _ hrba hac + · simp [hac, h] + exact takeWhile_orderedInsert' r a b hr l + + + + + + + +lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) (n : ℕ) : (l : List α) → + (hn : n + 2 < (a :: b :: l).length) → + insertionSortEquiv r (a :: b :: l) ⟨n + 2, hn⟩ = (finCongr (by simp)) + (insertionSortEquiv r (b :: a :: l) ⟨n + 2, hn⟩):= by + have hrba : r b a:= by + have ht := IsTotal.total (r := r) a b + simp_all + intro l hn + simp [insertionSortEquiv] + conv_lhs => erw [equivCons_succ] + conv_rhs => erw [equivCons_succ] + simp + conv_lhs => + rhs + rhs + erw [orderedInsertEquiv_succ] + conv_lhs => erw [orderedInsertEquiv_fin_succ] + simp + conv_rhs => + rhs + rhs + erw [orderedInsertEquiv_succ] + conv_rhs => erw [orderedInsertEquiv_fin_succ] + ext + simp + let a1 : Fin ((List.orderedInsert r b (List.insertionSort r l)).length + 1) := ⟨↑(orderedInsertPos r (List.orderedInsert r b (List.insertionSort r l)) a), orderedInsertPos_lt_length r (List.orderedInsert r b (List.insertionSort r l)) a⟩ + let b1 : Fin ((List.insertionSort r l).length + 1) := ⟨↑(orderedInsertPos r (List.insertionSort r l) b), orderedInsertPos_lt_length r (List.insertionSort r l) b⟩ + let a2 : Fin ((List.insertionSort r l).length + 1) := ⟨↑(orderedInsertPos r (List.insertionSort r l) a), orderedInsertPos_lt_length r (List.insertionSort r l) a⟩ + let b2 : Fin ((List.orderedInsert r a (List.insertionSort r l)).length + 1) := ⟨↑(orderedInsertPos r (List.orderedInsert r a (List.insertionSort r l)) b), orderedInsertPos_lt_length r (List.orderedInsert r a (List.insertionSort r l)) b⟩ + have ht : (List.takeWhile (fun c => !decide (r b c)) (List.insertionSort r l)) + = (List.takeWhile (fun c => !decide (r b c)) ((List.takeWhile (fun c => !decide (r a c)) (List.insertionSort r l)))) := by + rw [List.takeWhile_takeWhile] + simp + congr + funext c + simp + intro hbc hac + refine hbc ?_ + exact IsTrans.trans _ _ _ hrba hac + have ha1 : b1.1 ≤ a2.1 := by + simp [a1, a2, orderedInsertPos] + rw [ht] + apply List.Sublist.length_le + exact List.takeWhile_sublist fun c => !decide (r b c) + have ha2 : a1.1 = a2.1 + 1 := by + simp [a1, a2, orderedInsertPos] + rw [takeWhile_orderedInsert] + exact hr + have hb : b1.1 = b2.1 := by + simp [b1, b2, orderedInsertPos] + rw [takeWhile_orderedInsert'] + exact hr + let n := ((insertionSortEquiv r l) ⟨n, by simpa using hn⟩) + change (a1.succAbove ⟨b1.succAbove n, _⟩).1 = (b2.succAbove ⟨a2.succAbove n, _⟩).1 + trans if (b1.succAbove n).1 < a1.1 then (b1.succAbove n).1 else (b1.succAbove n).1 + 1 + · rw [Fin.succAbove] + simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk] + by_cases ha : (b1.succAbove n).1 < a1.1 + · simp [ha] + · simp [ha] + trans if (a2.succAbove n).1 < b2.1 then (a2.succAbove n).1 else (a2.succAbove n).1 + 1 + swap + · conv_rhs => rw [Fin.succAbove] + simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk] + by_cases ha : (a2.succAbove n).1 < b2.1 + · simp [ha] + · simp [ha] + have hbs1 : (b1.succAbove n).1 = if n.1 < b1.1 then n.1 else n.1 + 1 := by + rw [Fin.succAbove] + simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk] + by_cases ha : n.1 < b1.1 + · simp [ha] + · simp [ha] + have has2 : (a2.succAbove n).1 = if n.1 < a2.1 then n.1 else n.1 + 1 := by + rw [Fin.succAbove] + simp only [Fin.castSucc_mk, Fin.lt_def, Fin.succ_mk] + by_cases ha : n.1 < a2.1 + · simp [ha] + · simp [ha] + rw [hbs1, has2, hb, ha2] + have hnat (a2 b2 n : ℕ) (h : b2 ≤ a2) : (if (if ↑n < ↑b2 then ↑n else ↑n + 1) < ↑a2 + 1 then if ↑n < ↑b2 then ↑n else ↑n + 1 + else (if ↑n < ↑b2 then ↑n else ↑n + 1) + 1) = + if (if ↑n < ↑a2 then ↑n else ↑n + 1) < ↑b2 then if ↑n < ↑a2 then ↑n else ↑n + 1 + else (if ↑n < ↑a2 then ↑n else ↑n + 1) + 1 := by + by_cases hnb2 : n < b2 + · simp [hnb2] + have h1 : n < a2 + 1 := by omega + have h2 : n < a2 := by omega + simp [h1, h2, hnb2] + · simp [hnb2] + by_cases ha2 : n < a2 + · simp [ha2, hnb2] + · simp [ha2] + rw [if_neg] + omega + apply hnat + rw [← hb] + exact ha1 + + + +lemma insertionSortEquiv_orderedInsert_append {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a a2 : α) : (l1 l2 : List α) → + (insertionSortEquiv r (List.orderedInsert r a l1 ++ a2 :: l2) ⟨l1.length + 1, by + simp⟩) + = (finCongr (by simp; omega)) + ((insertionSortEquiv r ( a :: l1 ++ a2 :: l2)) ⟨l1.length + 1, by simp⟩) + | [], l2 => by + simp + | b :: l1, l2 => by + by_cases h : r a b + · have h1 : (List.orderedInsert r a (b :: l1) ++ a2 :: l2) = (a :: b :: l1 ++ a2 :: l2) := by + simp [h] + rw [insertionSortEquiv_congr _ _ h1] + simp + · have h1 : (List.orderedInsert r a (b :: l1) ++ a2 :: l2) = (b :: List.orderedInsert r a (l1) ++ a2 :: l2) := by + simp [h] + rw [insertionSortEquiv_congr _ _ h1] + simp + conv_lhs => simp [insertionSortEquiv] + rw [insertionSortEquiv_orderedInsert_append r a] + have hl : (List.insertionSort r (List.orderedInsert r a l1 ++ a2 :: l2)) = + List.insertionSort r (a :: l1 ++ a2 :: l2) := by + exact insertionSort_orderedInsert_append r a l1 (a2 :: l2) + rw [orderedInsertEquiv_congr _ _ _ hl] + simp + change Fin.cast _ ((insertionSortEquiv r (b :: a :: (l1 ++ a2 :: l2))) ⟨l1.length + 2, by simp⟩) = _ + have hl : l1.length + 1 +1 = l1.length + 2 := by omega + simp [hl] + conv_rhs => + erw [insertionSortEquiv_commute _ _ _ h _ _] + simp + + + + +lemma insertionSortEquiv_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) : (l1 l2 : List α) → + (insertionSortEquiv r (List.insertionSort r l1 ++ a :: l2) ⟨l1.length, by simp⟩) + = finCongr (by simp) (insertionSortEquiv r (l1 ++ a :: l2) ⟨l1.length, by simp⟩) + | [], l2 => by + simp + | b :: l1, l2 => by + simp + have hl := insertionSortEquiv_orderedInsert_append r b a (List.insertionSort r l1) l2 + simp at hl + rw [hl] + have ih := insertionSortEquiv_insertionSort_append r a l1 l2 + simp [insertionSortEquiv] + rw [ih] + have hl : (List.insertionSort r (List.insertionSort r l1 ++ a :: l2)) = (List.insertionSort r (l1 ++ a :: l2)) := by + exact insertionSort_insertionSort_append r l1 (a :: l2) + rw [orderedInsertEquiv_congr _ _ _ hl] + simp + + + + + + end HepLean.List diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean index a96e7d5e..89862492 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean @@ -14,7 +14,6 @@ import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic namespace FieldSpecification open CrAnAlgebra open HepLean.List -open WickContraction open FieldStatistic namespace FieldOpAlgebra diff --git a/HepLean/PerturbationTheory/FieldSpecification/TimeOrder.lean b/HepLean/PerturbationTheory/FieldSpecification/TimeOrder.lean index 51f17544..b9b757b4 100644 --- a/HepLean/PerturbationTheory/FieldSpecification/TimeOrder.lean +++ b/HepLean/PerturbationTheory/FieldSpecification/TimeOrder.lean @@ -234,27 +234,10 @@ lemma crAnTimeOrderSign_pair_not_ordered {φ ψ : 𝓕.CrAnStates} (h : ¬ crAnT rw [if_neg h] simp [FieldStatistic.exchangeSign_eq_if] -lemma crAnTimeOrderSign_swap_eq_time_cons {φ ψ : 𝓕.CrAnStates} - (h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) (φs' : List 𝓕.CrAnStates) : - crAnTimeOrderSign (φ :: ψ :: φs') = crAnTimeOrderSign (ψ :: φ :: φs') := by - simp only [crAnTimeOrderSign, Wick.koszulSign, ← mul_assoc, mul_eq_mul_right_iff] - left - rw [mul_comm] - simp [Wick.koszulSignInsert, h1, h2] - lemma crAnTimeOrderSign_swap_eq_time {φ ψ : 𝓕.CrAnStates} - (h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) : (φs φs' : List 𝓕.CrAnStates) → - crAnTimeOrderSign (φs ++ φ :: ψ :: φs') = crAnTimeOrderSign (φs ++ ψ :: φ :: φs') - | [], φs' => by - simp only [crAnTimeOrderSign, List.nil_append] - exact crAnTimeOrderSign_swap_eq_time_cons h1 h2 φs' - | φ'' :: φs, φs' => by - simp only [crAnTimeOrderSign, Wick.koszulSign, List.append_eq] - rw [← crAnTimeOrderSign, ← crAnTimeOrderSign] - rw [crAnTimeOrderSign_swap_eq_time h1 h2] - congr 1 - apply Wick.koszulSignInsert_eq_perm - exact List.Perm.append_left φs (List.Perm.swap ψ φ φs') + (h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) (φs φs' : List 𝓕.CrAnStates) : + crAnTimeOrderSign (φs ++ φ :: ψ :: φs') = crAnTimeOrderSign (φs ++ ψ :: φ :: φs') := by + exact Wick.koszulSign_swap_eq_rel _ _ h1 h2 _ _ /-- Sort a list of `CrAnStates` based on `crAnTimeOrderRel`. -/ def crAnTimeOrderList (φs : List 𝓕.CrAnStates) : List 𝓕.CrAnStates := diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean index a7271602..0763d3ee 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Tooby-Smith -/ import HepLean.PerturbationTheory.Koszul.KoszulSignInsert +import HepLean.Mathematics.List.InsertionSort /-! # Koszul sign @@ -259,4 +260,95 @@ lemma koszulSign_eraseIdx_insertionSortMinPos [IsTotal 𝓕 le] [IsTrans 𝓕 le apply Or.inl rfl +lemma koszulSign_swap_eq_rel_cons {ψ φ : 𝓕} + (h1 : le φ ψ) (h2 : le ψ φ) (φs' : List 𝓕): + koszulSign q le (φ :: ψ :: φs') = koszulSign q le (ψ :: φ :: φs') := by + simp only [Wick.koszulSign, ← mul_assoc, mul_eq_mul_right_iff] + left + rw [mul_comm] + simp [Wick.koszulSignInsert, h1, h2] + +lemma koszulSign_swap_eq_rel {ψ φ : 𝓕} (h1 : le φ ψ) (h2 : le ψ φ) : (φs φs' : List 𝓕) → + koszulSign q le (φs ++ φ :: ψ :: φs') = koszulSign q le (φs ++ ψ :: φ :: φs') + | [], φs' => by + simp only [List.nil_append] + exact koszulSign_swap_eq_rel_cons q le h1 h2 φs' + | φ'' :: φs, φs' => by + simp only [Wick.koszulSign, List.append_eq] + rw [koszulSign_swap_eq_rel h1 h2] + congr 1 + apply Wick.koszulSignInsert_eq_perm + exact List.Perm.append_left φs (List.Perm.swap ψ φ φs') + +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 + rw [koszulSign_of_sorted φs h.2] + simp + exact koszulSignInsert_of_le_mem _ _ _ _ h.1 + +@[simp] +lemma koszulSign_of_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φs : List 𝓕) : + koszulSign q le (List.insertionSort le φs) = 1 := by + apply koszulSign_of_sorted + exact List.sorted_insertionSort le φs + +lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕 le] : (φs φs' : List 𝓕) → + koszulSign q le (φs ++ φs') = + koszulSign q le (List.insertionSort le φs ++ φs') * koszulSign q le φs + | φs, [] => by + simp + | φs, φ :: φs' => by + have h1 : (φs ++ φ :: φs') = List.insertIdx φs.length φ (φs ++ φs') := by + rw [insertIdx_length_fst_append] + have h2 : (List.insertionSort le φs ++ φ :: φs') = List.insertIdx (List.insertionSort le φs).length φ (List.insertionSort le φs ++ φs') := by + rw [insertIdx_length_fst_append] + rw [h1, h2] + rw [koszulSign_insertIdx] + simp + rw [koszulSign_insertIdx] + simp [mul_assoc] + left + rw [koszulSign_of_append_eq_insertionSort_left φs φs'] + simp [mul_assoc] + left + simp [mul_comm] + left + congr 3 + · have h2 : (List.insertionSort le φs ++ φ :: φs') = List.insertIdx φs.length φ (List.insertionSort le φs ++ φs') := by + rw [← insertIdx_length_fst_append] + simp + rw [insertionSortEquiv_congr _ _ h2.symm] + simp + rw [insertionSortEquiv_insertionSort_append] + simp + rw [insertionSortEquiv_congr _ _ h1.symm] + simp + · rw [insertIdx_length_fst_append] + rw [show φs.length = (List.insertionSort le φs).length by simp] + rw [insertIdx_length_fst_append] + symm + apply insertionSort_insertionSort_append + · simp + · simp + +lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] : (φs'' φs φs' : List 𝓕) → + koszulSign q le (φs'' ++ φs ++ φs') = + koszulSign q le (φs'' ++ List.insertionSort le φs ++ φs') * koszulSign q le φs + | [], φs, φs'=> by + simp + exact koszulSign_of_append_eq_insertionSort_left q le φs φs' + | φ'' :: φs'', φs, φs' => by + simp only [koszulSign, List.append_eq] + rw [koszulSign_of_append_eq_insertionSort φs'' φs φs', ← mul_assoc] + congr 2 + apply koszulSignInsert_eq_perm + refine (List.perm_append_right_iff φs').mpr ?_ + refine List.Perm.append_left φs'' ?_ + exact List.Perm.symm (List.perm_insertionSort le φs) + end Wick diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean index 0ce28d02..7eb5ee15 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean @@ -235,4 +235,17 @@ lemma koszulSignInsert_cons (r0 r1 : 𝓕) (r : List 𝓕) : koszulSignInsert q le r0 r := by simp [koszulSignInsert, koszulSignCons] +lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ b ∈ φs, le φ0 b) → + koszulSignInsert q le φ0 φs = 1 + | [], _ => by + simp [koszulSignInsert] + | φ1 :: φs, h => by + simp [koszulSignInsert] + rw [if_pos] + · apply koszulSignInsert_of_le_mem + · intro b hb + exact h b (List.mem_cons_of_mem _ hb) + · exact h φ1 (List.mem_cons_self _ _) + + end Wick From c2d89cc093bf5dd434a6370089300f103d249974 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Wed, 29 Jan 2025 12:09:02 +0000 Subject: [PATCH 3/8] feat: Property of time-order w.r.t. superCommute --- HepLean/Mathematics/List/InsertionSort.lean | 278 +++++++++++++++++- .../Algebras/CrAnAlgebra/SuperCommute.lean | 34 +++ .../Algebras/CrAnAlgebra/TimeOrder.lean | 173 +++++++++++ .../Algebras/FieldOpAlgebra/Basic.lean | 4 +- .../Algebras/FieldOpAlgebra/TimeOrder.lean | 175 +++++++++++ .../FieldStatistics/Basic.lean | 4 + .../PerturbationTheory/Koszul/KoszulSign.lean | 83 ++++++ .../Koszul/KoszulSignInsert.lean | 38 +++ 8 files changed, 781 insertions(+), 8 deletions(-) create mode 100644 HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean diff --git a/HepLean/Mathematics/List/InsertionSort.lean b/HepLean/Mathematics/List/InsertionSort.lean index 97b840be..2a914d88 100644 --- a/HepLean/Mathematics/List/InsertionSort.lean +++ b/HepLean/Mathematics/List/InsertionSort.lean @@ -150,6 +150,15 @@ lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [D simp rw [insertionSort_insertionSort_append r l1 l2] +lemma insertionSort_append_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] : (l1 l2 l3 : List α) → + List.insertionSort r (l1 ++ List.insertionSort r l2 ++ l3) = List.insertionSort r (l1 ++ l2 ++ l3) + | [], l2, l3 => by + simp + exact insertionSort_insertionSort_append r l2 l3 + | a :: l1, l2, l3 => by + simp only [List.insertionSort, List.append_eq] + rw [insertionSort_append_insertionSort_append r l1 l2 l3] @[simp] lemma orderedInsert_length {α : Type} (r : α → α → Prop) [DecidableRel r] (a : α) (l : List α) : @@ -210,12 +219,6 @@ lemma takeWhile_orderedInsert' {α : Type} (r : α → α → Prop) [DecidableRe · simp [hac, h] exact takeWhile_orderedInsert' r a b hr l - - - - - - lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) (n : ℕ) : (l : List α) → (hn : n + 2 < (a :: b :: l).length) → @@ -373,8 +376,271 @@ lemma insertionSortEquiv_insertionSort_append {α : Type} (r : α → α → Pro simp +/-! + +## Insertion sort with equal fields + +-/ +lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] + [IsTrans α r] (a : α) (p : α → Prop) [DecidablePred p] (h : p a) : (l : List α) → + (hl : l.Sorted r) → + List.filter p (List.orderedInsert r a l) = List.orderedInsert r a (List.filter p l) + | [], hl => by + simp + exact h + | b :: l, hl => by + simp + by_cases hpb : p b <;> by_cases hab : r a b + · simp [hpb, hab] + rw [List.filter_cons_of_pos (by simp [h])] + rw [List.filter_cons_of_pos (by simp [hpb])] + · simp [hab] + rw [List.filter_cons_of_pos (by simp [hpb])] + rw [List.filter_cons_of_pos (by simp [hpb])] + simp [hab] + simp at hl + exact orderedInsert_filter_of_pos r a p h l hl.2 + · simp [hab] + rw [List.filter_cons_of_pos (by simp [h]), + List.filter_cons_of_neg (by simp [hpb])] + rw [List.orderedInsert_eq_take_drop] + have hl : List.takeWhile (fun b => decide ¬r a b) (List.filter (fun b => decide (p b)) l) = [] := by + rw [List.takeWhile_eq_nil_iff] + intro c hc + simp at hc + apply hc + apply IsTrans.trans a b _ hab + simp at hl + apply hl.1 + have hlf : (List.filter (fun b => decide (p b)) l)[0] ∈ (List.filter (fun b => decide (p b)) l) := by + exact List.getElem_mem c + simp [- List.getElem_mem] at hlf + exact hlf.1 + rw [hl] + simp + conv_lhs => rw [← List.takeWhile_append_dropWhile (fun b => decide ¬r a b) (List.filter (fun b => decide (p b)) l )] + rw [hl] + simp + · simp [hab] + rw [List.filter_cons_of_neg (by simp [hpb])] + rw [List.filter_cons_of_neg (by simp [hpb])] + simp at hl + exact orderedInsert_filter_of_pos r a p h l hl.2 + +lemma orderedInsert_filter_of_neg {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] + [IsTrans α r] (a : α) (p : α → Prop) [DecidablePred p] (h : ¬ p a) (l : List α) : + List.filter p (List.orderedInsert r a l) = (List.filter p l) := by + rw [List.orderedInsert_eq_take_drop] + simp + rw [List.filter_cons_of_neg] + rw [← List.filter_append] + congr + exact List.takeWhile_append_dropWhile (fun b => !decide (r a b)) l + simp [h] + + + +lemma insertionSort_filter {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] + [IsTrans α r] (p : α → Prop) [DecidablePred p] : (l : List α) → + List.insertionSort r (List.filter p l) = + List.filter p (List.insertionSort r l) + | [] => by simp + | a :: l => by + simp + by_cases h : p a + · rw [orderedInsert_filter_of_pos] + rw [List.filter_cons_of_pos] + simp + rw [insertionSort_filter] + simp_all + simp_all + exact List.sorted_insertionSort r l + · rw [orderedInsert_filter_of_neg] + rw [List.filter_cons_of_neg] + rw [insertionSort_filter] + simp_all + exact h +lemma takeWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → + List.takeWhile (fun c => ¬ r a c) l = List.filter (fun c => ¬ r a c) l + | [], _ => by simp + | b :: l, hl => by + simp at hl + by_cases hb : ¬ r a b + · simp [hb] + simpa using takeWhile_sorted_eq_filter r a l hl.2 + · simp_all + intro c hc + apply IsTrans.trans a b c hb + exact hl.1 c hc + +lemma dropWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → + List.dropWhile (fun c => ¬ r a c) l = List.filter (fun c => r a c) l + | [], _ => by simp + | b :: l, hl => by + simp at hl + by_cases hb : ¬ r a b + · simp [hb] + simpa using dropWhile_sorted_eq_filter r a l hl.2 + · simp_all + symm + rw [List.filter_eq_self] + intro c hc + simp + apply IsTrans.trans a b c hb + exact hl.1 c hc + +lemma dropWhile_sorted_eq_filter_filter {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) :(l : List α) → (hl : l.Sorted r) → + List.filter (fun c => r a c) l = + List.filter (fun c => r a c ∧ r c a) l ++ List.filter (fun c => r a c ∧ ¬ r c a) l + | [], _ => by + simp + | b :: l, hl => by + simp at hl + by_cases hb : ¬ r a b + · simp [hb] + simpa using dropWhile_sorted_eq_filter_filter r a l hl.2 + · simp_all + by_cases hba : r b a + · simp [hba] + rw [List.filter_cons_of_pos] + rw [dropWhile_sorted_eq_filter_filter] + simp + exact hl.2 + simp_all + · simp[hba] + have h1 : List.filter (fun c => decide (r a c) && decide (r c a)) l = [] := by + rw [@List.filter_eq_nil_iff] + intro c hc + simp + intro hac hca + apply hba + apply IsTrans.trans b c a _ hca + exact hl.1 c hc + rw [h1] + rw [dropWhile_sorted_eq_filter_filter] + simp [h1] + rw [List.filter_cons_of_pos] + simp_all + exact hl.2 + +lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) :(l : List α) → + List.filter (fun c => r a c ∧ r c a) (l.insertionSort r) = + List.filter (fun c => r a c ∧ r c a) l + | [] => by simp + | b :: l => by + simp only [ List.insertionSort] + by_cases h : r a b ∧ r b a + · have hl := orderedInsert_filter_of_pos r b (fun c => r a c ∧ r c a) h (List.insertionSort r l) + (by exact List.sorted_insertionSort r l) + simp at hl ⊢ + rw [hl] + rw [List.orderedInsert_eq_take_drop] + have ht : List.takeWhile (fun b_1 => decide ¬r b b_1) + (List.filter (fun b => decide (r a b) && decide (r b a)) (List.insertionSort r l)) = [] := by + rw [List.takeWhile_eq_nil_iff] + intro hl + simp + have hx := List.getElem_mem hl + simp [- List.getElem_mem] at hx + apply IsTrans.trans b a _ h.2 + simp_all + rw [ht] + simp + rw [List.filter_cons_of_pos] + simp + have ih := filter_rel_eq_insertionSort r a l + simp at ih + rw [← ih] + have htd := List.takeWhile_append_dropWhile (fun b_1 => decide ¬r b b_1) (List.filter (fun b => decide (r a b) && decide (r b a)) (List.insertionSort r l)) + simp [decide_not, - List.takeWhile_append_dropWhile] at htd + conv_rhs => rw [← htd] + simp [- List.takeWhile_append_dropWhile] + intro hl + have hx := List.getElem_mem hl + simp [- List.getElem_mem] at hx + apply IsTrans.trans b a _ h.2 + simp_all + simp_all + · have hl := orderedInsert_filter_of_neg r b (fun c => r a c ∧ r c a) h (List.insertionSort r l) + simp at hl ⊢ + rw [hl] + rw [List.filter_cons_of_neg] + have ih := filter_rel_eq_insertionSort r a l + simp_all + simpa using h + + +lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) (l1 l l2 : List α) + (h : ∀ b ∈ l, r a b ∧ r b a) : + List.insertionSort r (l1 ++ l ++ l2) = + (List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r)) + ++ (List.filter (fun c => r a c ∧ r c a) l1) + ++ l + ++ (List.filter (fun c => r a c ∧ r c a) l2) + ++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) + := by + have hl : List.insertionSort r (l1 ++ l ++ l2) = + List.takeWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) ++ + List.dropWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) := by + exact (List.takeWhile_append_dropWhile (fun c => decide ¬r a c) + (List.insertionSort r (l1 ++ l ++ l2))).symm + have hlt : List.takeWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) + = List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r) := by + rw [takeWhile_sorted_eq_filter, takeWhile_sorted_eq_filter ] + rw [← insertionSort_filter, ← insertionSort_filter] + congr 1 + simp + exact fun b hb => (h b hb).1 + exact List.sorted_insertionSort r (l1 ++ l2) + exact List.sorted_insertionSort r (l1 ++ l ++ l2) + conv_lhs => rw [hl, hlt] + simp only [decide_not, Bool.decide_and] + simp + have h1 := dropWhile_sorted_eq_filter r a (List.insertionSort r (l1 ++ (l ++ l2))) + simp at h1 + rw [h1] + rw [dropWhile_sorted_eq_filter_filter, filter_rel_eq_insertionSort] + simp + congr 1 + simp + exact fun a a_1 => h a a_1 + congr 1 + have h1 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ (l ++ l2)) + simp at h1 + rw [← h1] + have h2 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ l2) + simp at h2 + rw [← h2] + congr + have hl : List.filter (fun b => decide (r a b) && !decide (r b a)) l = [] := by + rw [@List.filter_eq_nil_iff] + intro c hc + simp_all + rw [hl] + simp + exact List.sorted_insertionSort r (l1 ++ (l ++ l2)) + exact List.sorted_insertionSort r (l1 ++ (l ++ l2)) + +lemma insertionSort_of_takeWhile_filter {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a : α) (l1 l2 : List α) : + List.insertionSort r (l1 ++ l2) = + (List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r)) + ++ (List.filter (fun c => r a c ∧ r c a) l1) + ++ (List.filter (fun c => r a c ∧ r c a) l2) + ++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) + := by + trans List.insertionSort r (l1 ++ [] ++ l2) + simp + rw [insertionSort_of_eq_list r a l1 [] l2] + simp + simp end HepLean.List diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean index 59a66053..ce9e4a61 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean @@ -439,6 +439,40 @@ lemma superCommute_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) : · simp · simp [Finset.mul_sum, smul_smul, ofStateList_cons, mul_assoc, FieldStatistic.ofList_cons_eq_mul, mul_comm] + +lemma summerCommute_jacobi_ofCrAnList (φs1 φs2 φs3 : List 𝓕.CrAnStates) : + [ofCrAnList φs1, [ofCrAnList φs2, ofCrAnList φs3]ₛca]ₛca = + 𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs3) • + (- 𝓢(𝓕 |>ₛ φ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 + 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] + abel + · simp [h1, h2, exchangeSign_bosonic, bosonic_exchangeSign, mul_one, one_smul] + abel + · simp [h1, bosonic_exchangeSign, h3, exchangeSign_bosonic, mul_one, one_smul] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel /-! ## Interaction with grading. diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean index 16859ba7..29f12d30 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean @@ -39,6 +39,66 @@ lemma timeOrder_ofCrAnList (φs : List 𝓕.CrAnStates) : rw [← ofListBasis_eq_ofList] simp only [timeOrder, Basis.constr_basis] +lemma timeOrder_timeOrder_mid (a b c : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b * c) = 𝓣ᶠ(a * 𝓣ᶠ(b) * c) := by + let pc (c : 𝓕.CrAnAlgebra) (hc : c ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) : + Prop := 𝓣ᶠ(a * b * c) = 𝓣ᶠ(a * 𝓣ᶠ(b) * c) + change pc c (Basis.mem_span _ c) + apply Submodule.span_induction + · intro x hx + obtain ⟨φs, rfl⟩ := hx + simp [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] + 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] + rw [timeOrder_ofCrAnList] + simp only [← ofCrAnList_append, Algebra.mul_smul_comm, + Algebra.smul_mul_assoc, map_smul] + rw [timeOrder_ofCrAnList, timeOrder_ofCrAnList, smul_smul] + congr 1 + · simp only [crAnTimeOrderSign, crAnTimeOrderList] + rw [Wick.koszulSign_of_append_eq_insertionSort, mul_comm] + · congr 1 + simp only [crAnTimeOrderList] + rw [insertionSort_append_insertionSort_append] + · simp [pa] + · intro x y hx hy h1 h2 + simp_all [pa, add_mul] + · intro x hx h + simp_all [pa] + · simp [pb] + · intro x y hx hy h1 h2 + simp_all [pb, mul_add, add_mul] + · intro x hx h + simp_all [pb] + · simp [pc] + · intro x y hx hy h1 h2 + simp_all [pc, mul_add] + · intro x hx h hp + simp_all [pc] + +lemma timeOrder_timeOrder_right (a b : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b) = 𝓣ᶠ(a * 𝓣ᶠ(b)) := by + trans 𝓣ᶠ(a * b * 1) + · simp + · rw [timeOrder_timeOrder_mid] + simp + +lemma timeOrder_timeOrder_left (a b : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b) = 𝓣ᶠ(𝓣ᶠ(a) * b) := by + trans 𝓣ᶠ(1 * a * b) + · simp + · rw [timeOrder_timeOrder_mid] + simp + lemma timeOrder_ofStateList (φs : List 𝓕.States) : 𝓣ᶠ(ofStateList φs) = timeOrderSign φs • ofStateList (timeOrderList φs) := by conv_lhs => @@ -100,6 +160,119 @@ lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel · rw [crAnTimeOrderList_pair_ordered] simp_all +lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right + {φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a : 𝓕.CrAnAlgebra) : + 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca) = 0 := by + rw [timeOrder_timeOrder_right, + timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h] + simp + + +lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left + {φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a : 𝓕.CrAnAlgebra) : + 𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca * a) = 0 := by + rw [timeOrder_timeOrder_left, + timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h] + simp + +lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_mid + {φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a b : 𝓕.CrAnAlgebra) : + 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = 0 := by + rw [timeOrder_timeOrder_mid, + timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h] + simp + +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 + rw [bosonic_superCommute (Submodule.coe_mem (bosonicProj a))] + simp + rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h] + rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h] + simp + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] + rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h' | h' + · rw [superCommute_bonsonic h'] + simp [ofCrAnList_singleton] + 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] + rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left h] + rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right h] + simp + +lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel + {φ1 φ2 φ3 : 𝓕.CrAnStates} (h12 : ¬ crAnTimeOrderRel φ1 φ2) + (h13 : ¬ crAnTimeOrderRel φ1 φ3) : + 𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton] + rw [summerCommute_jacobi_ofCrAnList] + simp [ofCrAnList_singleton] + right + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12 ] + rw [superCommute_ofCrAnState_ofCrAnState_symm φ3] + simp + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13] + simp + +lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel' + {φ1 φ2 φ3 : 𝓕.CrAnStates} (h12 : ¬ crAnTimeOrderRel φ2 φ1) + (h13 : ¬ crAnTimeOrderRel φ3 φ1) : + 𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton] + rw [summerCommute_jacobi_ofCrAnList] + simp [ofCrAnList_singleton] + right + rw [superCommute_ofCrAnState_ofCrAnState_symm φ1] + simp + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12 ] + simp + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13] + simp + +lemma timeOrder_superCommute_ofCrAnState_superCommute_all_not_crAnTimeOrderRel + (φ1 φ2 φ3 : 𝓕.CrAnStates) (h : ¬ ( + crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧ + 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 + by_cases h23 : ¬ crAnTimeOrderRel φ2 φ3 + · simp_all + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h23] + simp_all + by_cases h32 : ¬ crAnTimeOrderRel φ3 φ2 + · simp_all + rw [superCommute_ofCrAnState_ofCrAnState_symm] + simp + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h32] + simp + simp_all + 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 + have h13 : crAnTimeOrderRel φ1 φ3 := IsTrans.trans φ1 φ2 φ3 h12 h23 + simp_all + by_cases h21 : ¬ crAnTimeOrderRel φ2 φ1 + · simp_all + 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 + refine False.elim (h ?_) + exact IsTrans.trans φ3 φ2 φ1 h32 h21 + + lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_eq_time {φ ψ : 𝓕.CrAnStates} (h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) : 𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca) = [ofCrAnState φ, ofCrAnState ψ]ₛca := by diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean index 9fba2622..aa71461c 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean @@ -92,7 +92,7 @@ lemma ι_superCommute_of_annihilate_annihilate (φa φa' : 𝓕.CrAnStates) left use φa, φa', hφa, hφa' -lemma ι_superCommute_of_diff_statistic (φ ψ : 𝓕.CrAnStates) +lemma ι_superCommute_of_diff_statistic {φ ψ : 𝓕.CrAnStates} (h : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ)) : ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by apply ι_of_mem_fieldOpIdealSet simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq] @@ -107,7 +107,7 @@ lemma ι_superCommute_zero_of_fermionic (φ ψ : 𝓕.CrAnStates) rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] at h ⊢ rcases statistic_neq_of_superCommute_fermionic h with h | h · simp [ofCrAnList_singleton] - apply ι_superCommute_of_diff_statistic _ _ + apply ι_superCommute_of_diff_statistic simpa using h · simp [h] diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean new file mode 100644 index 00000000..6db44c8d --- /dev/null +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean @@ -0,0 +1,175 @@ +/- +Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Tooby-Smith +-/ +import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.TimeOrder +import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic +/-! + +# Time Ordering on Field operator algebra + +-/ + +namespace FieldSpecification +open CrAnAlgebra +open HepLean.List +open FieldStatistic + +namespace FieldOpAlgebra +variable {𝓕 : FieldSpecification} + +lemma ι_timeOrder_superCommute_time {φ ψ : 𝓕.CrAnStates} + (hφψ : crAnTimeOrderRel φ ψ) (hψφ : crAnTimeOrderRel ψ φ) (a b : 𝓕.CrAnAlgebra) : + ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = + ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a * b)) := by + let pb (b : 𝓕.CrAnAlgebra) (hc : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) : + Prop := ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = + ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a * b)) + change pb b (Basis.mem_span _ b) + apply Submodule.span_induction + · intro x hx + obtain ⟨φs, rfl⟩ := hx + simp [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)) + change pa a (Basis.mem_span _ a) + apply Submodule.span_induction + · intro x hx + obtain ⟨φs', rfl⟩ := hx + simp [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 + trans crAnTimeOrderSign (φs' ++ [φ, ψ] ++ φs) + simp + rw [crAnTimeOrderSign] + have hp : List.Perm [φ,ψ] [ψ,φ] := by exact List.Perm.swap ψ φ [] + rw [Wick.koszulSign_perm_eq _ _ φ _ _ _ _ _ hp] + simp + rfl + simp_all + rw [h1] + simp + have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ φs' [φ, ψ] φs + (by simp_all) + rw [crAnTimeOrderList, show φs' ++ φ :: ψ :: φs = φs' ++ [φ, ψ] ++ φs by simp, h1] + have h2 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ φs' [ψ, φ] φs + (by simp_all) + rw [crAnTimeOrderList, show φs' ++ ψ :: φ :: φs = φs' ++ [ψ, φ] ++ φs by simp, h2] + repeat rw [ofCrAnList_append] + rw [smul_smul, mul_comm, ← smul_smul, ← smul_sub] + rw [map_mul, map_mul, map_mul, map_mul, map_mul, map_mul, map_mul, map_mul] + rw [← mul_smul_comm] + rw [mul_assoc, mul_assoc, mul_assoc ,mul_assoc ,mul_assoc ,mul_assoc] + rw [← mul_sub, ← mul_sub, mul_smul_comm, mul_smul_comm, ← smul_mul_assoc, + ← smul_mul_assoc] + rw [← sub_mul] + have h1 : (ι (ofCrAnList [φ, ψ]) - (exchangeSign (𝓕.crAnStatistics φ)) (𝓕.crAnStatistics ψ) • ι (ofCrAnList [ψ, φ])) = + ι [ofCrAnState φ, ofCrAnState ψ]ₛca := by + rw [superCommute_ofCrAnState_ofCrAnState] + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append] + simp only [instCommGroup.eq_1, List.singleton_append, Algebra.smul_mul_assoc, map_sub, + map_smul] + rw [← ofCrAnList_append] + simp + rw [h1] + have hc : ι ((superCommute (ofCrAnState φ)) (ofCrAnState ψ)) ∈ Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by + apply ι_superCommute_ofCrAnState_ofCrAnState_mem_center + rw [Subalgebra.mem_center_iff] at hc + repeat rw [← mul_assoc] + rw [hc] + repeat rw [mul_assoc] + rw [smul_mul_assoc] + 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 ⊢ + rw [← h1] + rw [← crAnTimeOrderList] + by_cases hq : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ) + · rw [ι_superCommute_of_diff_statistic hq] + simp + · rw [crAnTimeOrderSign, Wick.koszulSign_eq_rel_eq_stat _ _, ← crAnTimeOrderSign] + rw [timeOrder_ofCrAnList] + simp + exact hψφ + exact hφψ + simpa using hq + · simp [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] + · intro x y hx hy hpx hpy + simp_all [pb,mul_add, add_mul] + · intro x hx hpx + simp_all [pb, hpx] + +/-! + +## Defining normal order for `FiedOpAlgebra`. + +-/ + +lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra) + (h : a ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : ι 𝓣ᶠ(a) = 0 := by + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] at h + let p {k : Set 𝓕.CrAnAlgebra} (a : CrAnAlgebra 𝓕) (h : a ∈ AddSubgroup.closure k) := ι 𝓣ᶠ(a) = 0 + change p a h + apply AddSubgroup.closure_induction + · intro x hx + obtain ⟨a, ha, b, hb, rfl⟩ := Set.mem_mul.mp hx + obtain ⟨a, ha, c, hc, rfl⟩ := ha + simp only [p] + simp only [fieldOpIdealSet, exists_prop, exists_and_left, Set.mem_setOf_eq] at hc + match hc with + | Or.inl hc => + obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc + sorry + | Or.inr (Or.inl hc) => + obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc + sorry + | Or.inr (Or.inr (Or.inl hc)) => + obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc + sorry + | Or.inr (Or.inr (Or.inr hc)) => + obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc + sorry + · simp [p] + · intro x y hx hy + simp only [map_add, p] + intro h1 h2 + simp [h1, h2] + · intro x hx + simp [p] + +lemma ι_timeOrder_eq_of_equiv (a b : 𝓕.CrAnAlgebra) (h : a ≈ b) : + ι 𝓣ᶠ(a) = ι 𝓣ᶠ(b) := by + rw [equiv_iff_sub_mem_ideal] at h + rw [LinearMap.sub_mem_ker_iff.mp] + simp only [LinearMap.mem_ker, ← map_sub] + exact ι_timeOrder_zero_of_mem_ideal (a - b) h + +/-- Normal ordering on `FieldOpAlgebra`. -/ +noncomputable def timeOrder : FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where + toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.timeOrder) ι_timeOrder_eq_of_equiv + map_add' x y := by + obtain ⟨x, hx⟩ := ι_surjective x + obtain ⟨y, hy⟩ := ι_surjective y + subst hx hy + rw [← map_add, ι_apply, ι_apply, ι_apply] + rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk] + simp + map_smul' c y := by + obtain ⟨y, hy⟩ := ι_surjective y + subst hy + rw [← map_smul, ι_apply, ι_apply] + simp + +end FieldOpAlgebra +end FieldSpecification diff --git a/HepLean/PerturbationTheory/FieldStatistics/Basic.lean b/HepLean/PerturbationTheory/FieldStatistics/Basic.lean index 6faa8304..4e209eff 100644 --- a/HepLean/PerturbationTheory/FieldStatistics/Basic.lean +++ b/HepLean/PerturbationTheory/FieldStatistics/Basic.lean @@ -235,6 +235,10 @@ lemma ofList_map_eq_finset_prod (s : 𝓕 → FieldStatistic) : simp only [List.length_cons, List.nodup_cons] at hl exact hl.2 +lemma ofList_pair (s : 𝓕 → FieldStatistic) (φ1 φ2 : 𝓕) : + ofList s [φ1, φ2] = s φ1 * s φ2 := by + rw [ofList_cons_eq_mul, ofList_singleton] + /-! ## ofList and take diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean index 0763d3ee..0f3205e1 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean @@ -280,6 +280,40 @@ lemma koszulSign_swap_eq_rel {ψ φ : 𝓕} (h1 : le φ ψ) (h2 : le ψ φ) : ( apply Wick.koszulSignInsert_eq_perm exact List.Perm.append_left φs (List.Perm.swap ψ φ φs') +lemma koszulSign_eq_rel_eq_stat_append {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal 𝓕 le] + (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) → + koszulSign q le (φ :: ψ :: φs) = koszulSign q le φs := by + intro φs + simp [koszulSign, ← mul_assoc] + trans 1 * koszulSign q le φs + swap + simp + congr + simp [koszulSignInsert] + simp_all + rw [koszulSignInsert_eq_rel_eq_stat q le h1 h2 hq] + simp + +lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal 𝓕 le] + (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs' φs : List 𝓕) → + koszulSign q le (φs' ++ φ :: ψ :: φs) = koszulSign q le (φs' ++ φs) + | [], φs => by + simp + exact koszulSign_eq_rel_eq_stat_append q le h1 h2 hq φs + | φ'' :: φs', φs => by + simp [koszulSign] + rw [koszulSign_eq_rel_eq_stat h1 h2 hq φs' φs] + simp + 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 + lemma koszulSign_of_sorted : (φs : List 𝓕) → (hs : List.Sorted le φs) → koszulSign q le φs = 1 | [], _ => by @@ -351,4 +385,53 @@ lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] refine List.Perm.append_left φs'' ?_ exact List.Perm.symm (List.perm_insertionSort le φs) +/-! + +# koszulSign with permutations + +-/ + +lemma koszulSign_perm_eq_append [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) ( φs φs' φs2 : List 𝓕) + (hp : φs.Perm φs') : (h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) → + koszulSign q le (φs ++ φs2) = koszulSign q le (φs' ++ φs2) := by + let motive (φs φs' : List 𝓕) (hp : φs.Perm φs') : Prop := + (h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) → + koszulSign q le (φs ++ φs2) = koszulSign q le (φs' ++ φs2) + change motive φs φs' hp + apply List.Perm.recOn + · simp [motive] + · intro x l1 l2 h ih hxφ + simp_all [motive] + simp [koszulSign, ih] + left + apply koszulSignInsert_eq_perm + exact (List.perm_append_right_iff φs2).mpr h + · intro x y l h + simp_all [motive] + 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] + refine (ih2 ?_) + intro φ' hφ + refine h φ' ?_ + exact (List.Perm.mem_iff (id (List.Perm.symm h1))).mp hφ + +lemma koszulSign_perm_eq [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) : (φs1 φs φs' φs2 : List 𝓕) → + (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 + 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] + have ih := koszulSign_perm_eq φ φs1 φs φs' φs2 h hp + rw [ih] + congr 1 + apply koszulSignInsert_eq_perm + refine (List.perm_append_right_iff φs2).mpr ?_ + exact List.Perm.append_left φs1 hp + + end Wick diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean index 7eb5ee15..2f22ffa2 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean @@ -248,4 +248,42 @@ lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ · exact h φ1 (List.mem_cons_self _ _) +lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTotal 𝓕 le] [IsTrans 𝓕 le] + (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) → + koszulSignInsert q le φ φs = koszulSignInsert q le ψ φs + | [] => by + simp [koszulSignInsert] + | φ' :: φs => by + simp [koszulSignInsert] + simp_all + by_cases hr : le φ φ' + · simp [hr] + have h1' : le ψ φ' := by + apply IsTrans.trans ψ φ φ' h2 hr + simp [h1'] + 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ψφ'] + 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] + by_cases hφ'φ : le φ' φ + · have hφ'ψ : le φ' ψ := by + apply IsTrans.trans φ' φ ψ hφ'φ h1 + simp [hφ'φ, hφ'ψ] + · have hφ'ψ : ¬ le φ' ψ := by + intro hφ'ψ + apply hφ'φ + apply IsTrans.trans φ' ψ φ hφ'ψ h2 + simp_all [hφ'φ, hφ'ψ] + + + end Wick From 91071156204fcb1aa2a3c872ca5a318881f184fa Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Wed, 29 Jan 2025 15:08:43 +0000 Subject: [PATCH 4/8] feat: Add SuperCommute for FieldOpAlgebra --- .../Algebras/CrAnAlgebra/Grading.lean | 102 +++++++- .../Algebras/CrAnAlgebra/SuperCommute.lean | 47 ++++ .../Algebras/FieldOpAlgebra/Basic.lean | 236 ++++++++++++++++++ .../Algebras/FieldOpAlgebra/SuperCommute.lean | 114 +++++++++ .../Algebras/FieldOpAlgebra/TimeOrder.lean | 218 +++++++++++++++- 5 files changed, 708 insertions(+), 9 deletions(-) create mode 100644 HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean index f5a34ef4..00c4f43c 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean @@ -37,6 +37,11 @@ lemma ofCrAnList_bosonic_or_fermionic (φs : List 𝓕.CrAnStates) : · right exact ofCrAnList_mem_statisticSubmodule_of φs fermionic (by simpa using h) +lemma ofCrAnState_bosonic_or_fermionic (φ : 𝓕.CrAnStates) : + ofCrAnState φ ∈ statisticSubmodule bosonic ∨ ofCrAnState φ ∈ statisticSubmodule fermionic := by + rw [← ofCrAnList_singleton] + exact ofCrAnList_bosonic_or_fermionic [φ] + /-- The projection of an element of `CrAnAlgebra` onto it's bosonic part. -/ def bosonicProj : 𝓕.CrAnAlgebra →ₗ[ℂ] statisticSubmodule (𝓕 := 𝓕) bosonic := Basis.constr ofCrAnListBasis ℂ fun φs => @@ -183,6 +188,7 @@ lemma bosonicProj_add_fermionicProj (a : 𝓕.CrAnAlgebra) : · simp [h] · simp [h] + lemma coeAddMonoidHom_apply_eq_bosonic_plus_fermionic (a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) : DirectSum.coeAddMonoidHom statisticSubmodule a = a.1 bosonic + a.1 fermionic := by @@ -229,7 +235,7 @@ lemma directSum_eq_bosonic_plus_fermionic conv_lhs => rw [hx, hy] abel -instance : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where +instance crAnAlgebraGrade : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where one_mem := by simp only [statisticSubmodule] refine Submodule.mem_span.mpr fun p a => a ?_ @@ -296,6 +302,100 @@ lemma eq_zero_of_bosonic_and_fermionic {a : 𝓕.CrAnAlgebra} rw [ha, hb] at hc simpa using hc +lemma bosonicProj_mul (a b : 𝓕.CrAnAlgebra) : + (a * b).bosonicProj.1 = a.bosonicProj.1 * b.bosonicProj.1 + + a.fermionicProj.1 * b.fermionicProj.1 := by + conv_lhs => + rw [← bosonicProj_add_fermionicProj a] + rw [← bosonicProj_add_fermionicProj b] + simp [mul_add, add_mul] + rw [bosonicProj_of_mem_bosonic] + conv_lhs => + left + right + rw [bosonicProj_of_mem_fermionic _ + (by + have h1 : fermionic = fermionic + bosonic := by simp + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp)] + conv_lhs => + right + left + rw [bosonicProj_of_mem_fermionic _ + (by + have h1 : fermionic = bosonic + fermionic := by simp + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp)] + conv_lhs => + right + right + rw [bosonicProj_of_mem_bosonic _ + (by + have h1 : bosonic = fermionic + fermionic := by simp; rfl + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp)] + simp + · have h1 : bosonic = bosonic + bosonic := by simp; rfl + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp + +lemma fermionicProj_mul (a b : 𝓕.CrAnAlgebra) : + (a * b).fermionicProj.1 = a.bosonicProj.1 * b.fermionicProj.1 + + a.fermionicProj.1 * b.bosonicProj.1 := by + conv_lhs => + rw [← bosonicProj_add_fermionicProj a] + rw [← bosonicProj_add_fermionicProj b] + simp [mul_add, add_mul] + conv_lhs => + left + left + rw [fermionicProj_of_mem_bosonic _ + (by + have h1 : bosonic = bosonic + bosonic := by simp; rfl + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp)] + conv_lhs => + left + right + rw [fermionicProj_of_mem_fermionic _ + (by + have h1 : fermionic = fermionic + bosonic := by simp + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp)] + conv_lhs => + right + left + rw [fermionicProj_of_mem_fermionic _ + (by + have h1 : fermionic = bosonic + fermionic := by simp + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp)] + conv_lhs => + right + right + rw [fermionicProj_of_mem_bosonic _ + (by + have h1 : bosonic = fermionic + fermionic := by simp; rfl + conv_lhs => rw [h1] + apply crAnAlgebraGrade.mul_mem + simp + simp)] + simp + abel end diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean index ce9e4a61..f5ceffce 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean @@ -688,6 +688,18 @@ lemma superCommute_fermionic_fermionic_symm {a b : 𝓕.CrAnAlgebra} rw [superCommute_fermionic_fermionic hb ha] abel +lemma superCommute_expand_bosonicProj_fermionicProj (a b : 𝓕.CrAnAlgebra) : + [a, b]ₛca = bosonicProj a * bosonicProj b - bosonicProj b * bosonicProj a + + bosonicProj a * fermionicProj b - fermionicProj b * bosonicProj a + + 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), + superCommute_fermionic_bonsonic (by simp) (by simp), + superCommute_fermionic_fermionic (by simp) (by simp)] + abel + lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List 𝓕.CrAnStates) : [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule bosonic ∨ [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic := by @@ -725,6 +737,41 @@ lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h1) apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2) +lemma superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic (φ φ' : 𝓕.CrAnStates) : + [ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule bosonic ∨ + [ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule fermionic := by + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] + exact superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [φ'] + +lemma superCommute_superCommute_ofCrAnState_bosonic_or_fermionic (φ1 φ2 φ3 : 𝓕.CrAnStates) : + [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule bosonic ∨ + [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca ∈ statisticSubmodule fermionic := by + rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ2 φ3 with hs | hs + <;> rcases ofCrAnState_bosonic_or_fermionic φ1 with h1 | h1 + · left + have h : bosonic = bosonic + bosonic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade h1 hs + · right + have h : fermionic = fermionic + bosonic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade h1 hs + · right + have h : fermionic = bosonic + fermionic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade h1 hs + · left + have h : bosonic = fermionic + fermionic := by + simp only [add_eq_mul, instCommGroup, mul_self] + rfl + rw [h] + apply superCommute_grade h1 hs lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates) (ha : a ∈ statisticSubmodule bosonic) : diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean index aa71461c..10993f0c 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean @@ -191,5 +191,241 @@ lemma ι_superCommute_ofCrAnState_ofCrAnState_mem_center (φ ψ : 𝓕.CrAnState rw [← h0] abel +/-! + +## The kernal of ι +-/ + +lemma ι_eq_zero_iff_mem_ideal (x : CrAnAlgebra 𝓕) : + ι x = 0 ↔ x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by + rw [ι_apply] + change ⟦x⟧ = ⟦0⟧ ↔ _ + simp only [ringConGen, Quotient.eq] + rw [TwoSidedIdeal.mem_iff] + simp only + rfl + +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 + 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 + · left + rw [bosonicProj_of_mem_bosonic _ h] + simpa using hx' + · right + rw [bosonicProj_of_mem_fermionic _ h] + · rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φc φc' with h | h + · left + rw [bosonicProj_of_mem_bosonic _ h] + simpa using hx' + · right + rw [bosonicProj_of_mem_fermionic _ h] + · rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φa φa' with h | h + · left + rw [bosonicProj_of_mem_bosonic _ h] + simpa using hx' + · right + rw [bosonicProj_of_mem_fermionic _ h] + · rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ φ' with h | h + · left + rw [bosonicProj_of_mem_bosonic _ h] + simpa using hx' + · right + rw [bosonicProj_of_mem_fermionic _ h] + + +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 + 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 + · right + rw [fermionicProj_of_mem_bosonic _ h] + · left + rw [fermionicProj_of_mem_fermionic _ h] + simpa using hx' + · rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φc φc' with h | h + · right + rw [fermionicProj_of_mem_bosonic _ h] + · left + rw [fermionicProj_of_mem_fermionic _ h] + simpa using hx' + · rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φa φa' with h | h + · right + rw [fermionicProj_of_mem_bosonic _ h] + · left + rw [fermionicProj_of_mem_fermionic _ h] + simpa using hx' + · rcases superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic φ φ' with h | h + · right + rw [fermionicProj_of_mem_bosonic _ h] + · left + rw [fermionicProj_of_mem_fermionic _ h] + simpa using hx' + +lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : + x.bosonicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] at hx + let p {k : Set 𝓕.CrAnAlgebra} (a : CrAnAlgebra 𝓕) (h : a ∈ AddSubgroup.closure k) : Prop := + a.bosonicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet + change p x hx + apply AddSubgroup.closure_induction + · intro x hx + simp [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] + 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 + apply TwoSidedIdeal.add_mem + · /- boson, boson, boson mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(bosonicProj d) * ↑(bosonicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use bosonicProj d + simp + use (bosonicProj y).1 + simp [hby] + · use ↑(bosonicProj b) + simp + · /- fermion, fermion, boson mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(fermionicProj d) * ↑(fermionicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use fermionicProj d + simp + use (fermionicProj y).1 + simp [hby, hfy] + · use ↑(bosonicProj b) + simp + apply TwoSidedIdeal.add_mem + · /- boson, fermion, fermion mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(bosonicProj d) * ↑(fermionicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use bosonicProj d + simp + use (fermionicProj y).1 + simp [hby, hfy] + · use ↑(fermionicProj b) + simp + · /- fermion, boson, fermion mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(fermionicProj d) * ↑(bosonicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use fermionicProj d + simp + use (bosonicProj y).1 + simp [hby, hfy] + · use ↑(fermionicProj b) + simp + · simp [hby] + apply TwoSidedIdeal.add_mem + · /- fermion, fermion, boson mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(fermionicProj d) * ↑(fermionicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use fermionicProj d + simp + use (fermionicProj y).1 + simp [hby, hfy] + · use ↑(bosonicProj b) + simp + · /- boson, fermion, fermion mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(bosonicProj d) * ↑(fermionicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use bosonicProj d + simp + use (fermionicProj y).1 + simp [hby, hfy] + · use ↑(fermionicProj b) + simp + · simp [hfy] + apply TwoSidedIdeal.add_mem + · /- boson, boson, boson mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(bosonicProj d) * ↑(bosonicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use bosonicProj d + simp + use (bosonicProj y).1 + simp [hby] + · use ↑(bosonicProj b) + simp + · /- fermion, boson, fermion mem-/ + rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] + refine Set.mem_of_mem_of_subset ?_ AddSubgroup.subset_closure + apply Set.mem_mul.mpr + use ↑(fermionicProj d) * ↑(bosonicProj y) + apply And.intro + · apply Set.mem_mul.mpr + use fermionicProj d + simp + use (bosonicProj y).1 + simp [hby, hfy] + · use ↑(fermionicProj b) + simp + · simp [hfy, hby] + · simp [p] + · intro x y hx hy hpx hpy + simp_all [p] + apply TwoSidedIdeal.add_mem + exact hpx + exact hpy + · intro x hx + simp [p] + +lemma fermionicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : + x.fermionicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by + 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_all + +lemma ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero (x : CrAnAlgebra 𝓕) : + ι x = 0 ↔ ι x.bosonicProj.1 = 0 ∧ ι x.fermionicProj.1 = 0 := by + apply Iff.intro + · intro h + rw [@ι_eq_zero_iff_mem_ideal] at h ⊢ + rw [ι_eq_zero_iff_mem_ideal] + apply And.intro + · exact bosonicProj_mem_ideal x h + · exact fermionicProj_mem_ideal x h + · intro h + rw [← bosonicProj_add_fermionicProj x] + simp_all + + + end FieldOpAlgebra end FieldSpecification diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean new file mode 100644 index 00000000..e5559773 --- /dev/null +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Tooby-Smith +-/ +import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.TimeOrder +import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic +/-! + +# SuperCommute on Field operator algebra + +-/ + +namespace FieldSpecification +open CrAnAlgebra +open HepLean.List +open FieldStatistic + +namespace FieldOpAlgebra +variable {𝓕 : FieldSpecification} + +lemma ι_superCommute_eq_zero_of_ι_right_zero (a b : 𝓕.CrAnAlgebra) (h : ι b = 0) : + ι [a, b]ₛca = 0 := by + rw [superCommute_expand_bosonicProj_fermionicProj] + rw [ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero] at h + simp_all + +lemma ι_superCommute_eq_zero_of_ι_left_zero (a b : 𝓕.CrAnAlgebra) (h : ι a = 0) : + ι [a, b]ₛca = 0 := by + rw [superCommute_expand_bosonicProj_fermionicProj] + rw [ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero] at h + simp_all + +/-! + +## Defining normal order for `FiedOpAlgebra`. + +-/ + +lemma ι_superCommute_right_zero_of_mem_ideal (a b : 𝓕.CrAnAlgebra) + (h : b ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : ι [a, b]ₛca = 0 := by + apply ι_superCommute_eq_zero_of_ι_right_zero + exact (ι_eq_zero_iff_mem_ideal b).mpr h + +lemma ι_superCommute_eq_of_equiv_right (a b1 b2 : 𝓕.CrAnAlgebra) (h : b1 ≈ b2) : + ι [a, b1]ₛca = ι [a, b2]ₛca := by + rw [equiv_iff_sub_mem_ideal] at h + rw [LinearMap.sub_mem_ker_iff.mp] + simp only [LinearMap.mem_ker, ← map_sub] + exact ι_superCommute_right_zero_of_mem_ideal a (b1 - b2) h + +noncomputable def superCommuteRight (a : 𝓕.CrAnAlgebra) : FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where + toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.superCommute a) (ι_superCommute_eq_of_equiv_right a) + map_add' x y := by + obtain ⟨x, hx⟩ := ι_surjective x + obtain ⟨y, hy⟩ := ι_surjective y + subst hx hy + rw [← map_add, ι_apply, ι_apply, ι_apply] + rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk] + simp + map_smul' c y := by + obtain ⟨y, hy⟩ := ι_surjective y + subst hy + rw [← map_smul, ι_apply, ι_apply] + simp + +lemma superCommuteRight_apply_ι (a b : 𝓕.CrAnAlgebra) : superCommuteRight a (ι b) = ι [a, b]ₛca := by + rfl + +lemma superCommuteRight_apply_quot (a b : 𝓕.CrAnAlgebra) : superCommuteRight a ⟦b⟧= ι [a, b]ₛca := by + rfl + +lemma superCommuteRight_eq_of_equiv (a1 a2 : 𝓕.CrAnAlgebra) (h : a1 ≈ a2) : + superCommuteRight a1 = superCommuteRight a2 := by + rw [equiv_iff_sub_mem_ideal] at h + ext b + obtain ⟨b, rfl⟩ := ι_surjective b + have ha1b1 : (superCommuteRight (a1 - a2)) (ι b) = 0 := by + rw [superCommuteRight_apply_ι] + apply ι_superCommute_eq_zero_of_ι_left_zero + exact (ι_eq_zero_iff_mem_ideal (a1 - a2)).mpr h + simp_all [superCommuteRight_apply_ι] + trans ι ((superCommute a2) b) + 0 + rw [← ha1b1] + simp + simp + +noncomputable def superCommute : FieldOpAlgebra 𝓕 →ₗ[ℂ] + FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where + toFun := Quotient.lift superCommuteRight superCommuteRight_eq_of_equiv + map_add' x y := by + obtain ⟨x, rfl⟩ := ι_surjective x + obtain ⟨y, rfl⟩ := ι_surjective y + ext b + obtain ⟨b, rfl⟩ := ι_surjective b + rw [← map_add, ι_apply, ι_apply, ι_apply, ι_apply] + rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk] + simp + rw [superCommuteRight_apply_quot, superCommuteRight_apply_quot, superCommuteRight_apply_quot] + simp + map_smul' c y := by + obtain ⟨y, rfl⟩ := ι_surjective y + ext b + obtain ⟨b, rfl⟩ := ι_surjective b + rw [← map_smul, ι_apply, ι_apply, ι_apply] + simp + rw [superCommuteRight_apply_quot, superCommuteRight_apply_quot] + simp + +lemma ι_superCommute (a b : 𝓕.CrAnAlgebra) : ι [a, b]ₛca = superCommute (ι a) (ι b) := by + rfl + +end FieldOpAlgebra +end FieldSpecification diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean index 6db44c8d..16b1e76f 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean @@ -19,7 +19,164 @@ open FieldStatistic namespace FieldOpAlgebra variable {𝓕 : FieldSpecification} -lemma ι_timeOrder_superCommute_time {φ ψ : 𝓕.CrAnStates} +lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates} + (φs1 φs2 : List 𝓕.CrAnStates) (h : + crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧ + crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧ + crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2): + ι 𝓣ᶠ(ofCrAnList φs1 * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ofCrAnList φs2) + = 0 := by + let l1 := + (List.takeWhile (fun c => ¬ crAnTimeOrderRel φ1 c) ((φs1 ++ φs2).insertionSort crAnTimeOrderRel)) + ++ (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ crAnTimeOrderRel c φ1) φs1) + let l2 := (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ crAnTimeOrderRel c φ1) φs2) + ++ (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ ¬ crAnTimeOrderRel c φ1) ((φs1 ++ φs2).insertionSort crAnTimeOrderRel)) + have h123 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ1 :: φ2 :: φ3 :: φs2)) = + crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ2, φ3]) * ι (ofCrAnList l2)):= by + have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ1, φ2, φ3] φs2 + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ2 :: φ3 :: φs2 = φs1 ++ [φ1, φ2, φ3] ++ φs2 by simp, + crAnTimeOrderList, h1] + simp only [List.append_assoc, List.singleton_append, decide_not, + Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] + have h132 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ1 :: φ3 :: φ2 :: φs2)) = + crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ3, φ2]) * ι (ofCrAnList l2)):= by + have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ1, φ3, φ2] φs2 + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ3 :: φ2 :: φs2 = φs1 ++ [φ1, φ3, φ2] ++ φs2 by simp, + crAnTimeOrderList, h1] + simp only [List.singleton_append, decide_not, + Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] + congr 1 + have hp : List.Perm [φ1, φ3, φ2] [φ1, φ2, φ3] := by + refine List.Perm.cons φ1 ?_ + exact List.Perm.swap φ2 φ3 [] + rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp, ← crAnTimeOrderSign] + · simp + · intro φ4 hφ4 + simp at hφ4 + rcases hφ4 with hφ4 | hφ4 | hφ4 + all_goals + subst hφ4 + simp_all + have hp231 : List.Perm [φ2, φ3, φ1] [φ1, φ2, φ3] := by + refine List.Perm.trans (l₂ := [φ2, φ1, φ3]) ?_ ?_ + refine List.Perm.cons φ2 (List.Perm.swap φ1 φ3 []) + exact List.Perm.swap φ1 φ2 [φ3] + have h231 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ2 :: φ3 :: φ1 :: φs2)) = + crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ2, φ3, φ1]) * ι (ofCrAnList l2)):= by + have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ2, φ3, φ1] φs2 + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ2 :: φ3 :: φ1 :: φs2 = φs1 ++ [φ2, φ3, φ1] ++ φs2 by simp, + crAnTimeOrderList, h1] + simp only [List.singleton_append, decide_not, + Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] + congr 1 + rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp231, ← crAnTimeOrderSign] + · simp + · intro φ4 hφ4 + simp at hφ4 + rcases hφ4 with hφ4 | hφ4 | hφ4 + all_goals + subst hφ4 + simp_all + have h321 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ3 :: φ2 :: φ1 :: φs2)) = + crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ3, φ2, φ1]) * ι (ofCrAnList l2)):= by + have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ3, φ2, φ1] φs2 + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ3 :: φ2 :: φ1 :: φs2 = φs1 ++ [φ3, φ2, φ1] ++ φs2 by simp, + crAnTimeOrderList, h1] + simp only [List.singleton_append, decide_not, + Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] + congr 1 + have hp : List.Perm [φ3, φ2, φ1] [φ1, φ2, φ3] := by + refine List.Perm.trans ?_ hp231 + exact List.Perm.swap φ2 φ3 [φ1] + rw [crAnTimeOrderSign, Wick.koszulSign_perm_eq _ _ φ1 _ _ _ _ _ hp, ← crAnTimeOrderSign] + · simp + · intro φ4 hφ4 + simp 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 + rw [superCommute_ofCrAnList_ofCrAnList, superCommute_ofCrAnList_ofCrAnList] + simp [mul_sub, sub_mul, ← ofCrAnList_append] + rw [h123, h132, h231, h321] + simp [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 + 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] + rw [← smul_mul_assoc, ← mul_smul_comm] + repeat rw [mul_assoc] + rw [← mul_sub, ← mul_sub, ← mul_sub] + rw [← sub_mul, ← sub_mul, ← sub_mul] + trans ι (ofCrAnList l1) * ι [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ι (ofCrAnList l2) + rw [mul_assoc] + congr + rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton] + rw [superCommute_ofCrAnList_ofCrAnList] + simp + rw [superCommute_ofCrAnList_ofCrAnList, superCommute_ofCrAnList_ofCrAnList] + simp [smul_sub] + simp_all + +lemma ι_timeOrder_superCommute_superCommute_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates} + (φs1 φs2 : List 𝓕.CrAnStates): + ι 𝓣ᶠ(ofCrAnList φs1 * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ofCrAnList φs2) + = 0 := by + by_cases h : + crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧ + crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧ + crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2 + · exact ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList φs1 φs2 h + · rw [timeOrder_timeOrder_mid] + rw [timeOrder_superCommute_ofCrAnState_superCommute_all_not_crAnTimeOrderRel _ _ _ h] + simp + +@[simp] +lemma ι_timeOrder_superCommute_superCommute {φ1 φ2 φ3 : 𝓕.CrAnStates} (a b : 𝓕.CrAnAlgebra): + ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * b) = 0 := by + let pb (b : 𝓕.CrAnAlgebra) (hc : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) : + Prop := ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * b) = 0 + change pb b (Basis.mem_span _ b) + apply Submodule.span_induction + · intro x hx + obtain ⟨φs, rfl⟩ := hx + simp [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] + exact ι_timeOrder_superCommute_superCommute_ofCrAnList φs' φs + · simp [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] + · intro x y hx hy hpx hpy + simp_all [pb,mul_add, add_mul] + · intro x hx hpx + simp_all [pb, hpx] + +lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates} (hφψ : crAnTimeOrderRel φ ψ) (hψφ : crAnTimeOrderRel ψ φ) (a b : 𝓕.CrAnAlgebra) : ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a * b)) := by @@ -110,6 +267,25 @@ lemma ι_timeOrder_superCommute_time {φ ψ : 𝓕.CrAnStates} · intro x hx hpx simp_all [pb, hpx] + +lemma ι_timeOrder_superCommute_neq_time {φ ψ : 𝓕.CrAnStates} + (hφψ : ¬ (crAnTimeOrderRel φ ψ ∧ crAnTimeOrderRel ψ φ)) (a b : 𝓕.CrAnAlgebra) : + ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = 0 := by + rw [timeOrder_timeOrder_mid] + have hφψ : ¬ (crAnTimeOrderRel φ ψ) ∨ ¬ (crAnTimeOrderRel ψ φ) := by + exact Decidable.not_and_iff_or_not.mp hφψ + rcases hφψ with hφψ | hφψ + · rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel ] + have ht := IsTotal.total (r := crAnTimeOrderRel) φ ψ + simp_all + simp_all + · rw [superCommute_ofCrAnState_ofCrAnState_symm] + simp + rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel ] + simp + simp_all + + /-! ## Defining normal order for `FiedOpAlgebra`. @@ -130,16 +306,42 @@ lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra) match hc with | Or.inl hc => obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc - sorry + simp | Or.inr (Or.inl hc) => - obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc - sorry + obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc + by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) + · rw [ι_timeOrder_superCommute_eq_time] + simp + rw [ι_superCommute_of_create_create] + simp + · exact hφa + · exact hφb + · exact heqt.1 + · exact heqt.2 + · rw [ι_timeOrder_superCommute_neq_time heqt] | Or.inr (Or.inr (Or.inl hc)) => - obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc - sorry + obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc + by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) + · rw [ι_timeOrder_superCommute_eq_time] + simp + rw [ι_superCommute_of_annihilate_annihilate] + simp + · exact hφa + · exact hφb + · exact heqt.1 + · exact heqt.2 + · rw [ι_timeOrder_superCommute_neq_time heqt] | Or.inr (Or.inr (Or.inr hc)) => - obtain ⟨φa, φa', hφa, hφa', rfl⟩ := hc - sorry + obtain ⟨φa, φb, hdiff, rfl⟩ := hc + by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) + · rw [ι_timeOrder_superCommute_eq_time] + simp + rw [ι_superCommute_of_diff_statistic] + simp + · exact hdiff + · exact heqt.1 + · exact heqt.2 + · rw [ι_timeOrder_superCommute_neq_time heqt] · simp [p] · intro x y hx hy simp only [map_add, p] From a329661c2413e92a26853f1fe5dbd32636414bae Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Wed, 29 Jan 2025 15:33:40 +0000 Subject: [PATCH 5/8] refactor: Simplify normal order --- .../Algebras/FieldOpAlgebra/NormalOrder.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean index 89862492..b518f2e2 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/NormalOrder.lean @@ -221,15 +221,13 @@ lemma ι_normalOrder_eq_of_equiv (a b : 𝓕.CrAnAlgebra) (h : a ≈ b) : noncomputable def normalOrder : FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.normalOrder) ι_normalOrder_eq_of_equiv map_add' x y := by - obtain ⟨x, hx⟩ := ι_surjective x - obtain ⟨y, hy⟩ := ι_surjective y - subst hx hy + obtain ⟨x, rfl⟩ := ι_surjective x + obtain ⟨y, rfl⟩ := ι_surjective y rw [← map_add, ι_apply, ι_apply, ι_apply] rw [Quotient.lift_mk, Quotient.lift_mk, Quotient.lift_mk] simp map_smul' c y := by - obtain ⟨y, hy⟩ := ι_surjective y - subst hy + obtain ⟨y, rfl⟩ := ι_surjective y rw [← map_smul, ι_apply, ι_apply] simp From e5c85ac109016c5cebf97e5accb4310c34657bb2 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Wed, 29 Jan 2025 16:06:28 +0000 Subject: [PATCH 6/8] refactor: Lint --- HepLean/Mathematics/List/InsertIdx.lean | 2 +- HepLean/Mathematics/List/InsertionSort.lean | 147 ++++++++++-------- .../Algebras/CrAnAlgebra/Grading.lean | 15 +- .../Algebras/CrAnAlgebra/SuperCommute.lean | 106 ++++++------- .../Algebras/CrAnAlgebra/TimeOrder.lean | 14 +- .../Algebras/FieldOpAlgebra/Basic.lean | 11 +- .../Algebras/FieldOpAlgebra/SuperCommute.lean | 2 +- .../Algebras/FieldOpAlgebra/TimeOrder.lean | 24 +-- .../FieldStatistics/ExchangeSign.lean | 2 +- .../PerturbationTheory/Koszul/KoszulSign.lean | 31 ++-- .../Koszul/KoszulSignInsert.lean | 7 +- 11 files changed, 179 insertions(+), 182 deletions(-) diff --git a/HepLean/Mathematics/List/InsertIdx.lean b/HepLean/Mathematics/List/InsertIdx.lean index 99943e96..c26e125d 100644 --- a/HepLean/Mathematics/List/InsertIdx.lean +++ b/HepLean/Mathematics/List/InsertIdx.lean @@ -119,7 +119,7 @@ lemma insertIdx_eraseIdx_fin {I : Type} : exact insertIdx_eraseIdx_fin as ⟨n, Nat.lt_of_succ_lt_succ h⟩ lemma insertIdx_length_fst_append {I : Type} (φ : I) : (φs φs' : List I) → - List.insertIdx φs.length φ (φs ++ φs') = (φs ++ φ :: φs') + List.insertIdx φs.length φ (φs ++ φs') = (φs ++ φ :: φs') | [], φs' => by simp | φ' :: φs, φs' => by simp diff --git a/HepLean/Mathematics/List/InsertionSort.lean b/HepLean/Mathematics/List/InsertionSort.lean index 2a914d88..b979ac23 100644 --- a/HepLean/Mathematics/List/InsertionSort.lean +++ b/HepLean/Mathematics/List/InsertionSort.lean @@ -92,14 +92,15 @@ lemma insertionSortMin_lt_mem_insertionSortDropMinPos_of_lt {α : Type} (r : α exact Fin.succAbove_ne (insertionSortMinPosFin r a l) i lemma insertionSort_insertionSort {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (l1 : List α): + [IsTotal α r] [IsTrans α r] (l1 : List α) : List.insertionSort r (List.insertionSort r l1) = List.insertionSort r l1 := by apply List.Sorted.insertionSort_eq exact List.sorted_insertionSort r l1 lemma orderedInsert_commute {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) : (l : List α) → - List.orderedInsert r a (List.orderedInsert r b l) = List.orderedInsert r b (List.orderedInsert r a l) + List.orderedInsert r a (List.orderedInsert r b l) = + List.orderedInsert r b (List.orderedInsert r a l) | [] => by have hrb : r b a := by have ht := IsTotal.total (r := r) a b @@ -125,11 +126,11 @@ lemma orderedInsert_commute {α : Type} (r : α → α → Prop) [DecidableRel r exact orderedInsert_commute r a b hr l lemma insertionSort_orderedInsert_append {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) : (l1 l2 : List α) → - List.insertionSort r (List.orderedInsert r a l1 ++ l2) = List.insertionSort r (a :: l1 ++ l2) + [IsTotal α r] [IsTrans α r] (a : α) : (l1 l2 : List α) → + List.insertionSort r (List.orderedInsert r a l1 ++ l2) = List.insertionSort r (a :: l1 ++ l2) | [], l2 => by simp - | b :: l1, l2 => by + | b :: l1, l2 => by conv_lhs => simp by_cases h : r a b · simp [h] @@ -138,25 +139,25 @@ lemma insertionSort_orderedInsert_append {α : Type} (r : α → α → Prop) [D simp rw [orderedInsert_commute r a b h] - lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] : (l1 l2 : List α) → + [IsTotal α r] [IsTrans α r] : (l1 l2 : List α) → List.insertionSort r (List.insertionSort r l1 ++ l2) = List.insertionSort r (l1 ++ l2) | [], l2 => by simp - | a :: l1, l2 => by + | a :: l1, l2 => by conv_lhs => simp rw [insertionSort_orderedInsert_append] simp rw [insertionSort_insertionSort_append r l1 l2] lemma insertionSort_append_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] : (l1 l2 l3 : List α) → - List.insertionSort r (l1 ++ List.insertionSort r l2 ++ l3) = List.insertionSort r (l1 ++ l2 ++ l3) + [IsTotal α r] [IsTrans α r] : (l1 l2 l3 : List α) → + List.insertionSort r (l1 ++ List.insertionSort r l2 ++ l3) = + List.insertionSort r (l1 ++ l2 ++ l3) | [], l2, l3 => by simp exact insertionSort_insertionSort_append r l2 l3 - | a :: l1, l2, l3 => by + | a :: l1, l2, l3 => by simp only [List.insertionSort, List.append_eq] rw [insertionSort_append_insertionSort_append r l1 l2 l3] @@ -168,7 +169,7 @@ lemma orderedInsert_length {α : Type} (r : α → α → Prop) [DecidableRel r] lemma takeWhile_orderedInsert {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] [IsTrans α r] - (a b : α) (hr : ¬ r a b) : (l : List α) → + (a b : α) (hr : ¬ r a b) : (l : List α) → (List.takeWhile (fun c => !decide (r a c)) (List.orderedInsert r b l)).length = (List.takeWhile (fun c => !decide (r a c)) l).length + 1 | [] => by @@ -193,7 +194,7 @@ lemma takeWhile_orderedInsert {α : Type} (r : α → α → Prop) [DecidableRel lemma takeWhile_orderedInsert' {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] [IsTrans α r] - (a b : α) (hr : ¬ r a b) : (l : List α) → + (a b : α) (hr : ¬ r a b) : (l : List α) → (List.takeWhile (fun c => !decide (r b c)) (List.orderedInsert r a l)).length = (List.takeWhile (fun c => !decide (r b c)) l).length | [] => by @@ -219,11 +220,11 @@ lemma takeWhile_orderedInsert' {α : Type} (r : α → α → Prop) [DecidableRe · simp [hac, h] exact takeWhile_orderedInsert' r a b hr l -lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) (n : ℕ) : (l : List α) → +lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [DecidableRel r] + [IsTotal α r] [IsTrans α r] (a b : α) (hr : ¬ r a b) (n : ℕ) : (l : List α) → (hn : n + 2 < (a :: b :: l).length) → - insertionSortEquiv r (a :: b :: l) ⟨n + 2, hn⟩ = (finCongr (by simp)) - (insertionSortEquiv r (b :: a :: l) ⟨n + 2, hn⟩):= by + insertionSortEquiv r (a :: b :: l) ⟨n + 2, hn⟩ = (finCongr (by simp)) + (insertionSortEquiv r (b :: a :: l) ⟨n + 2, hn⟩) := by have hrba : r b a:= by have ht := IsTotal.total (r := r) a b simp_all @@ -245,12 +246,21 @@ lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [Decidabl conv_rhs => erw [orderedInsertEquiv_fin_succ] ext simp - let a1 : Fin ((List.orderedInsert r b (List.insertionSort r l)).length + 1) := ⟨↑(orderedInsertPos r (List.orderedInsert r b (List.insertionSort r l)) a), orderedInsertPos_lt_length r (List.orderedInsert r b (List.insertionSort r l)) a⟩ - let b1 : Fin ((List.insertionSort r l).length + 1) := ⟨↑(orderedInsertPos r (List.insertionSort r l) b), orderedInsertPos_lt_length r (List.insertionSort r l) b⟩ - let a2 : Fin ((List.insertionSort r l).length + 1) := ⟨↑(orderedInsertPos r (List.insertionSort r l) a), orderedInsertPos_lt_length r (List.insertionSort r l) a⟩ - let b2 : Fin ((List.orderedInsert r a (List.insertionSort r l)).length + 1) := ⟨↑(orderedInsertPos r (List.orderedInsert r a (List.insertionSort r l)) b), orderedInsertPos_lt_length r (List.orderedInsert r a (List.insertionSort r l)) b⟩ + let a1 : Fin ((List.orderedInsert r b (List.insertionSort r l)).length + 1) := + ⟨↑(orderedInsertPos r (List.orderedInsert r b (List.insertionSort r l)) a), + orderedInsertPos_lt_length r (List.orderedInsert r b (List.insertionSort r l)) a⟩ + let b1 : Fin ((List.insertionSort r l).length + 1) := + ⟨↑(orderedInsertPos r (List.insertionSort r l) b), + orderedInsertPos_lt_length r (List.insertionSort r l) b⟩ + let a2 : Fin ((List.insertionSort r l).length + 1) := + ⟨↑(orderedInsertPos r (List.insertionSort r l) a), + orderedInsertPos_lt_length r (List.insertionSort r l) a⟩ + let b2 : Fin ((List.orderedInsert r a (List.insertionSort r l)).length + 1) := + ⟨↑(orderedInsertPos r (List.orderedInsert r a (List.insertionSort r l)) b), + orderedInsertPos_lt_length r (List.orderedInsert r a (List.insertionSort r l)) b⟩ have ht : (List.takeWhile (fun c => !decide (r b c)) (List.insertionSort r l)) - = (List.takeWhile (fun c => !decide (r b c)) ((List.takeWhile (fun c => !decide (r a c)) (List.insertionSort r l)))) := by + = (List.takeWhile (fun c => !decide (r b c)) + ((List.takeWhile (fun c => !decide (r a c)) (List.insertionSort r l)))) := by rw [List.takeWhile_takeWhile] simp congr @@ -300,13 +310,14 @@ lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [Decidabl · simp [ha] · simp [ha] rw [hbs1, has2, hb, ha2] - have hnat (a2 b2 n : ℕ) (h : b2 ≤ a2) : (if (if ↑n < ↑b2 then ↑n else ↑n + 1) < ↑a2 + 1 then if ↑n < ↑b2 then ↑n else ↑n + 1 + have hnat (a2 b2 n : ℕ) (h : b2 ≤ a2) : (if (if ↑n < ↑b2 then ↑n else ↑n + 1) < ↑a2 + 1 then + if ↑n < ↑b2 then ↑n else ↑n + 1 else (if ↑n < ↑b2 then ↑n else ↑n + 1) + 1) = if (if ↑n < ↑a2 then ↑n else ↑n + 1) < ↑b2 then if ↑n < ↑a2 then ↑n else ↑n + 1 else (if ↑n < ↑a2 then ↑n else ↑n + 1) + 1 := by by_cases hnb2 : n < b2 · simp [hnb2] - have h1 : n < a2 + 1 := by omega + have h1 : n < a2 + 1 := by omega have h2 : n < a2 := by omega simp [h1, h2, hnb2] · simp [hnb2] @@ -319,23 +330,22 @@ lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [Decidabl rw [← hb] exact ha1 - - lemma insertionSortEquiv_orderedInsert_append {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] [IsTrans α r] (a a2 : α) : (l1 l2 : List α) → (insertionSortEquiv r (List.orderedInsert r a l1 ++ a2 :: l2) ⟨l1.length + 1, by simp⟩) = (finCongr (by simp; omega)) - ((insertionSortEquiv r ( a :: l1 ++ a2 :: l2)) ⟨l1.length + 1, by simp⟩) + ((insertionSortEquiv r (a :: l1 ++ a2 :: l2)) ⟨l1.length + 1, by simp⟩) | [], l2 => by simp - | b :: l1, l2 => by + | b :: l1, l2 => by by_cases h : r a b · have h1 : (List.orderedInsert r a (b :: l1) ++ a2 :: l2) = (a :: b :: l1 ++ a2 :: l2) := by simp [h] rw [insertionSortEquiv_congr _ _ h1] simp - · have h1 : (List.orderedInsert r a (b :: l1) ++ a2 :: l2) = (b :: List.orderedInsert r a (l1) ++ a2 :: l2) := by + · have h1 : (List.orderedInsert r a (b :: l1) ++ a2 :: l2) = + (b :: List.orderedInsert r a (l1) ++ a2 :: l2) := by simp [h] rw [insertionSortEquiv_congr _ _ h1] simp @@ -346,36 +356,34 @@ lemma insertionSortEquiv_orderedInsert_append {α : Type} (r : α → α → Pro exact insertionSort_orderedInsert_append r a l1 (a2 :: l2) rw [orderedInsertEquiv_congr _ _ _ hl] simp - change Fin.cast _ ((insertionSortEquiv r (b :: a :: (l1 ++ a2 :: l2))) ⟨l1.length + 2, by simp⟩) = _ + change Fin.cast _ + ((insertionSortEquiv r (b :: a :: (l1 ++ a2 :: l2))) ⟨l1.length + 2, by simp⟩) = _ have hl : l1.length + 1 +1 = l1.length + 2 := by omega simp [hl] conv_rhs => erw [insertionSortEquiv_commute _ _ _ h _ _] simp - - - lemma insertionSortEquiv_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] [IsTrans α r] (a : α) : (l1 l2 : List α) → (insertionSortEquiv r (List.insertionSort r l1 ++ a :: l2) ⟨l1.length, by simp⟩) = finCongr (by simp) (insertionSortEquiv r (l1 ++ a :: l2) ⟨l1.length, by simp⟩) | [], l2 => by simp - | b :: l1, l2 => by + | b :: l1, l2 => by simp have hl := insertionSortEquiv_orderedInsert_append r b a (List.insertionSort r l1) l2 simp at hl rw [hl] - have ih := insertionSortEquiv_insertionSort_append r a l1 l2 + have ih := insertionSortEquiv_insertionSort_append r a l1 l2 simp [insertionSortEquiv] rw [ih] - have hl : (List.insertionSort r (List.insertionSort r l1 ++ a :: l2)) = (List.insertionSort r (l1 ++ a :: l2)) := by + have hl : (List.insertionSort r (List.insertionSort r l1 ++ a :: l2)) = + (List.insertionSort r (l1 ++ a :: l2)) := by exact insertionSort_insertionSort_append r l1 (a :: l2) rw [orderedInsertEquiv_congr _ _ _ hl] simp - /-! ## Insertion sort with equal fields @@ -389,7 +397,7 @@ lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [Decidabl | [], hl => by simp exact h - | b :: l, hl => by + | b :: l, hl => by simp by_cases hpb : p b <;> by_cases hab : r a b · simp [hpb, hab] @@ -405,7 +413,8 @@ lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [Decidabl rw [List.filter_cons_of_pos (by simp [h]), List.filter_cons_of_neg (by simp [hpb])] rw [List.orderedInsert_eq_take_drop] - have hl : List.takeWhile (fun b => decide ¬r a b) (List.filter (fun b => decide (p b)) l) = [] := by + have hl : List.takeWhile (fun b => decide ¬r a b) + (List.filter (fun b => decide (p b)) l) = [] := by rw [List.takeWhile_eq_nil_iff] intro c hc simp at hc @@ -413,13 +422,15 @@ lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [Decidabl apply IsTrans.trans a b _ hab simp at hl apply hl.1 - have hlf : (List.filter (fun b => decide (p b)) l)[0] ∈ (List.filter (fun b => decide (p b)) l) := by + have hlf : (List.filter (fun b => decide (p b)) l)[0] ∈ + (List.filter (fun b => decide (p b)) l) := by exact List.getElem_mem c - simp [- List.getElem_mem] at hlf + simp [- List.getElem_mem] at hlf exact hlf.1 rw [hl] simp - conv_lhs => rw [← List.takeWhile_append_dropWhile (fun b => decide ¬r a b) (List.filter (fun b => decide (p b)) l )] + conv_lhs => rw [← List.takeWhile_append_dropWhile (fun b => decide ¬r a b) + (List.filter (fun b => decide (p b)) l)] rw [hl] simp · simp [hab] @@ -439,10 +450,8 @@ lemma orderedInsert_filter_of_neg {α : Type} (r : α → α → Prop) [Decidabl exact List.takeWhile_append_dropWhile (fun b => !decide (r a b)) l simp [h] - - lemma insertionSort_filter {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] - [IsTrans α r] (p : α → Prop) [DecidablePred p] : (l : List α) → + [IsTrans α r] (p : α → Prop) [DecidablePred p] : (l : List α) → List.insertionSort r (List.filter p l) = List.filter p (List.insertionSort r l) | [] => by simp @@ -468,7 +477,7 @@ lemma takeWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [Decidable | [], _ => by simp | b :: l, hl => by simp at hl - by_cases hb : ¬ r a b + by_cases hb : ¬ r a b · simp [hb] simpa using takeWhile_sorted_eq_filter r a l hl.2 · simp_all @@ -477,12 +486,12 @@ lemma takeWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [Decidable exact hl.1 c hc lemma dropWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → - List.dropWhile (fun c => ¬ r a c) l = List.filter (fun c => r a c) l + [IsTotal α r] [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → + List.dropWhile (fun c => ¬ r a c) l = List.filter (fun c => r a c) l | [], _ => by simp | b :: l, hl => by simp at hl - by_cases hb : ¬ r a b + by_cases hb : ¬ r a b · simp [hb] simpa using dropWhile_sorted_eq_filter r a l hl.2 · simp_all @@ -494,14 +503,14 @@ lemma dropWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [Decidable exact hl.1 c hc lemma dropWhile_sorted_eq_filter_filter {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) :(l : List α) → (hl : l.Sorted r) → - List.filter (fun c => r a c) l = + [IsTotal α r] [IsTrans α r] (a : α) :(l : List α) → (hl : l.Sorted r) → + List.filter (fun c => r a c) l = List.filter (fun c => r a c ∧ r c a) l ++ List.filter (fun c => r a c ∧ ¬ r c a) l | [], _ => by simp | b :: l, hl => by simp at hl - by_cases hb : ¬ r a b + by_cases hb : ¬ r a b · simp [hb] simpa using dropWhile_sorted_eq_filter_filter r a l hl.2 · simp_all @@ -529,25 +538,26 @@ lemma dropWhile_sorted_eq_filter_filter {α : Type} (r : α → α → Prop) [De exact hl.2 lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) :(l : List α) → + [IsTotal α r] [IsTrans α r] (a : α) : (l : List α) → List.filter (fun c => r a c ∧ r c a) (l.insertionSort r) = List.filter (fun c => r a c ∧ r c a) l | [] => by simp | b :: l => by - simp only [ List.insertionSort] + simp only [List.insertionSort] by_cases h : r a b ∧ r b a - · have hl := orderedInsert_filter_of_pos r b (fun c => r a c ∧ r c a) h (List.insertionSort r l) - (by exact List.sorted_insertionSort r l) + · have hl := orderedInsert_filter_of_pos r b (fun c => r a c ∧ r c a) h + (List.insertionSort r l) (by exact List.sorted_insertionSort r l) simp at hl ⊢ rw [hl] rw [List.orderedInsert_eq_take_drop] have ht : List.takeWhile (fun b_1 => decide ¬r b b_1) - (List.filter (fun b => decide (r a b) && decide (r b a)) (List.insertionSort r l)) = [] := by + (List.filter (fun b => decide (r a b) && decide (r b a)) + (List.insertionSort r l)) = [] := by rw [List.takeWhile_eq_nil_iff] intro hl simp have hx := List.getElem_mem hl - simp [- List.getElem_mem] at hx + simp [- List.getElem_mem] at hx apply IsTrans.trans b a _ h.2 simp_all rw [ht] @@ -557,13 +567,14 @@ lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [Decidabl have ih := filter_rel_eq_insertionSort r a l simp at ih rw [← ih] - have htd := List.takeWhile_append_dropWhile (fun b_1 => decide ¬r b b_1) (List.filter (fun b => decide (r a b) && decide (r b a)) (List.insertionSort r l)) + have htd := List.takeWhile_append_dropWhile (fun b_1 => decide ¬r b b_1) + (List.filter (fun b => decide (r a b) && decide (r b a)) (List.insertionSort r l)) simp [decide_not, - List.takeWhile_append_dropWhile] at htd conv_rhs => rw [← htd] simp [- List.takeWhile_append_dropWhile] intro hl have hx := List.getElem_mem hl - simp [- List.getElem_mem] at hx + simp [- List.getElem_mem] at hx apply IsTrans.trans b a _ h.2 simp_all simp_all @@ -575,9 +586,8 @@ lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [Decidabl simp_all simpa using h - lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) (l1 l l2 : List α) + [IsTotal α r] [IsTrans α r] (a : α) (l1 l l2 : List α) (h : ∀ b ∈ l, r a b ∧ r b a) : List.insertionSort r (l1 ++ l ++ l2) = (List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r)) @@ -586,14 +596,14 @@ lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRe ++ (List.filter (fun c => r a c ∧ r c a) l2) ++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) := by - have hl : List.insertionSort r (l1 ++ l ++ l2) = + have hl : List.insertionSort r (l1 ++ l ++ l2) = List.takeWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) ++ List.dropWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) := by exact (List.takeWhile_append_dropWhile (fun c => decide ¬r a c) (List.insertionSort r (l1 ++ l ++ l2))).symm have hlt : List.takeWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) = List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r) := by - rw [takeWhile_sorted_eq_filter, takeWhile_sorted_eq_filter ] + rw [takeWhile_sorted_eq_filter, takeWhile_sorted_eq_filter] rw [← insertionSort_filter, ← insertionSort_filter] congr 1 simp @@ -601,7 +611,7 @@ lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRe exact List.sorted_insertionSort r (l1 ++ l2) exact List.sorted_insertionSort r (l1 ++ l ++ l2) conv_lhs => rw [hl, hlt] - simp only [decide_not, Bool.decide_and] + simp only [decide_not, Bool.decide_and] simp have h1 := dropWhile_sorted_eq_filter r a (List.insertionSort r (l1 ++ (l ++ l2))) simp at h1 @@ -615,13 +625,13 @@ lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRe have h1 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ (l ++ l2)) simp at h1 rw [← h1] - have h2 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ l2) + have h2 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ l2) simp at h2 rw [← h2] congr have hl : List.filter (fun b => decide (r a b) && !decide (r b a)) l = [] := by rw [@List.filter_eq_nil_iff] - intro c hc + intro c hc simp_all rw [hl] simp @@ -629,7 +639,7 @@ lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRe exact List.sorted_insertionSort r (l1 ++ (l ++ l2)) lemma insertionSort_of_takeWhile_filter {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) (l1 l2 : List α) : + [IsTotal α r] [IsTrans α r] (a : α) (l1 l2 : List α) : List.insertionSort r (l1 ++ l2) = (List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r)) ++ (List.filter (fun c => r a c ∧ r c a) l1) @@ -642,5 +652,4 @@ lemma insertionSort_of_takeWhile_filter {α : Type} (r : α → α → Prop) [De simp simp - end HepLean.List diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean index 00c4f43c..5ee7578f 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean @@ -21,12 +21,12 @@ namespace CrAnAlgebra noncomputable section /-- The submodule of `CrAnAlgebra` spanned by lists of field statistic `f`. -/ -def statisticSubmodule (f : FieldStatistic) : Submodule ℂ 𝓕.CrAnAlgebra := +def statisticSubmodule (f : FieldStatistic) : Submodule ℂ 𝓕.CrAnAlgebra := Submodule.span ℂ {a | ∃ φs, a = ofCrAnList φs ∧ (𝓕 |>ₛ φs) = f} lemma ofCrAnList_mem_statisticSubmodule_of (φs : List 𝓕.CrAnStates) (f : FieldStatistic) (h : (𝓕 |>ₛ φs) = f) : - ofCrAnList φs ∈ statisticSubmodule f := by + ofCrAnList φs ∈ statisticSubmodule f := by refine Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩ lemma ofCrAnList_bosonic_or_fermionic (φs : List 𝓕.CrAnStates) : @@ -119,7 +119,7 @@ lemma fermionicProj_ofCrAnList (φs : List 𝓕.CrAnStates) : lemma fermionicProj_ofCrAnList_if_bosonic (φs : List 𝓕.CrAnStates) : fermionicProj (ofCrAnList φs) = if h : (𝓕 |>ₛ φs) = bosonic then 0 else ⟨ofCrAnList φs, Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, - by simpa using h ⟩⟩⟩ := by + by simpa using h⟩⟩⟩ := by rw [fermionicProj_ofCrAnList] by_cases h1 : (𝓕 |>ₛ φs) = fermionic · simp [h1] @@ -188,7 +188,6 @@ lemma bosonicProj_add_fermionicProj (a : 𝓕.CrAnAlgebra) : · simp [h] · simp [h] - lemma coeAddMonoidHom_apply_eq_bosonic_plus_fermionic (a : DirectSum FieldStatistic (fun i => (statisticSubmodule (𝓕 := 𝓕) i))) : DirectSum.coeAddMonoidHom statisticSubmodule a = a.1 bosonic + a.1 fermionic := by @@ -245,7 +244,7 @@ instance crAnAlgebraGrade : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmo rfl mul_mem f1 f2 a1 a2 h1 h2 := by let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop := - a1 * a2 ∈ statisticSubmodule (f1 + f2) + a1 * a2 ∈ statisticSubmodule (f1 + f2) change p a2 h2 apply Submodule.span_induction (p := p) · intro x hx @@ -281,10 +280,10 @@ instance crAnAlgebraGrade : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmo simp only [Algebra.mul_smul_comm, p] exact Submodule.smul_mem _ _ h1 · exact h2 - decompose' a := DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) bosonic (bosonicProj a) - + DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) fermionic (fermionicProj a) + decompose' a := DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) bosonic (bosonicProj a) + + DirectSum.of (fun i => (statisticSubmodule (𝓕 := 𝓕) i)) fermionic (fermionicProj a) left_inv a := by - trans a.bosonicProj + fermionicProj a + trans a.bosonicProj + fermionicProj a · simp · exact bosonicProj_add_fermionicProj a right_inv a := by diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean index f5ceffce..c15e41bf 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean @@ -441,38 +441,38 @@ lemma superCommute_ofCrAnList_ofStateList_eq_sum (φs : List 𝓕.CrAnStates) : FieldStatistic.ofList_cons_eq_mul, mul_comm] lemma summerCommute_jacobi_ofCrAnList (φs1 φs2 φs3 : List 𝓕.CrAnStates) : - [ofCrAnList φs1, [ofCrAnList φs2, ofCrAnList φs3]ₛca]ₛca = - 𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs3) • - (- 𝓢(𝓕 |>ₛ φ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 - 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] - abel - · simp [h1, h2, exchangeSign_bosonic, bosonic_exchangeSign, mul_one, one_smul] - abel - · simp [h1, bosonic_exchangeSign, h3, exchangeSign_bosonic, mul_one, one_smul] - abel - · simp at h1 h2 h3 - simp [h1, h2, h3] - abel - · simp at h1 h2 h3 - simp [h1, h2, h3] - abel - · simp at h1 h2 h3 - simp [h1, h2, h3] - abel - · simp at h1 h2 h3 - simp [h1, h2, h3] - abel - · simp at h1 h2 h3 - simp [h1, h2, h3] - abel + [ofCrAnList φs1, [ofCrAnList φs2, ofCrAnList φs3]ₛca]ₛca = + 𝓢(𝓕 |>ₛ φs1, 𝓕 |>ₛ φs3) • + (- 𝓢(𝓕 |>ₛ φ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 + 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] + abel + · simp [h1, h2, exchangeSign_bosonic, bosonic_exchangeSign, mul_one, one_smul] + abel + · simp [h1, bosonic_exchangeSign, h3, exchangeSign_bosonic, mul_one, one_smul] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel + · simp at h1 h2 h3 + simp [h1, h2, h3] + abel /-! ## Interaction with grading. @@ -483,14 +483,14 @@ lemma superCommute_grade {a b : 𝓕.CrAnAlgebra} {f1 f2 : FieldStatistic} (ha : a ∈ statisticSubmodule f1) (hb : b ∈ statisticSubmodule f2) : [a, b]ₛca ∈ statisticSubmodule (f1 + f2) := by let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f2) : Prop := - [a, a2]ₛca ∈ statisticSubmodule (f1 + f2) + [a, a2]ₛca ∈ statisticSubmodule (f1 + f2) change p b hb apply Submodule.span_induction (p := p) · intro x hx obtain ⟨φs, rfl, hφs⟩ := hx simp [p] let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule f1) : Prop := - [a2 , ofCrAnList φs]ₛca ∈ statisticSubmodule (f1 + f2) + [a2, ofCrAnList φs]ₛca ∈ statisticSubmodule (f1 + f2) change p a ha apply Submodule.span_induction (p := p) · intro x hx @@ -525,13 +525,13 @@ lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule bosonic) : [a, b]ₛca = a * b - b * a := by let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := - [a, a2]ₛca = a * a2 - a2 * a + [a, a2]ₛca = a * a2 - a2 * a change p b hb apply Submodule.span_induction (p := p) · intro x hx obtain ⟨φs, rfl, hφs⟩ := hx let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := - [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 + [a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 change p a ha apply Submodule.span_induction (p := p) · intro x hx @@ -554,18 +554,17 @@ lemma superCommute_bosonic_bosonic {a b : 𝓕.CrAnAlgebra} simp_all [p, smul_sub] · exact hb - lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) (hb : b ∈ statisticSubmodule fermionic) : [a, b]ₛca = a * b - b * a := by let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := - [a, a2]ₛca = a * a2 - a2 * a + [a, a2]ₛca = a * a2 - a2 * a change p b hb apply Submodule.span_induction (p := p) · intro x hx obtain ⟨φs, rfl, hφs⟩ := hx let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := - [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 + [a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 change p a ha apply Submodule.span_induction (p := p) · intro x hx @@ -588,18 +587,17 @@ lemma superCommute_bosonic_fermionic {a b : 𝓕.CrAnAlgebra} simp_all [p, smul_sub] · exact hb - lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule bosonic) : [a, b]ₛca = a * b - b * a := by let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule bosonic) : Prop := - [a, a2]ₛca = a * a2 - a2 * a + [a, a2]ₛca = a * a2 - a2 * a change p b hb apply Submodule.span_induction (p := p) · intro x hx obtain ⟨φs, rfl, hφs⟩ := hx let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := - [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 + [a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs - ofCrAnList φs * a2 change p a ha apply Submodule.span_induction (p := p) · intro x hx @@ -622,7 +620,7 @@ lemma superCommute_fermionic_bonsonic {a b : 𝓕.CrAnAlgebra} simp_all [p, smul_sub] · exact hb -lemma superCommute_bonsonic {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) : +lemma superCommute_bonsonic {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) : [a, b]ₛca = a * b - b * a := by rw [← bosonicProj_add_fermionicProj a] simp only [map_add, LinearMap.add_apply] @@ -630,7 +628,7 @@ lemma superCommute_bonsonic {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmo simp only [add_mul, mul_add] abel -lemma bosonic_superCommute {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) : +lemma bosonic_superCommute {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) : [a, b]ₛca = a * b - b * a := by rw [← bosonicProj_add_fermionicProj b] simp only [map_add, LinearMap.add_apply] @@ -638,12 +636,12 @@ lemma bosonic_superCommute {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmod simp only [add_mul, mul_add] abel -lemma superCommute_bonsonic_symm {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) : +lemma superCommute_bonsonic_symm {a b : 𝓕.CrAnAlgebra} (hb : b ∈ statisticSubmodule bosonic) : [a, b]ₛca = - [b, a]ₛca := by rw [bosonic_superCommute hb, superCommute_bonsonic hb] simp -lemma bonsonic_superCommute_symm {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) : +lemma bonsonic_superCommute_symm {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule bosonic) : [a, b]ₛca = - [b, a]ₛca := by rw [bosonic_superCommute ha, superCommute_bonsonic ha] simp @@ -652,13 +650,13 @@ lemma superCommute_fermionic_fermionic {a b : 𝓕.CrAnAlgebra} (ha : a ∈ statisticSubmodule fermionic) (hb : b ∈ statisticSubmodule fermionic) : [a, b]ₛca = a * b + b * a := by let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := - [a, a2]ₛca = a * a2 + a2 * a + [a, a2]ₛca = a * a2 + a2 * a change p b hb apply Submodule.span_induction (p := p) · intro x hx obtain ⟨φs, rfl, hφs⟩ := hx let p (a2 : 𝓕.CrAnAlgebra) (hx : a2 ∈ statisticSubmodule fermionic) : Prop := - [a2 , ofCrAnList φs]ₛca = a2 * ofCrAnList φs + ofCrAnList φs * a2 + [a2, ofCrAnList φs]ₛca = a2 * ofCrAnList φs + ofCrAnList φs * a2 change p a ha apply Submodule.span_induction (p := p) · intro x hx @@ -701,7 +699,7 @@ lemma superCommute_expand_bosonicProj_fermionicProj (a b : 𝓕.CrAnAlgebra) : abel lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List 𝓕.CrAnStates) : - [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule bosonic ∨ + [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule bosonic ∨ [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic := by by_cases h1 : (𝓕 |>ₛ φs) = bosonic <;> by_cases h2 : (𝓕 |>ₛ φs') = bosonic · left @@ -738,7 +736,7 @@ lemma superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic (φs φs' : List apply ofCrAnList_mem_statisticSubmodule_of _ _ (by simpa using h2) lemma superCommute_ofCrAnState_ofCrAnState_bosonic_or_fermionic (φ φ' : 𝓕.CrAnStates) : - [ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule bosonic ∨ + [ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule bosonic ∨ [ofCrAnState φ, ofCrAnState φ']ₛca ∈ statisticSubmodule fermionic := by rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] exact superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [φ'] @@ -779,7 +777,7 @@ lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * ofCrAnList (φs.drop (n + 1)) := by let p (a : 𝓕.CrAnAlgebra) (ha : a ∈ statisticSubmodule bosonic) : Prop := - [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), + [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * ofCrAnList (φs.drop (n + 1)) change p a ha @@ -802,14 +800,13 @@ lemma superCommute_bosonic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List simp_all [p, Finset.smul_sum] · exact ha - lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : List 𝓕.CrAnStates) (ha : a ∈ statisticSubmodule fermionic) : - [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) • + [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) • ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * ofCrAnList (φs.drop (n + 1)) := by let p (a : 𝓕.CrAnAlgebra) (ha : a ∈ statisticSubmodule fermionic) : Prop := - [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) • + [a, ofCrAnList φs]ₛca = ∑ (n : Fin φs.length), 𝓢(fermionic, 𝓕 |>ₛ φs.take n) • ofCrAnList (φs.take n) * [a, ofCrAnState (φs.get n)]ₛca * ofCrAnList (φs.drop (n + 1)) change p a ha @@ -835,7 +832,6 @@ lemma superCommute_fermionic_ofCrAnList_eq_sum (a : 𝓕.CrAnAlgebra) (φs : Lis simp [smul_smul, mul_comm] · exact ha - lemma statistic_neq_of_superCommute_fermionic {φs φs' : List 𝓕.CrAnStates} (h : [ofCrAnList φs, ofCrAnList φs']ₛca ∈ statisticSubmodule fermionic) : (𝓕 |>ₛ φs) ≠ (𝓕 |>ₛ φs') ∨ [ofCrAnList φs, ofCrAnList φs']ₛca = 0 := by diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean index 29f12d30..92b0b739 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean @@ -66,7 +66,7 @@ lemma timeOrder_timeOrder_mid (a b c : 𝓕.CrAnAlgebra) : 𝓣ᶠ(a * b * c) = Algebra.smul_mul_assoc, map_smul] rw [timeOrder_ofCrAnList, timeOrder_ofCrAnList, smul_smul] congr 1 - · simp only [crAnTimeOrderSign, crAnTimeOrderList] + · simp only [crAnTimeOrderSign, crAnTimeOrderList] rw [Wick.koszulSign_of_append_eq_insertionSort, mul_comm] · congr 1 simp only [crAnTimeOrderList] @@ -167,7 +167,6 @@ lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_right timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel h] simp - lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_left {φ ψ : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ ψ) (a : 𝓕.CrAnAlgebra) : 𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca * a) = 0 := by @@ -183,7 +182,7 @@ lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel_mid simp lemma timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel - {φ1 φ2 : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ1 φ2) (a : 𝓕.CrAnAlgebra): + {φ1 φ2 : 𝓕.CrAnStates} (h : ¬ crAnTimeOrderRel φ1 φ2) (a : 𝓕.CrAnAlgebra) : 𝓣ᶠ([a, [ofCrAnState φ1, ofCrAnState φ2]ₛca]ₛca) = 0 := by rw [← bosonicProj_add_fermionicProj a] simp @@ -213,7 +212,7 @@ lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel rw [summerCommute_jacobi_ofCrAnList] simp [ofCrAnList_singleton] right - rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12 ] + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12] rw [superCommute_ofCrAnState_ofCrAnState_symm φ3] simp rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13] @@ -229,14 +228,14 @@ lemma timeOrder_superCommute_ofCrAnState_superCommute_not_crAnTimeOrderRel' right rw [superCommute_ofCrAnState_ofCrAnState_symm φ1] simp - rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12 ] + rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h12] simp rw [timeOrder_superCommute_superCommute_ofCrAnState_not_crAnTimeOrderRel h13] simp lemma timeOrder_superCommute_ofCrAnState_superCommute_all_not_crAnTimeOrderRel - (φ1 φ2 φ3 : 𝓕.CrAnStates) (h : ¬ ( - crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧ + (φ1 φ2 φ3 : 𝓕.CrAnStates) (h : ¬ + (crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧ crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧ crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2)) : 𝓣ᶠ([ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca) = 0 := by @@ -272,7 +271,6 @@ lemma timeOrder_superCommute_ofCrAnState_superCommute_all_not_crAnTimeOrderRel refine False.elim (h ?_) exact IsTrans.trans φ3 φ2 φ1 h32 h21 - lemma timeOrder_superCommute_ofCrAnState_ofCrAnState_eq_time {φ ψ : 𝓕.CrAnStates} (h1 : crAnTimeOrderRel φ ψ) (h2 : crAnTimeOrderRel ψ φ) : 𝓣ᶠ([ofCrAnState φ, ofCrAnState ψ]ₛca) = [ofCrAnState φ, ofCrAnState ψ]ₛca := by diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean index 10993f0c..6dae02a8 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean @@ -112,8 +112,8 @@ lemma ι_superCommute_zero_of_fermionic (φ ψ : 𝓕.CrAnStates) · simp [h] lemma ι_superCommute_ofCrAnState_ofCrAnState_bosonic_or_zero (φ ψ : 𝓕.CrAnStates) : - [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ statisticSubmodule bosonic ∨ - ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by + [ofCrAnState φ, ofCrAnState ψ]ₛca ∈ statisticSubmodule bosonic ∨ + ι [ofCrAnState φ, ofCrAnState ψ]ₛca = 0 := by rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ] [ψ] with h | h · simp_all [ofCrAnList_singleton] · simp_all [ofCrAnList_singleton] @@ -236,7 +236,6 @@ lemma bosonicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x ∈ · right rw [bosonicProj_of_mem_fermionic _ h] - lemma fermionicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x ∈ 𝓕.fieldOpIdealSet) : x.fermionicProj.1 ∈ 𝓕.fieldOpIdealSet ∨ x.fermionicProj = 0 := by have hx' := hx @@ -269,7 +268,7 @@ lemma fermionicProj_mem_fieldOpIdealSet_or_zero (x : CrAnAlgebra 𝓕) (hx : x simpa using hx' lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : - x.bosonicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by + x.bosonicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by rw [TwoSidedIdeal.mem_span_iff_mem_addSubgroup_closure] at hx let p {k : Set 𝓕.CrAnAlgebra} (a : CrAnAlgebra 𝓕) (h : a ∈ AddSubgroup.closure k) : Prop := a.bosonicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet @@ -405,7 +404,7 @@ lemma bosonicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.spa simp [p] lemma fermionicProj_mem_ideal (x : CrAnAlgebra 𝓕) (hx : x ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet) : - x.fermionicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by + x.fermionicProj.1 ∈ TwoSidedIdeal.span 𝓕.fieldOpIdealSet := by have hb := bosonicProj_mem_ideal x hx rw [← ι_eq_zero_iff_mem_ideal] at hx hb ⊢ rw [← bosonicProj_add_fermionicProj x] at hx @@ -425,7 +424,5 @@ lemma ι_eq_zero_iff_ι_bosonicProj_fermonicProj_zero (x : CrAnAlgebra 𝓕) : rw [← bosonicProj_add_fermionicProj x] simp_all - - end FieldOpAlgebra end FieldSpecification diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean index e5559773..d0b2c1a6 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean @@ -80,7 +80,7 @@ lemma superCommuteRight_eq_of_equiv (a1 a2 : 𝓕.CrAnAlgebra) (h : a1 ≈ a2) : apply ι_superCommute_eq_zero_of_ι_left_zero exact (ι_eq_zero_iff_mem_ideal (a1 - a2)).mpr h simp_all [superCommuteRight_apply_ι] - trans ι ((superCommute a2) b) + 0 + trans ι ((superCommute a2) b) + 0 rw [← ha1b1] simp simp diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean index 16b1e76f..3863ccb3 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean @@ -20,7 +20,7 @@ namespace FieldOpAlgebra variable {𝓕 : FieldSpecification} lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates} - (φs1 φs2 : List 𝓕.CrAnStates) (h : + (φs1 φs2 : List 𝓕.CrAnStates) (h : crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧ crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧ crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2): @@ -115,12 +115,12 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : rw [← smul_sub, ← smul_sub, smul_smul, mul_comm, ← smul_smul, ← smul_sub] simp right - rw [← smul_mul_assoc, ← mul_smul_comm, mul_assoc] + rw [← smul_mul_assoc, ← mul_smul_comm, mul_assoc] - rw [← smul_mul_assoc, ← mul_smul_comm] + rw [← smul_mul_assoc, ← mul_smul_comm] rw [smul_sub] - rw [← smul_mul_assoc, ← mul_smul_comm] - rw [← smul_mul_assoc, ← mul_smul_comm] + rw [← smul_mul_assoc, ← mul_smul_comm] + rw [← smul_mul_assoc, ← mul_smul_comm] repeat rw [mul_assoc] rw [← mul_sub, ← mul_sub, ← mul_sub] rw [← sub_mul, ← sub_mul, ← sub_mul] @@ -181,7 +181,7 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates} ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a * b)) := by let pb (b : 𝓕.CrAnAlgebra) (hc : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) : - Prop := ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = + Prop := ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = ι ([ofCrAnState φ, ofCrAnState ψ]ₛca * 𝓣ᶠ(a * b)) change pb b (Basis.mem_span _ b) apply Submodule.span_induction @@ -201,7 +201,7 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates} simp [mul_sub, sub_mul, ← ofCrAnList_append] rw [timeOrder_ofCrAnList, timeOrder_ofCrAnList] have h1 : crAnTimeOrderSign (φs' ++ φ :: ψ :: φs) = crAnTimeOrderSign (φs' ++ ψ :: φ :: φs) := by - trans crAnTimeOrderSign (φs' ++ [φ, ψ] ++ φs) + trans crAnTimeOrderSign (φs' ++ [φ, ψ] ++ φs) simp rw [crAnTimeOrderSign] have hp : List.Perm [φ,ψ] [ψ,φ] := by exact List.Perm.swap ψ φ [] @@ -244,7 +244,7 @@ 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 at h1 ⊢ rw [← h1] rw [← crAnTimeOrderList] by_cases hq : (𝓕 |>ₛ φ) ≠ (𝓕 |>ₛ ψ) @@ -269,7 +269,7 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates} lemma ι_timeOrder_superCommute_neq_time {φ ψ : 𝓕.CrAnStates} - (hφψ : ¬ (crAnTimeOrderRel φ ψ ∧ crAnTimeOrderRel ψ φ)) (a b : 𝓕.CrAnAlgebra) : + (hφψ : ¬ (crAnTimeOrderRel φ ψ ∧ crAnTimeOrderRel ψ φ)) (a b : 𝓕.CrAnAlgebra) : ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = 0 := by rw [timeOrder_timeOrder_mid] have hφψ : ¬ (crAnTimeOrderRel φ ψ) ∨ ¬ (crAnTimeOrderRel ψ φ) := by @@ -309,7 +309,7 @@ lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra) simp | Or.inr (Or.inl hc) => obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc - by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) + by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) · rw [ι_timeOrder_superCommute_eq_time] simp rw [ι_superCommute_of_create_create] @@ -321,7 +321,7 @@ lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra) · rw [ι_timeOrder_superCommute_neq_time heqt] | Or.inr (Or.inr (Or.inl hc)) => obtain ⟨φa, hφa, φb, hφb, rfl⟩ := hc - by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) + by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) · rw [ι_timeOrder_superCommute_eq_time] simp rw [ι_superCommute_of_annihilate_annihilate] @@ -333,7 +333,7 @@ lemma ι_timeOrder_zero_of_mem_ideal (a : 𝓕.CrAnAlgebra) · rw [ι_timeOrder_superCommute_neq_time heqt] | Or.inr (Or.inr (Or.inr hc)) => obtain ⟨φa, φb, hdiff, rfl⟩ := hc - by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) + by_cases heqt : (crAnTimeOrderRel φa φb ∧ crAnTimeOrderRel φb φa) · rw [ι_timeOrder_superCommute_eq_time] simp rw [ι_superCommute_of_diff_statistic] diff --git a/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean b/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean index b83dd38c..53ca3a7b 100644 --- a/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean +++ b/HepLean/PerturbationTheory/FieldStatistics/ExchangeSign.lean @@ -72,7 +72,7 @@ lemma bosonic_exchangeSign (a : FieldStatistic) : 𝓢(bosonic, a) = 1 := by rw [exchangeSign_symm, exchangeSign_bosonic] @[simp] -lemma fermionic_exchangeSign_fermionic : 𝓢(fermionic, fermionic) = - 1 := by +lemma fermionic_exchangeSign_fermionic : 𝓢(fermionic, fermionic) = - 1 := by rfl lemma exchangeSign_eq_if (a b : FieldStatistic) : diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean index 0f3205e1..928902be 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean @@ -261,7 +261,7 @@ lemma koszulSign_eraseIdx_insertionSortMinPos [IsTotal 𝓕 le] [IsTrans 𝓕 le rfl lemma koszulSign_swap_eq_rel_cons {ψ φ : 𝓕} - (h1 : le φ ψ) (h2 : le ψ φ) (φs' : List 𝓕): + (h1 : le φ ψ) (h2 : le ψ φ) (φs' : List 𝓕) : koszulSign q le (φ :: ψ :: φs') = koszulSign q le (ψ :: φ :: φs') := by simp only [Wick.koszulSign, ← mul_assoc, mul_eq_mul_right_iff] left @@ -285,7 +285,7 @@ lemma koszulSign_eq_rel_eq_stat_append {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal koszulSign q le (φ :: ψ :: φs) = koszulSign q le φs := by intro φs simp [koszulSign, ← mul_assoc] - trans 1 * koszulSign q le φs + trans 1 * koszulSign q le φs swap simp congr @@ -305,11 +305,11 @@ lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal 𝓕 l rw [koszulSign_eq_rel_eq_stat h1 h2 hq φs' φs] simp left - trans koszulSignInsert q le φ'' (φ :: ψ :: (φs' ++ φs) ) + 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 ] + rw [koszulSignInsert_eq_remove_same_stat_append q le] simp_all simp_all simp_all @@ -331,15 +331,16 @@ lemma koszulSign_of_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φs : Lis apply koszulSign_of_sorted exact List.sorted_insertionSort le φs -lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕 le] : (φs φs' : List 𝓕) → - koszulSign q le (φs ++ φs') = - koszulSign q le (List.insertionSort le φs ++ φs') * koszulSign q le φs +lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕 le] : + (φs φs' : List 𝓕) → koszulSign q le (φs ++ φs') = + koszulSign q le (List.insertionSort le φs ++ φs') * koszulSign q le φs | φs, [] => by simp | φs, φ :: φs' => by have h1 : (φs ++ φ :: φs') = List.insertIdx φs.length φ (φs ++ φs') := by rw [insertIdx_length_fst_append] - have h2 : (List.insertionSort le φs ++ φ :: φs') = List.insertIdx (List.insertionSort le φs).length φ (List.insertionSort le φs ++ φs') := by + have h2 : (List.insertionSort le φs ++ φ :: φs') = + List.insertIdx (List.insertionSort le φs).length φ (List.insertionSort le φs ++ φs') := by rw [insertIdx_length_fst_append] rw [h1, h2] rw [koszulSign_insertIdx] @@ -353,7 +354,8 @@ lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕 simp [mul_comm] left congr 3 - · have h2 : (List.insertionSort le φs ++ φ :: φs') = List.insertIdx φs.length φ (List.insertionSort le φs ++ φs') := by + · have h2 : (List.insertionSort le φs ++ φ :: φs') = + List.insertIdx φs.length φ (List.insertionSort le φs ++ φs') := by rw [← insertIdx_length_fst_append] simp rw [insertionSortEquiv_congr _ _ h2.symm] @@ -363,16 +365,16 @@ lemma koszulSign_of_append_eq_insertionSort_left [IsTotal 𝓕 le] [IsTrans 𝓕 rw [insertionSortEquiv_congr _ _ h1.symm] simp · rw [insertIdx_length_fst_append] - rw [show φs.length = (List.insertionSort le φs).length by simp] + rw [show φs.length = (List.insertionSort le φs).length by simp] rw [insertIdx_length_fst_append] symm apply insertionSort_insertionSort_append · simp · simp -lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] : (φs'' φs φs' : List 𝓕) → +lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] : (φs'' φs φs' : List 𝓕) → koszulSign q le (φs'' ++ φs ++ φs') = - koszulSign q le (φs'' ++ List.insertionSort le φs ++ φs') * koszulSign q le φs + koszulSign q le (φs'' ++ List.insertionSort le φs ++ φs') * koszulSign q le φs | [], φs, φs'=> by simp exact koszulSign_of_append_eq_insertionSort_left q le φs φs' @@ -391,10 +393,10 @@ lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] -/ -lemma koszulSign_perm_eq_append [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) ( φs φs' φs2 : List 𝓕) +lemma koszulSign_perm_eq_append [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) (φs φs' φs2 : List 𝓕) (hp : φs.Perm φs') : (h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) → koszulSign q le (φs ++ φs2) = koszulSign q le (φs' ++ φs2) := by - let motive (φs φs' : List 𝓕) (hp : φs.Perm φs') : Prop := + let motive (φs φs' : List 𝓕) (hp : φs.Perm φs') : Prop := (h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) → koszulSign q le (φs ++ φs2) = koszulSign q le (φs' ++ φs2) change motive φs φs' hp @@ -433,5 +435,4 @@ lemma koszulSign_perm_eq [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) : (φs1 refine (List.perm_append_right_iff φs2).mpr ?_ exact List.Perm.append_left φs1 hp - end Wick diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean index 2f22ffa2..b074142e 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean @@ -235,7 +235,7 @@ lemma koszulSignInsert_cons (r0 r1 : 𝓕) (r : List 𝓕) : koszulSignInsert q le r0 r := by simp [koszulSignInsert, koszulSignCons] -lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ b ∈ φs, le φ0 b) → +lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ b ∈ φs, le φ0 b) → koszulSignInsert q le φ0 φs = 1 | [], _ => by simp [koszulSignInsert] @@ -247,7 +247,6 @@ lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ exact h b (List.mem_cons_of_mem _ hb) · exact h φ1 (List.mem_cons_self _ _) - lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTotal 𝓕 le] [IsTrans 𝓕 le] (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) → koszulSignInsert q le φ φs = koszulSignInsert q le ψ φs @@ -270,7 +269,7 @@ lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTotal 𝓕 le] [IsTrans 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 𝓕) → + (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) → koszulSignInsert q le φ' (φ :: ψ :: φs) = koszulSignInsert q le φ' φs := by intro φs simp_all [koszulSignInsert] @@ -284,6 +283,4 @@ lemma koszulSignInsert_eq_remove_same_stat_append {ψ φ φ' : 𝓕} [IsTotal apply IsTrans.trans φ' ψ φ hφ'ψ h2 simp_all [hφ'φ, hφ'ψ] - - end Wick From 22636db6066dab779c1d7d94210c898fcc61fbe1 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Wed, 29 Jan 2025 16:41:10 +0000 Subject: [PATCH 7/8] refactor: Free simps --- .../Algebras/CrAnAlgebra/Grading.lean | 40 +++++--- .../Algebras/CrAnAlgebra/SuperCommute.lean | 97 +++++++++++-------- .../Algebras/CrAnAlgebra/TimeOrder.lean | 49 +++++----- .../Algebras/FieldOpAlgebra/Basic.lean | 38 ++++---- .../Algebras/FieldOpAlgebra/SuperCommute.lean | 8 +- .../Algebras/FieldOpAlgebra/TimeOrder.lean | 70 +++++++------ .../FieldStatistics/Basic.lean | 7 +- .../PerturbationTheory/Koszul/KoszulSign.lean | 53 +++++----- .../Koszul/KoszulSignInsert.lean | 14 +-- 9 files changed, 205 insertions(+), 171 deletions(-) diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean index 5ee7578f..812fecc1 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean @@ -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 diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean index c15e41bf..96c1eaa0 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/SuperCommute.lean @@ -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 diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean index 92b0b739..e2666471 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/TimeOrder.lean @@ -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 diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean index 6dae02a8..7fd9aba7 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean @@ -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 𝓕) : diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean index d0b2c1a6..60b2bd40 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean @@ -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 diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean index 3863ccb3..c3f4f5a1 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean @@ -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 diff --git a/HepLean/PerturbationTheory/FieldStatistics/Basic.lean b/HepLean/PerturbationTheory/FieldStatistics/Basic.lean index 4e209eff..f7595e6c 100644 --- a/HepLean/PerturbationTheory/FieldStatistics/Basic.lean +++ b/HepLean/PerturbationTheory/FieldStatistics/Basic.lean @@ -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 diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean index 928902be..2b8eaadf 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean @@ -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] diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean index b074142e..2d899df9 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean @@ -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 From 32aefb7eb74e49b423d831b3f2eae5090a5fc557 Mon Sep 17 00:00:00 2001 From: jstoobysmith <72603918+jstoobysmith@users.noreply.github.com> Date: Thu, 30 Jan 2025 05:35:42 +0000 Subject: [PATCH 8/8] refactor: Lint --- HepLean.lean | 2 + HepLean/Mathematics/List/InsertIdx.lean | 3 +- HepLean/Mathematics/List/InsertionSort.lean | 257 ++++++++++-------- .../Algebras/CrAnAlgebra/Grading.lean | 3 +- .../Algebras/FieldOpAlgebra/Basic.lean | 10 +- .../Algebras/FieldOpAlgebra/SuperCommute.lean | 16 +- .../Algebras/FieldOpAlgebra/TimeOrder.lean | 72 ++--- .../PerturbationTheory/Koszul/KoszulSign.lean | 8 +- .../Koszul/KoszulSignInsert.lean | 4 +- 9 files changed, 209 insertions(+), 166 deletions(-) diff --git a/HepLean.lean b/HepLean.lean index 1eb16e18..bde57c90 100644 --- a/HepLean.lean +++ b/HepLean.lean @@ -129,6 +129,8 @@ import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.TimeOrder import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.NormalOrder +import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.SuperCommute +import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeOrder import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.Basic import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.NormalOrder import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.TimeContraction diff --git a/HepLean/Mathematics/List/InsertIdx.lean b/HepLean/Mathematics/List/InsertIdx.lean index c26e125d..70431b8a 100644 --- a/HepLean/Mathematics/List/InsertIdx.lean +++ b/HepLean/Mathematics/List/InsertIdx.lean @@ -122,7 +122,8 @@ lemma insertIdx_length_fst_append {I : Type} (φ : I) : (φs φs' : List I) → List.insertIdx φs.length φ (φs ++ φs') = (φs ++ φ :: φs') | [], φs' => by simp | φ' :: φs, φs' => by - simp + simp only [List.length_cons, List.cons_append, List.insertIdx_succ_cons, List.cons.injEq, + true_and] exact insertIdx_length_fst_append φ φs φs' lemma get_eq_insertIdx_succAbove {I : Type} (i : I) (r : List I) (k : Fin r.length.succ) : diff --git a/HepLean/Mathematics/List/InsertionSort.lean b/HepLean/Mathematics/List/InsertionSort.lean index b979ac23..9957650d 100644 --- a/HepLean/Mathematics/List/InsertionSort.lean +++ b/HepLean/Mathematics/List/InsertionSort.lean @@ -110,19 +110,16 @@ lemma orderedInsert_commute {α : Type} (r : α → α → Prop) [DecidableRel r have hrb : r b a := by have ht := IsTotal.total (r := r) a b simp_all - simp + simp only [List.orderedInsert] by_cases h : r a c - · simp [h, hrb] + · simp only [h, ↓reduceIte, List.orderedInsert.eq_2, hrb] rw [if_pos] - simp [hrb, hr, h] + simp only [List.orderedInsert, hr, ↓reduceIte, h] exact IsTrans.trans (r :=r) _ _ _ hrb h - · simp [h] - have hrca : r c a := by - have ht := IsTotal.total (r := r) a c - simp_all + · simp only [h, ↓reduceIte, List.orderedInsert.eq_2] by_cases hbc : r b c · simp [hbc, hr, h] - · simp [hbc, h] + · simp only [hbc, ↓reduceIte, List.orderedInsert.eq_2, h, List.cons.injEq, true_and] exact orderedInsert_commute r a b hr l lemma insertionSort_orderedInsert_append {α : Type} (r : α → α → Prop) [DecidableRel r] @@ -136,7 +133,7 @@ lemma insertionSort_orderedInsert_append {α : Type} (r : α → α → Prop) [D · simp [h] conv_lhs => simp [h] rw [insertionSort_orderedInsert_append r a l1 l2] - simp + simp only [List.insertionSort, List.append_eq] rw [orderedInsert_commute r a b h] lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] @@ -147,7 +144,7 @@ lemma insertionSort_insertionSort_append {α : Type} (r : α → α → Prop) [D | a :: l1, l2 => by conv_lhs => simp rw [insertionSort_orderedInsert_append] - simp + simp only [List.insertionSort, List.append_eq] rw [insertionSort_insertionSort_append r l1 l2] lemma insertionSort_append_insertionSort_append {α : Type} (r : α → α → Prop) [DecidableRel r] @@ -155,7 +152,7 @@ lemma insertionSort_append_insertionSort_append {α : Type} (r : α → α → P List.insertionSort r (l1 ++ List.insertionSort r l2 ++ l3) = List.insertionSort r (l1 ++ l2 ++ l3) | [], l2, l3 => by - simp + simp only [List.nil_append] exact insertionSort_insertionSort_append r l2 l3 | a :: l1, l2, l3 => by simp only [List.insertionSort, List.append_eq] @@ -175,13 +172,13 @@ lemma takeWhile_orderedInsert {α : Type} (r : α → α → Prop) [DecidableRel | [] => by simp [hr] | c :: l => by - simp + simp only [List.orderedInsert] by_cases h : r b c - · simp [h] + · simp only [h, ↓reduceIte] rw [List.takeWhile_cons_of_pos] - simp + simp only [List.length_cons] simp [hr] - · simp [h] + · simp only [h, ↓reduceIte] have hrba : r b a:= by have ht := IsTotal.total (r := r) a b simp_all @@ -189,7 +186,8 @@ lemma takeWhile_orderedInsert {α : Type} (r : α → α → Prop) [DecidableRel by_contra hn apply h exact IsTrans.trans _ _ _ hrba hn - simp [hl] + simp only [hl, decide_false, Bool.not_false, List.takeWhile_cons_of_pos, List.length_cons, + add_left_inj] exact takeWhile_orderedInsert r a b hr l lemma takeWhile_orderedInsert' {α : Type} (r : α → α → Prop) [DecidableRel r] @@ -198,26 +196,29 @@ lemma takeWhile_orderedInsert' {α : Type} (r : α → α → Prop) [DecidableRe (List.takeWhile (fun c => !decide (r b c)) (List.orderedInsert r a l)).length = (List.takeWhile (fun c => !decide (r b c)) l).length | [] => by - simp + simp only [List.orderedInsert, List.takeWhile_nil, List.length_nil, List.length_eq_zero, + List.takeWhile_eq_nil_iff, List.length_singleton, zero_lt_one, Fin.zero_eta, Fin.isValue, + List.get_eq_getElem, Fin.val_eq_zero, List.getElem_cons_zero, Bool.not_eq_eq_eq_not, + Bool.not_true, decide_eq_false_iff_not, Decidable.not_not, forall_const] have ht := IsTotal.total (r := r) a b simp_all | c :: l => by have hrba : r b a:= by have ht := IsTotal.total (r := r) a b simp_all - simp + simp only [List.orderedInsert] by_cases h : r b c - · simp [h, hrba] + · simp only [h, decide_true, Bool.not_true, Bool.false_eq_true, not_false_eq_true, + List.takeWhile_cons_of_neg, List.length_nil, List.length_eq_zero, List.takeWhile_eq_nil_iff, + List.get_eq_getElem, Bool.not_eq_eq_eq_not, decide_eq_false_iff_not, Decidable.not_not] by_cases hac : r a c · simp [hac, hrba] · simp [hac, h] - · have hcb : r c b := by - have ht := IsTotal.total (r := r) b c - simp_all - by_cases hac : r a c + · by_cases hac : r a c · refine False.elim (h ?_) exact IsTrans.trans _ _ _ hrba hac - · simp [hac, h] + · simp only [hac, ↓reduceIte, h, decide_false, Bool.not_false, List.takeWhile_cons_of_pos, + List.length_cons, add_left_inj] exact takeWhile_orderedInsert' r a b hr l lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [DecidableRel r] @@ -229,23 +230,24 @@ lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [Decidable have ht := IsTotal.total (r := r) a b simp_all intro l hn - simp [insertionSortEquiv] + simp only [List.insertionSort, List.length_cons, insertionSortEquiv, Nat.succ_eq_add_one, + equivCons_trans, Equiv.trans_apply, equivCons_succ, finCongr_apply] conv_lhs => erw [equivCons_succ] conv_rhs => erw [equivCons_succ] - simp + simp only [Equiv.toFun_as_coe] conv_lhs => rhs rhs erw [orderedInsertEquiv_succ] conv_lhs => erw [orderedInsertEquiv_fin_succ] - simp + simp only [Fin.eta, Fin.coe_cast] conv_rhs => rhs rhs erw [orderedInsertEquiv_succ] conv_rhs => erw [orderedInsertEquiv_fin_succ] ext - simp + simp only [Fin.coe_cast, Fin.eta, Fin.cast_trans] let a1 : Fin ((List.orderedInsert r b (List.insertionSort r l)).length + 1) := ⟨↑(orderedInsertPos r (List.orderedInsert r b (List.insertionSort r l)) a), orderedInsertPos_lt_length r (List.orderedInsert r b (List.insertionSort r l)) a⟩ @@ -262,24 +264,25 @@ lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [Decidable = (List.takeWhile (fun c => !decide (r b c)) ((List.takeWhile (fun c => !decide (r a c)) (List.insertionSort r l)))) := by rw [List.takeWhile_takeWhile] - simp + simp only [Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, Bool.decide_and, + decide_not] congr funext c - simp + simp only [Bool.iff_self_and, Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not] intro hbc hac refine hbc ?_ exact IsTrans.trans _ _ _ hrba hac have ha1 : b1.1 ≤ a2.1 := by - simp [a1, a2, orderedInsertPos] + simp only [orderedInsertPos, decide_not] rw [ht] apply List.Sublist.length_le exact List.takeWhile_sublist fun c => !decide (r b c) have ha2 : a1.1 = a2.1 + 1 := by - simp [a1, a2, orderedInsertPos] + simp only [orderedInsertPos, decide_not] rw [takeWhile_orderedInsert] exact hr have hb : b1.1 = b2.1 := by - simp [b1, b2, orderedInsertPos] + simp only [orderedInsertPos, decide_not] rw [takeWhile_orderedInsert'] exact hr let n := ((insertionSortEquiv r l) ⟨n, by simpa using hn⟩) @@ -316,14 +319,14 @@ lemma insertionSortEquiv_commute {α : Type} (r : α → α → Prop) [Decidable if (if ↑n < ↑a2 then ↑n else ↑n + 1) < ↑b2 then if ↑n < ↑a2 then ↑n else ↑n + 1 else (if ↑n < ↑a2 then ↑n else ↑n + 1) + 1 := by by_cases hnb2 : n < b2 - · simp [hnb2] + · simp only [hnb2, ↓reduceIte] have h1 : n < a2 + 1 := by omega have h2 : n < a2 := by omega simp [h1, h2, hnb2] - · simp [hnb2] + · simp only [hnb2, ↓reduceIte, add_lt_add_iff_right] by_cases ha2 : n < a2 · simp [ha2, hnb2] - · simp [ha2] + · simp only [ha2, ↓reduceIte] rw [if_neg] omega apply hnat @@ -334,7 +337,10 @@ lemma insertionSortEquiv_orderedInsert_append {α : Type} (r : α → α → Pro [IsTotal α r] [IsTrans α r] (a a2 : α) : (l1 l2 : List α) → (insertionSortEquiv r (List.orderedInsert r a l1 ++ a2 :: l2) ⟨l1.length + 1, by simp⟩) - = (finCongr (by simp; omega)) + = (finCongr (by + simp only [List.insertionSort, List.append_eq, orderedInsert_length, List.length_cons, + List.length_insertionSort, List.length_append] + omega)) ((insertionSortEquiv r (a :: l1 ++ a2 :: l2)) ⟨l1.length + 1, by simp⟩) | [], l2 => by simp @@ -348,18 +354,22 @@ lemma insertionSortEquiv_orderedInsert_append {α : Type} (r : α → α → Pro (b :: List.orderedInsert r a (l1) ++ a2 :: l2) := by simp [h] rw [insertionSortEquiv_congr _ _ h1] - simp + simp only [List.orderedInsert.eq_2, List.cons_append, List.length_cons, List.insertionSort, + Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_mk, + finCongr_apply] conv_lhs => simp [insertionSortEquiv] rw [insertionSortEquiv_orderedInsert_append r a] have hl : (List.insertionSort r (List.orderedInsert r a l1 ++ a2 :: l2)) = List.insertionSort r (a :: l1 ++ a2 :: l2) := by exact insertionSort_orderedInsert_append r a l1 (a2 :: l2) rw [orderedInsertEquiv_congr _ _ _ hl] - simp + simp only [List.length_cons, List.cons_append, List.insertionSort, finCongr_apply, + Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_succ_eq, + Fin.cast_trans, Fin.cast_eq_self] change Fin.cast _ ((insertionSortEquiv r (b :: a :: (l1 ++ a2 :: l2))) ⟨l1.length + 2, by simp⟩) = _ have hl : l1.length + 1 +1 = l1.length + 2 := by omega - simp [hl] + simp only [List.insertionSort, List.length_cons, hl] conv_rhs => erw [insertionSortEquiv_commute _ _ _ h _ _] simp @@ -369,14 +379,17 @@ lemma insertionSortEquiv_insertionSort_append {α : Type} (r : α → α → Pro (insertionSortEquiv r (List.insertionSort r l1 ++ a :: l2) ⟨l1.length, by simp⟩) = finCongr (by simp) (insertionSortEquiv r (l1 ++ a :: l2) ⟨l1.length, by simp⟩) | [], l2 => by - simp + simp only [List.insertionSort, List.nil_append, List.length_cons, List.length_nil, Fin.zero_eta, + finCongr_refl, Equiv.refl_apply] | b :: l1, l2 => by - simp + simp only [List.insertionSort, List.length_cons, List.cons_append, finCongr_apply] have hl := insertionSortEquiv_orderedInsert_append r b a (List.insertionSort r l1) l2 - simp at hl + simp only [List.length_insertionSort, List.cons_append, List.insertionSort, List.length_cons, + finCongr_apply] at hl rw [hl] have ih := insertionSortEquiv_insertionSort_append r a l1 l2 - simp [insertionSortEquiv] + simp only [insertionSortEquiv, Nat.succ_eq_add_one, List.insertionSort, Equiv.trans_apply, + equivCons_succ] rw [ih] have hl : (List.insertionSort r (List.insertionSort r l1 ++ a :: l2)) = (List.insertionSort r (l1 ++ a :: l2)) := by @@ -390,26 +403,28 @@ lemma insertionSortEquiv_insertionSort_append {α : Type} (r : α → α → Pro -/ -lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] +lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTrans α r] (a : α) (p : α → Prop) [DecidablePred p] (h : p a) : (l : List α) → (hl : l.Sorted r) → List.filter p (List.orderedInsert r a l) = List.orderedInsert r a (List.filter p l) | [], hl => by - simp + simp only [List.orderedInsert, List.filter_eq_self, List.mem_singleton, decide_eq_true_eq, + forall_eq] exact h | b :: l, hl => by - simp + simp only [List.orderedInsert] by_cases hpb : p b <;> by_cases hab : r a b - · simp [hpb, hab] + · simp only [hab, ↓reduceIte, hpb, decide_true, List.filter_cons_of_pos, + List.orderedInsert.eq_2] rw [List.filter_cons_of_pos (by simp [h])] rw [List.filter_cons_of_pos (by simp [hpb])] - · simp [hab] + · simp only [hab, ↓reduceIte] rw [List.filter_cons_of_pos (by simp [hpb])] rw [List.filter_cons_of_pos (by simp [hpb])] - simp [hab] - simp at hl + simp only [List.orderedInsert, hab, ↓reduceIte, List.cons.injEq, true_and] + simp only [List.sorted_cons] at hl exact orderedInsert_filter_of_pos r a p h l hl.2 - · simp [hab] + · simp only [hab, ↓reduceIte] rw [List.filter_cons_of_pos (by simp [h]), List.filter_cons_of_neg (by simp [hpb])] rw [List.orderedInsert_eq_take_drop] @@ -417,33 +432,34 @@ lemma orderedInsert_filter_of_pos {α : Type} (r : α → α → Prop) [Decidabl (List.filter (fun b => decide (p b)) l) = [] := by rw [List.takeWhile_eq_nil_iff] intro c hc - simp at hc + simp only [List.get_eq_getElem, decide_not, Bool.not_eq_eq_eq_not, Bool.not_true, + decide_eq_false_iff_not] at hc apply hc apply IsTrans.trans a b _ hab - simp at hl + simp only [List.sorted_cons] at hl apply hl.1 have hlf : (List.filter (fun b => decide (p b)) l)[0] ∈ (List.filter (fun b => decide (p b)) l) := by exact List.getElem_mem c - simp [- List.getElem_mem] at hlf + simp only [List.mem_filter, decide_eq_true_eq] at hlf exact hlf.1 rw [hl] - simp + simp only [decide_not, List.nil_append, List.cons.injEq, true_and] conv_lhs => rw [← List.takeWhile_append_dropWhile (fun b => decide ¬r a b) (List.filter (fun b => decide (p b)) l)] rw [hl] simp - · simp [hab] + · simp only [hab, ↓reduceIte] rw [List.filter_cons_of_neg (by simp [hpb])] rw [List.filter_cons_of_neg (by simp [hpb])] - simp at hl + simp only [List.sorted_cons] at hl exact orderedInsert_filter_of_pos r a p h l hl.2 -lemma orderedInsert_filter_of_neg {α : Type} (r : α → α → Prop) [DecidableRel r] [IsTotal α r] - [IsTrans α r] (a : α) (p : α → Prop) [DecidablePred p] (h : ¬ p a) (l : List α) : +lemma orderedInsert_filter_of_neg {α : Type} (r : α → α → Prop) [DecidableRel r] + (a : α) (p : α → Prop) [DecidablePred p] (h : ¬ p a) (l : List α) : List.filter p (List.orderedInsert r a l) = (List.filter p l) := by rw [List.orderedInsert_eq_take_drop] - simp + simp only [decide_not, List.filter_append] rw [List.filter_cons_of_neg] rw [← List.filter_append] congr @@ -456,85 +472,96 @@ lemma insertionSort_filter {α : Type} (r : α → α → Prop) [DecidableRel r] List.filter p (List.insertionSort r l) | [] => by simp | a :: l => by - simp + simp only [List.insertionSort] by_cases h : p a · rw [orderedInsert_filter_of_pos] rw [List.filter_cons_of_pos] - simp + simp only [List.insertionSort] rw [insertionSort_filter] - simp_all - simp_all + simp_all only [decide_true] + simp_all only exact List.sorted_insertionSort r l · rw [orderedInsert_filter_of_neg] rw [List.filter_cons_of_neg] rw [insertionSort_filter] - simp_all + simp_all only [decide_false, Bool.false_eq_true, not_false_eq_true] exact h lemma takeWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → + [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → List.takeWhile (fun c => ¬ r a c) l = List.filter (fun c => ¬ r a c) l | [], _ => by simp | b :: l, hl => by - simp at hl + simp only [List.sorted_cons] at hl by_cases hb : ¬ r a b - · simp [hb] + · simp only [decide_not, hb, decide_false, Bool.not_false, List.takeWhile_cons_of_pos, + List.filter_cons_of_pos, List.cons.injEq, true_and] simpa using takeWhile_sorted_eq_filter r a l hl.2 - · simp_all + · simp_all only [Decidable.not_not, decide_not, decide_true, Bool.not_true, Bool.false_eq_true, + not_false_eq_true, List.takeWhile_cons_of_neg, List.filter_cons_of_neg, List.nil_eq, + List.filter_eq_nil_iff, Bool.not_eq_eq_eq_not, decide_eq_false_iff_not] intro c hc apply IsTrans.trans a b c hb exact hl.1 c hc lemma dropWhile_sorted_eq_filter {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → + [IsTrans α r] (a : α) : (l : List α) → (hl : l.Sorted r) → List.dropWhile (fun c => ¬ r a c) l = List.filter (fun c => r a c) l | [], _ => by simp | b :: l, hl => by - simp at hl + simp only [List.sorted_cons] at hl by_cases hb : ¬ r a b - · simp [hb] + · simp only [decide_not, hb, decide_false, Bool.not_false, List.dropWhile_cons_of_pos, + Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg] simpa using dropWhile_sorted_eq_filter r a l hl.2 - · simp_all + · simp_all only [Decidable.not_not, decide_not, decide_true, Bool.not_true, Bool.false_eq_true, + not_false_eq_true, List.dropWhile_cons_of_neg, List.filter_cons_of_pos, List.cons.injEq, + true_and] symm rw [List.filter_eq_self] intro c hc - simp + simp only [decide_eq_true_eq] apply IsTrans.trans a b c hb exact hl.1 c hc lemma dropWhile_sorted_eq_filter_filter {α : Type} (r : α → α → Prop) [DecidableRel r] - [IsTotal α r] [IsTrans α r] (a : α) :(l : List α) → (hl : l.Sorted r) → + [IsTrans α r] (a : α) :(l : List α) → (hl : l.Sorted r) → List.filter (fun c => r a c) l = List.filter (fun c => r a c ∧ r c a) l ++ List.filter (fun c => r a c ∧ ¬ r c a) l | [], _ => by simp | b :: l, hl => by - simp at hl + simp only [List.sorted_cons] at hl by_cases hb : ¬ r a b - · simp [hb] + · simp only [hb, decide_false, Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg, + Bool.decide_and, Bool.false_and, decide_not] simpa using dropWhile_sorted_eq_filter_filter r a l hl.2 - · simp_all + · simp_all only [Decidable.not_not, decide_true, List.filter_cons_of_pos, Bool.decide_and, + decide_not] by_cases hba : r b a - · simp [hba] + · simp only [hba, decide_true, Bool.not_true, Bool.and_false, Bool.false_eq_true, + not_false_eq_true, List.filter_cons_of_neg] rw [List.filter_cons_of_pos] rw [dropWhile_sorted_eq_filter_filter] - simp + simp only [Bool.decide_and, decide_not, List.cons_append] exact hl.2 simp_all - · simp[hba] + · simp only [hba, decide_false, Bool.and_false, Bool.false_eq_true, not_false_eq_true, + List.filter_cons_of_neg] have h1 : List.filter (fun c => decide (r a c) && decide (r c a)) l = [] := by rw [@List.filter_eq_nil_iff] intro c hc - simp + simp only [Bool.and_eq_true, decide_eq_true_eq, not_and] intro hac hca apply hba apply IsTrans.trans b c a _ hca exact hl.1 c hc rw [h1] rw [dropWhile_sorted_eq_filter_filter] - simp [h1] + simp only [Bool.decide_and, h1, decide_not, List.nil_append] rw [List.filter_cons_of_pos] - simp_all + simp_all only [List.filter_eq_nil_iff, Bool.and_eq_true, decide_eq_true_eq, not_and, + decide_true, decide_false, Bool.not_false, Bool.and_self] exact hl.2 lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [DecidableRel r] @@ -547,7 +574,7 @@ lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [Decidabl by_cases h : r a b ∧ r b a · have hl := orderedInsert_filter_of_pos r b (fun c => r a c ∧ r c a) h (List.insertionSort r l) (by exact List.sorted_insertionSort r l) - simp at hl ⊢ + simp only [Bool.decide_and] at hl ⊢ rw [hl] rw [List.orderedInsert_eq_take_drop] have ht : List.takeWhile (fun b_1 => decide ¬r b b_1) @@ -555,35 +582,40 @@ lemma filter_rel_eq_insertionSort {α : Type} (r : α → α → Prop) [Decidabl (List.insertionSort r l)) = [] := by rw [List.takeWhile_eq_nil_iff] intro hl - simp + simp only [List.get_eq_getElem, decide_not, Bool.not_eq_eq_eq_not, Bool.not_true, + decide_eq_false_iff_not, Decidable.not_not] have hx := List.getElem_mem hl - simp [- List.getElem_mem] at hx + simp only [List.mem_filter, List.mem_insertionSort, Bool.and_eq_true, + decide_eq_true_eq] at hx apply IsTrans.trans b a _ h.2 simp_all rw [ht] - simp + simp only [decide_not, List.nil_append] rw [List.filter_cons_of_pos] - simp + simp only [List.cons.injEq, true_and] have ih := filter_rel_eq_insertionSort r a l - simp at ih + simp only [Bool.decide_and] at ih rw [← ih] have htd := List.takeWhile_append_dropWhile (fun b_1 => decide ¬r b b_1) (List.filter (fun b => decide (r a b) && decide (r b a)) (List.insertionSort r l)) - simp [decide_not, - List.takeWhile_append_dropWhile] at htd + simp only [decide_not] at htd conv_rhs => rw [← htd] - simp [- List.takeWhile_append_dropWhile] + simp only [List.self_eq_append_left, List.takeWhile_eq_nil_iff, List.get_eq_getElem, + Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, Decidable.not_not] intro hl have hx := List.getElem_mem hl - simp [- List.getElem_mem] at hx + simp only [List.mem_filter, List.mem_insertionSort, Bool.and_eq_true, decide_eq_true_eq] at hx apply IsTrans.trans b a _ h.2 - simp_all + simp_all only [decide_not, List.takeWhile_eq_nil_iff, List.get_eq_getElem, + Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, Decidable.not_not, + List.takeWhile_append_dropWhile] simp_all · have hl := orderedInsert_filter_of_neg r b (fun c => r a c ∧ r c a) h (List.insertionSort r l) - simp at hl ⊢ + simp only [Bool.decide_and] at hl ⊢ rw [hl] rw [List.filter_cons_of_neg] have ih := filter_rel_eq_insertionSort r a l - simp_all + simp_all only [not_and, Bool.decide_and] simpa using h lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRel r] @@ -594,8 +626,7 @@ lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRe ++ (List.filter (fun c => r a c ∧ r c a) l1) ++ l ++ (List.filter (fun c => r a c ∧ r c a) l2) - ++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) - := by + ++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) := by have hl : List.insertionSort r (l1 ++ l ++ l2) = List.takeWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) ++ List.dropWhile (fun c => ¬ r a c) ((l1 ++ l ++ l2).insertionSort r) := by @@ -606,27 +637,32 @@ lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRe rw [takeWhile_sorted_eq_filter, takeWhile_sorted_eq_filter] rw [← insertionSort_filter, ← insertionSort_filter] congr 1 - simp + simp only [decide_not, List.append_assoc, List.filter_append, List.append_cancel_left_eq, + List.append_left_eq_self, List.filter_eq_nil_iff, Bool.not_eq_eq_eq_not, Bool.not_true, + decide_eq_false_iff_not, Decidable.not_not] exact fun b hb => (h b hb).1 exact List.sorted_insertionSort r (l1 ++ l2) exact List.sorted_insertionSort r (l1 ++ l ++ l2) conv_lhs => rw [hl, hlt] simp only [decide_not, Bool.decide_and] - simp + simp only [List.append_assoc, List.append_cancel_left_eq] have h1 := dropWhile_sorted_eq_filter r a (List.insertionSort r (l1 ++ (l ++ l2))) - simp at h1 + simp only [decide_not] at h1 rw [h1] rw [dropWhile_sorted_eq_filter_filter, filter_rel_eq_insertionSort] - simp + simp only [Bool.decide_and, List.filter_append, decide_not, List.append_assoc, + List.append_cancel_left_eq] congr 1 - simp + simp only [List.filter_eq_self, Bool.and_eq_true, decide_eq_true_eq] exact fun a a_1 => h a a_1 congr 1 have h1 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ (l ++ l2)) - simp at h1 + simp only [Bool.and_eq_true, decide_eq_true_eq, Bool.not_eq_eq_eq_not, Bool.not_true, + decide_eq_false_iff_not, Bool.decide_and, decide_not, List.filter_append] at h1 rw [← h1] have h2 := insertionSort_filter r (fun c => decide (r a c) && !decide (r c a)) (l1 ++ l2) - simp at h2 + simp only [Bool.and_eq_true, decide_eq_true_eq, Bool.not_eq_eq_eq_not, Bool.not_true, + decide_eq_false_iff_not, Bool.decide_and, decide_not, List.filter_append] at h2 rw [← h2] congr have hl : List.filter (fun b => decide (r a b) && !decide (r b a)) l = [] := by @@ -634,7 +670,7 @@ lemma insertionSort_of_eq_list {α : Type} (r : α → α → Prop) [DecidableRe intro c hc simp_all rw [hl] - simp + simp only [List.nil_append] exact List.sorted_insertionSort r (l1 ++ (l ++ l2)) exact List.sorted_insertionSort r (l1 ++ (l ++ l2)) @@ -644,12 +680,11 @@ lemma insertionSort_of_takeWhile_filter {α : Type} (r : α → α → Prop) [De (List.takeWhile (fun c => ¬ r a c) ((l1 ++ l2).insertionSort r)) ++ (List.filter (fun c => r a c ∧ r c a) l1) ++ (List.filter (fun c => r a c ∧ r c a) l2) - ++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) - := by + ++ (List.filter (fun c => r a c ∧ ¬ r c a) ((l1 ++ l2).insertionSort r)) := by trans List.insertionSort r (l1 ++ [] ++ l2) - simp + simp only [List.append_nil] rw [insertionSort_of_eq_list r a l1 [] l2] - simp + simp only [decide_not, Bool.decide_and, List.append_nil, List.append_assoc] simp end HepLean.List diff --git a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean index 812fecc1..2365c480 100644 --- a/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean +++ b/HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Grading.lean @@ -25,7 +25,7 @@ def statisticSubmodule (f : FieldStatistic) : Submodule ℂ 𝓕.CrAnAlgebra := Submodule.span ℂ {a | ∃ φs, a = ofCrAnList φs ∧ (𝓕 |>ₛ φs) = f} lemma ofCrAnList_mem_statisticSubmodule_of (φs : List 𝓕.CrAnStates) (f : FieldStatistic) - (h : (𝓕 |>ₛ φs) = f) : + (h : (𝓕 |>ₛ φs) = f) : ofCrAnList φs ∈ statisticSubmodule f := by refine Submodule.mem_span.mpr fun _ a => a ⟨φs, ⟨rfl, h⟩⟩ @@ -234,6 +234,7 @@ lemma directSum_eq_bosonic_plus_fermionic conv_lhs => rw [hx, hy] abel +/-- The instance of a graded algebra on `CrAnAlgebra`. -/ instance crAnAlgebraGrade : GradedAlgebra (A := 𝓕.CrAnAlgebra) statisticSubmodule where one_mem := by simp only [statisticSubmodule] diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean index 7fd9aba7..35d0efc6 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/Basic.lean @@ -134,7 +134,6 @@ lemma ι_superCommute_ofCrAnState_superCommute_ofCrAnState_ofCrAnState (φ1 φ2 left use φ1, φ2, φ3 -@[simp] lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState (φ1 φ2 φ3 : 𝓕.CrAnStates) : ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, ofCrAnState φ3]ₛca = 0 := by rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton] @@ -147,16 +146,15 @@ lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState (φ1 φ2 · rw [superCommute_fermionic_fermionic_symm h h'] simp [ofCrAnList_singleton] -@[simp] lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnList (φ1 φ2 : 𝓕.CrAnStates) (φs : List 𝓕.CrAnStates) : ι [[ofCrAnState φ1, ofCrAnState φ2]ₛca, ofCrAnList φs]ₛca = 0 := by rw [← ofCrAnList_singleton, ← ofCrAnList_singleton] rcases superCommute_ofCrAnList_ofCrAnList_bosonic_or_fermionic [φ1] [φ2] with h | h · rw [superCommute_bosonic_ofCrAnList_eq_sum _ _ h] - simp [ofCrAnList_singleton] + simp [ofCrAnList_singleton, ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState] · rw [superCommute_fermionic_ofCrAnList_eq_sum _ _ h] - simp [ofCrAnList_singleton] + simp [ofCrAnList_singleton, ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnState] @[simp] lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_crAnAlgebra (φ1 φ2 : 𝓕.CrAnStates) @@ -164,7 +162,7 @@ lemma ι_superCommute_superCommute_ofCrAnState_ofCrAnState_crAnAlgebra (φ1 φ2 change (ι.toLinearMap ∘ₗ superCommute [ofCrAnState φ1, ofCrAnState φ2]ₛca) a = _ have h1 : (ι.toLinearMap ∘ₗ superCommute [ofCrAnState φ1, ofCrAnState φ2]ₛca) = 0 := by apply (ofCrAnListBasis.ext fun l ↦ ?_) - simp + simp [ι_superCommute_superCommute_ofCrAnState_ofCrAnState_ofCrAnList] rw [h1] simp @@ -239,7 +237,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 only [fieldOpIdealSet, exists_prop, Set.mem_setOf_eq] 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 diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean index 60b2bd40..9736cd16 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/SuperCommute.lean @@ -49,8 +49,11 @@ lemma ι_superCommute_eq_of_equiv_right (a b1 b2 : 𝓕.CrAnAlgebra) (h : b1 ≈ simp only [LinearMap.mem_ker, ← map_sub] exact ι_superCommute_right_zero_of_mem_ideal a (b1 - b2) h -noncomputable def superCommuteRight (a : 𝓕.CrAnAlgebra) : FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where - toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.superCommute a) (ι_superCommute_eq_of_equiv_right a) +/-- The super commutor on the `FieldOpAlgebra` defined as a linear map `[a,_]ₛ`. -/ +noncomputable def superCommuteRight (a : 𝓕.CrAnAlgebra) : + FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where + toFun := Quotient.lift (ι.toLinearMap ∘ₗ CrAnAlgebra.superCommute a) + (ι_superCommute_eq_of_equiv_right a) map_add' x y := by obtain ⟨x, hx⟩ := ι_surjective x obtain ⟨y, hy⟩ := ι_surjective y @@ -64,11 +67,11 @@ noncomputable def superCommuteRight (a : 𝓕.CrAnAlgebra) : FieldOpAlgebra 𝓕 rw [← map_smul, ι_apply, ι_apply] simp -lemma superCommuteRight_apply_ι (a b : 𝓕.CrAnAlgebra) : superCommuteRight a (ι b) = ι [a, b]ₛca := by - rfl +lemma superCommuteRight_apply_ι (a b : 𝓕.CrAnAlgebra) : + superCommuteRight a (ι b) = ι [a, b]ₛca := by rfl -lemma superCommuteRight_apply_quot (a b : 𝓕.CrAnAlgebra) : superCommuteRight a ⟦b⟧= ι [a, b]ₛca := by - rfl +lemma superCommuteRight_apply_quot (a b : 𝓕.CrAnAlgebra) : + superCommuteRight a ⟦b⟧= ι [a, b]ₛca := by rfl lemma superCommuteRight_eq_of_equiv (a1 a2 : 𝓕.CrAnAlgebra) (h : a1 ≈ a2) : superCommuteRight a1 = superCommuteRight a2 := by @@ -85,6 +88,7 @@ lemma superCommuteRight_eq_of_equiv (a1 a2 : 𝓕.CrAnAlgebra) (h : a1 ≈ a2) : simp only [add_sub_cancel] simp +/-- The super commutor on the `FieldOpAlgebra`. -/ noncomputable def superCommute : FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 →ₗ[ℂ] FieldOpAlgebra 𝓕 where toFun := Quotient.lift superCommuteRight superCommuteRight_eq_of_equiv diff --git a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean index c3f4f5a1..b2e7d279 100644 --- a/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean +++ b/HepLean/PerturbationTheory/Algebras/FieldOpAlgebra/TimeOrder.lean @@ -20,33 +20,35 @@ namespace FieldOpAlgebra variable {𝓕 : FieldSpecification} lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates} - (φs1 φs2 : List 𝓕.CrAnStates) (h : + (φs1 φs2 : List 𝓕.CrAnStates) (h : crAnTimeOrderRel φ1 φ2 ∧ crAnTimeOrderRel φ1 φ3 ∧ crAnTimeOrderRel φ2 φ1 ∧ crAnTimeOrderRel φ2 φ3 ∧ - crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2): - ι 𝓣ᶠ(ofCrAnList φs1 * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ofCrAnList φs2) - = 0 := by + crAnTimeOrderRel φ3 φ1 ∧ crAnTimeOrderRel φ3 φ2) : + ι 𝓣ᶠ(ofCrAnList φs1 * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * + ofCrAnList φs2) = 0 := by let l1 := - (List.takeWhile (fun c => ¬ crAnTimeOrderRel φ1 c) ((φs1 ++ φs2).insertionSort crAnTimeOrderRel)) + (List.takeWhile (fun c => ¬ crAnTimeOrderRel φ1 c) + ((φs1 ++ φs2).insertionSort crAnTimeOrderRel)) ++ (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ crAnTimeOrderRel c φ1) φs1) let l2 := (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ crAnTimeOrderRel c φ1) φs2) - ++ (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ ¬ crAnTimeOrderRel c φ1) ((φs1 ++ φs2).insertionSort crAnTimeOrderRel)) + ++ (List.filter (fun c => crAnTimeOrderRel φ1 c ∧ ¬ crAnTimeOrderRel c φ1) + ((φs1 ++ φs2).insertionSort crAnTimeOrderRel)) have h123 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ1 :: φ2 :: φ3 :: φs2)) = crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) - • (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ2, φ3]) * ι (ofCrAnList l2)):= by + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ2, φ3]) * ι (ofCrAnList l2)) := by have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ1, φ2, φ3] φs2 - (by simp_all) - rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ2 :: φ3 :: φs2 = φs1 ++ [φ1, φ2, φ3] ++ φs2 by simp, - crAnTimeOrderList, h1] + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ2 :: φ3 :: φs2 = φs1 ++ [φ1, φ2, φ3] ++ φs2 + by simp, crAnTimeOrderList, h1] simp only [List.append_assoc, List.singleton_append, decide_not, Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] have h132 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ1 :: φ3 :: φ2 :: φs2)) = crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) - • (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ3, φ2]) * ι (ofCrAnList l2)):= by + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ1, φ3, φ2]) * ι (ofCrAnList l2)) := by have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ1, φ3, φ2] φs2 - (by simp_all) - rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ3 :: φ2 :: φs2 = φs1 ++ [φ1, φ3, φ2] ++ φs2 by simp, - crAnTimeOrderList, h1] + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ1 :: φ3 :: φ2 :: φs2 = φs1 ++ [φ1, φ3, φ2] ++ φs2 + by simp, crAnTimeOrderList, h1] simp only [List.singleton_append, decide_not, Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] congr 1 @@ -67,11 +69,11 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : exact List.Perm.swap φ1 φ2 [φ3] have h231 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ2 :: φ3 :: φ1 :: φs2)) = crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) - • (ι (ofCrAnList l1) * ι (ofCrAnList [φ2, φ3, φ1]) * ι (ofCrAnList l2)):= by + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ2, φ3, φ1]) * ι (ofCrAnList l2)) := by have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ2, φ3, φ1] φs2 - (by simp_all) - rw [timeOrder_ofCrAnList, show φs1 ++ φ2 :: φ3 :: φ1 :: φs2 = φs1 ++ [φ2, φ3, φ1] ++ φs2 by simp, - crAnTimeOrderList, h1] + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ2 :: φ3 :: φ1 :: φs2 = φs1 ++ [φ2, φ3, φ1] ++ φs2 + by simp, crAnTimeOrderList, h1] simp only [List.singleton_append, decide_not, Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] congr 1 @@ -85,11 +87,11 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : simp_all have h321 : ι 𝓣ᶠ(ofCrAnList (φs1 ++ φ3 :: φ2 :: φ1 :: φs2)) = crAnTimeOrderSign (φs1 ++ φ1 :: φ2 :: φ3 :: φs2) - • (ι (ofCrAnList l1) * ι (ofCrAnList [φ3, φ2, φ1]) * ι (ofCrAnList l2)):= by + • (ι (ofCrAnList l1) * ι (ofCrAnList [φ3, φ2, φ1]) * ι (ofCrAnList l2)) := by have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ1 φs1 [φ3, φ2, φ1] φs2 - (by simp_all) - rw [timeOrder_ofCrAnList, show φs1 ++ φ3 :: φ2 :: φ1 :: φs2 = φs1 ++ [φ3, φ2, φ1] ++ φs2 by simp, - crAnTimeOrderList, h1] + (by simp_all) + rw [timeOrder_ofCrAnList, show φs1 ++ φ3 :: φ2 :: φ1 :: φs2 = φs1 ++ [φ3, φ2, φ1] ++ φs2 + by simp, crAnTimeOrderList, h1] simp only [List.singleton_append, decide_not, Bool.decide_and, ofCrAnList_append, map_smul, map_mul, l1, l2, mul_assoc] congr 1 @@ -125,7 +127,8 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : repeat rw [mul_assoc] rw [← mul_sub, ← mul_sub, ← mul_sub] rw [← sub_mul, ← sub_mul, ← sub_mul] - trans ι (ofCrAnList l1) * ι [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ι (ofCrAnList l2) + trans ι (ofCrAnList l1) * ι [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * + ι (ofCrAnList l2) rw [mul_assoc] congr rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_singleton] @@ -137,7 +140,7 @@ lemma ι_timeOrder_superCommute_superCommute_eq_time_ofCrAnList {φ1 φ2 φ3 : simp_all lemma ι_timeOrder_superCommute_superCommute_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAnStates} - (φs1 φs2 : List 𝓕.CrAnStates): + (φs1 φs2 : List 𝓕.CrAnStates) : ι 𝓣ᶠ(ofCrAnList φs1 * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * ofCrAnList φs2) = 0 := by by_cases h : @@ -150,7 +153,7 @@ lemma ι_timeOrder_superCommute_superCommute_ofCrAnList {φ1 φ2 φ3 : 𝓕.CrAn simp @[simp] -lemma ι_timeOrder_superCommute_superCommute {φ1 φ2 φ3 : 𝓕.CrAnStates} (a b : 𝓕.CrAnAlgebra): +lemma ι_timeOrder_superCommute_superCommute {φ1 φ2 φ3 : 𝓕.CrAnStates} (a b : 𝓕.CrAnAlgebra) : ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * b) = 0 := by let pb (b : 𝓕.CrAnAlgebra) (hc : b ∈ Submodule.span ℂ (Set.range ofCrAnListBasis)) : Prop := ι 𝓣ᶠ(a * [ofCrAnState φ1, [ofCrAnState φ2, ofCrAnState φ3]ₛca]ₛca * b) = 0 @@ -215,20 +218,21 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates} rw [h1] simp only [map_smul] have h1 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ φs' [φ, ψ] φs - (by simp_all) + (by simp_all) rw [crAnTimeOrderList, show φs' ++ φ :: ψ :: φs = φs' ++ [φ, ψ] ++ φs by simp, h1] have h2 := insertionSort_of_eq_list 𝓕.crAnTimeOrderRel φ φs' [ψ, φ] φs - (by simp_all) + (by simp_all) rw [crAnTimeOrderList, show φs' ++ ψ :: φ :: φs = φs' ++ [ψ, φ] ++ φs by simp, h2] repeat rw [ofCrAnList_append] rw [smul_smul, mul_comm, ← smul_smul, ← smul_sub] rw [map_mul, map_mul, map_mul, map_mul, map_mul, map_mul, map_mul, map_mul] rw [← mul_smul_comm] - rw [mul_assoc, mul_assoc, mul_assoc ,mul_assoc ,mul_assoc ,mul_assoc] + rw [mul_assoc, mul_assoc, mul_assoc, mul_assoc, mul_assoc, mul_assoc] rw [← mul_sub, ← mul_sub, mul_smul_comm, mul_smul_comm, ← smul_mul_assoc, ← smul_mul_assoc] rw [← sub_mul] - have h1 : (ι (ofCrAnList [φ, ψ]) - (exchangeSign (𝓕.crAnStatistics φ)) (𝓕.crAnStatistics ψ) • ι (ofCrAnList [ψ, φ])) = + have h1 : (ι (ofCrAnList [φ, ψ]) - + (exchangeSign (𝓕.crAnStatistics φ)) (𝓕.crAnStatistics ψ) • ι (ofCrAnList [ψ, φ])) = ι [ofCrAnState φ, ofCrAnState ψ]ₛca := by rw [superCommute_ofCrAnState_ofCrAnState] rw [← ofCrAnList_singleton, ← ofCrAnList_singleton, ← ofCrAnList_append] @@ -237,7 +241,8 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates} rw [← ofCrAnList_append] simp rw [h1] - have hc : ι ((superCommute (ofCrAnState φ)) (ofCrAnState ψ)) ∈ Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by + have hc : ι ((superCommute (ofCrAnState φ)) (ofCrAnState ψ)) ∈ + Subalgebra.center ℂ 𝓕.FieldOpAlgebra := by apply ι_superCommute_ofCrAnState_ofCrAnState_mem_center rw [Subalgebra.mem_center_iff] at hc repeat rw [← mul_assoc] @@ -272,7 +277,6 @@ lemma ι_timeOrder_superCommute_eq_time {φ ψ : 𝓕.CrAnStates} · intro x hx hpx simp_all [pb, hpx] - lemma ι_timeOrder_superCommute_neq_time {φ ψ : 𝓕.CrAnStates} (hφψ : ¬ (crAnTimeOrderRel φ ψ ∧ crAnTimeOrderRel ψ φ)) (a b : 𝓕.CrAnAlgebra) : ι 𝓣ᶠ(a * [ofCrAnState φ, ofCrAnState ψ]ₛca * b) = 0 := by @@ -280,18 +284,16 @@ lemma ι_timeOrder_superCommute_neq_time {φ ψ : 𝓕.CrAnStates} have hφψ : ¬ (crAnTimeOrderRel φ ψ) ∨ ¬ (crAnTimeOrderRel ψ φ) := by exact Decidable.not_and_iff_or_not.mp hφψ rcases hφψ with hφψ | hφψ - · rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel ] - have ht := IsTotal.total (r := crAnTimeOrderRel) φ ψ + · rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel] 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 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 ] + rw [timeOrder_superCommute_ofCrAnState_ofCrAnState_not_crAnTimeOrderRel] simp only [mul_zero, zero_mul, map_zero, or_true] simp_all - /-! ## Defining normal order for `FiedOpAlgebra`. diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean index 2b8eaadf..d2eed3a9 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSign.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSign.lean @@ -280,7 +280,7 @@ lemma koszulSign_swap_eq_rel {ψ φ : 𝓕} (h1 : le φ ψ) (h2 : le ψ φ) : ( apply Wick.koszulSignInsert_eq_perm exact List.Perm.append_left φs (List.Perm.swap ψ φ φs') -lemma koszulSign_eq_rel_eq_stat_append {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal 𝓕 le] +lemma koszulSign_eq_rel_eq_stat_append {ψ φ : 𝓕} [IsTrans 𝓕 le] (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) → koszulSign q le (φ :: ψ :: φs) = koszulSign q le φs := by intro φs @@ -294,7 +294,7 @@ lemma koszulSign_eq_rel_eq_stat_append {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal rw [koszulSignInsert_eq_rel_eq_stat q le h1 h2 hq] simp -lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le] [IsTotal 𝓕 le] +lemma koszulSign_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le] (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs' φs : List 𝓕) → koszulSign q le (φs' ++ φ :: ψ :: φs) = koszulSign q le (φs' ++ φs) | [], φs => by @@ -395,7 +395,7 @@ lemma koszulSign_of_append_eq_insertionSort [IsTotal 𝓕 le] [IsTrans 𝓕 le] -/ -lemma koszulSign_perm_eq_append [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) (φs φs' φs2 : List 𝓕) +lemma koszulSign_perm_eq_append [IsTrans 𝓕 le] (φ : 𝓕) (φs φs' φs2 : List 𝓕) (hp : φs.Perm φs') : (h : ∀ φ' ∈ φs, le φ φ' ∧ le φ' φ) → koszulSign q le (φs ++ φs2) = koszulSign q le (φs' ++ φs2) := by let motive (φs φs' : List 𝓕) (hp : φs.Perm φs') : Prop := @@ -423,7 +423,7 @@ lemma koszulSign_perm_eq_append [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) refine h φ' ?_ exact (List.Perm.mem_iff (id (List.Perm.symm h1))).mp hφ -lemma koszulSign_perm_eq [IsTotal 𝓕 le] [IsTrans 𝓕 le] (φ : 𝓕) : (φs1 φs φs' φs2 : List 𝓕) → +lemma koszulSign_perm_eq [IsTrans 𝓕 le] (φ : 𝓕) : (φs1 φs φs' φs2 : List 𝓕) → (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 diff --git a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean index 2d899df9..79ef8e42 100644 --- a/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean +++ b/HepLean/PerturbationTheory/Koszul/KoszulSignInsert.lean @@ -247,7 +247,7 @@ lemma koszulSignInsert_of_le_mem (φ0 : 𝓕) : (φs : List 𝓕) → (h : ∀ b exact h b (List.mem_cons_of_mem _ hb) · exact h φ1 (List.mem_cons_self _ _) -lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTotal 𝓕 le] [IsTrans 𝓕 le] +lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTrans 𝓕 le] (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) → koszulSignInsert q le φ φs = koszulSignInsert q le ψ φs | [] => by @@ -268,7 +268,7 @@ lemma koszulSignInsert_eq_rel_eq_stat {ψ φ : 𝓕} [IsTotal 𝓕 le] [IsTrans 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] +lemma koszulSignInsert_eq_remove_same_stat_append {ψ φ φ' : 𝓕} [IsTrans 𝓕 le] (h1 : le φ ψ) (h2 : le ψ φ) (hq : q ψ = q φ) : (φs : List 𝓕) → koszulSignInsert q le φ' (φ :: ψ :: φs) = koszulSignInsert q le φ' φs := by intro φs