Merge pull request #316 from HEPLean/WickTheoremDoc

feat: Add static wick terms
This commit is contained in:
Joseph Tooby-Smith 2025-02-05 10:30:14 +00:00 committed by GitHub
commit 01818d7578
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
13 changed files with 257 additions and 129 deletions

View file

@ -131,6 +131,7 @@ import HepLean.PerturbationTheory.FieldOpAlgebra.Grading
import HepLean.PerturbationTheory.FieldOpAlgebra.NormalOrder.Basic
import HepLean.PerturbationTheory.FieldOpAlgebra.NormalOrder.Lemmas
import HepLean.PerturbationTheory.FieldOpAlgebra.NormalOrder.WickContractions
import HepLean.PerturbationTheory.FieldOpAlgebra.StaticWickTerm
import HepLean.PerturbationTheory.FieldOpAlgebra.StaticWickTheorem
import HepLean.PerturbationTheory.FieldOpAlgebra.SuperCommute
import HepLean.PerturbationTheory.FieldOpAlgebra.TimeContraction

View file

@ -39,6 +39,20 @@ lemma normalOrder_one_eq_one : normalOrder (𝓕 := 𝓕) 1 = 1 := by
rw [normalOrder_ofCrAnFieldOpList]
simp
@[simp]
lemma normalOrder_ofFieldOpList_nil : normalOrder (𝓕 := 𝓕) (ofFieldOpList []) = 1 := by
rw [ofFieldOpList]
rw [normalOrder_eq_ι_normalOrderF]
simp only [ofFieldOpListF_nil]
change normalOrder (𝓕 := 𝓕) 1 = _
simp
@[simp]
lemma normalOrder_ofCrAnFieldOpList_nil : normalOrder (𝓕 := 𝓕) (ofCrAnFieldOpList []) = 1 := by
rw [normalOrder_ofCrAnFieldOpList]
simp only [normalOrderSign_nil, normalOrderList_nil, one_smul]
rfl
lemma ofCrAnFieldOpList_eq_normalOrder (φs : List 𝓕.CrAnFieldOp) :
ofCrAnFieldOpList (normalOrderList φs) = normalOrderSign φs • 𝓝(ofCrAnFieldOpList φs) := by
rw [normalOrder_ofCrAnFieldOpList, smul_smul, normalOrderSign, Wick.koszulSign_mul_self,

View file

@ -0,0 +1,121 @@
/-
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.FieldOpAlgebra.NormalOrder.WickContractions
import HepLean.PerturbationTheory.WickContraction.Sign.InsertNone
import HepLean.PerturbationTheory.WickContraction.Sign.InsertSome
import HepLean.PerturbationTheory.WickContraction.StaticContract
/-!
# Static Wick's terms
-/
open FieldSpecification
variable {𝓕 : FieldSpecification}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open FieldOpAlgebra
open FieldStatistic
noncomputable section
/-- For a Wick contraction `φsΛ`, we define `staticWickTerm φsΛ` to be the element of
`𝓕.FieldOpAlgebra` given by `φsΛ.sign • φsΛ.staticContract * 𝓝([φsΛ]ᵘᶜ)`. -/
def staticWickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : 𝓕.FieldOpAlgebra :=
φsΛ.sign • φsΛ.staticContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
@[simp]
lemma staticWickTerm_empty_nil :
staticWickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
rw [staticWickTerm, uncontractedListGet, nil_zero_uncontractedList]
simp [sign, empty, staticContract]
lemma staticWickTerm_insert_zero_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) :
(φsΛ ↩Λ φ 0 none).staticWickTerm =
φsΛ.sign • φsΛ.staticContract * 𝓝(ofFieldOpList (φ :: [φsΛ]ᵘᶜ)) := by
symm
erw [staticWickTerm, sign_insert_none_zero]
simp only [staticContract_insertAndContract_none, insertAndContract_uncontractedList_none_zero,
Algebra.smul_mul_assoc]
lemma staticWickTerm_insert_zero_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (k : { x // x ∈ φsΛ.uncontracted }) :
(φsΛ ↩Λ φ 0 k).staticWickTerm =
sign φs φsΛ • (↑φsΛ.staticContract *
(contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedFieldOpEquiv φs φsΛ) (some k)) *
𝓝(ofFieldOpList (optionEraseZ [φsΛ]ᵘᶜ φ (uncontractedFieldOpEquiv φs φsΛ k))))) := by
symm
rw [staticWickTerm, normalOrder_uncontracted_some]
simp only [← mul_assoc]
rw [← smul_mul_assoc]
congr 1
rw [staticContract_insertAndContract_some_eq_mul_contractStateAtIndex_lt]
swap
· simp
rw [smul_smul]
by_cases hn : GradingCompliant φs φsΛ ∧ (𝓕|>ₛφ) = (𝓕|>ₛ φs[k.1])
· congr 1
swap
· have h1 := φsΛ.staticContract.2
rw [@Subalgebra.mem_center_iff] at h1
rw [h1]
erw [sign_insert_some]
rw [mul_assoc, mul_comm φsΛ.sign, ← mul_assoc]
rw [signInsertSome_mul_filter_contracted_of_not_lt]
simp only [instCommGroup.eq_1, Fin.zero_succAbove, Fin.not_lt_zero, Finset.filter_False,
ofFinset_empty, map_one, one_mul]
simp only [Fin.zero_succAbove, Fin.not_lt_zero, not_false_eq_true]
exact hn
· simp only [Fin.getElem_fin, not_and] at hn
by_cases h0 : ¬ GradingCompliant φs φsΛ
· rw [staticContract_of_not_gradingCompliant]
simp only [ZeroMemClass.coe_zero, zero_mul, smul_zero, instCommGroup.eq_1, mul_zero]
exact h0
· simp_all only [Finset.mem_univ, not_not, instCommGroup.eq_1, forall_const]
have h1 : contractStateAtIndex φ [φsΛ]ᵘᶜ (uncontractedFieldOpEquiv φs φsΛ k) = 0 := by
simp only [contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply,
instCommGroup.eq_1, Fin.coe_cast, Fin.getElem_fin, smul_eq_zero]
right
simp only [uncontractedListGet, List.getElem_map,
uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem]
rw [superCommute_anPart_ofFieldOpF_diff_grade_zero]
exact hn
rw [h1]
simp
lemma mul_staticWickTerm_eq_sum (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) :
ofFieldOp φ * φsΛ.staticWickTerm =
∑ (k : Option φsΛ.uncontracted), (φsΛ ↩Λ φ 0 k).staticWickTerm := by
trans (φsΛ.sign • φsΛ.staticContract * (ofFieldOp φ * normalOrder (ofFieldOpList [φsΛ]ᵘᶜ)))
· have ht := Subalgebra.mem_center_iff.mp (Subalgebra.smul_mem (Subalgebra.center _)
(φsΛ.staticContract).2 φsΛ.sign)
conv_rhs => rw [← mul_assoc, ← ht]
simp [mul_assoc, staticWickTerm]
rw [ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum]
rw [Finset.mul_sum]
rw [uncontractedFieldOpEquiv_list_sum]
refine Finset.sum_congr rfl (fun n _ => ?_)
match n with
| none =>
simp only [contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_none', one_mul, Algebra.smul_mul_assoc, Nat.succ_eq_add_one,
Fin.zero_eta, Fin.val_zero, List.insertIdx_zero, staticContract_insertAndContract_none,
insertAndContract_uncontractedList_none_zero]
rw [staticWickTerm_insert_zero_none]
simp only [Algebra.smul_mul_assoc]
rfl
| some n =>
simp only [Algebra.smul_mul_assoc, Nat.succ_eq_add_one, Fin.zero_eta, Fin.val_zero,
List.insertIdx_zero]
rw [staticWickTerm_insert_zero_some]
end
end WickContraction

View file

@ -3,16 +3,15 @@ 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.FieldOpAlgebra.NormalOrder.WickContractions
import HepLean.PerturbationTheory.WickContraction.Sign.InsertNone
import HepLean.PerturbationTheory.WickContraction.Sign.InsertSome
import HepLean.PerturbationTheory.WickContraction.StaticContract
import HepLean.PerturbationTheory.FieldOpAlgebra.StaticWickTerm
/-!
# Static Wick's theorem
-/
namespace FieldSpecification
variable {𝓕 : FieldSpecification}
open FieldOpFreeAlgebra
@ -22,12 +21,6 @@ open WickContraction
open FieldStatistic
namespace FieldOpAlgebra
lemma static_wick_theorem_nil : ofFieldOpList [] = ∑ (φsΛ : WickContraction [].length),
φsΛ.sign (𝓕 := 𝓕) • φsΛ.staticContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ) := by
simp only [ofFieldOpList, ofFieldOpListF_nil, map_one, List.length_nil]
rw [sum_WickContraction_nil, uncontractedListGet, nil_zero_uncontractedList]
simp [sign, empty, staticContract]
/--
The static Wicks theorem states that
`φ₀…φₙ` is equal to
@ -42,9 +35,11 @@ The inductive step works as follows:
- It also uses `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum`
-/
theorem static_wick_theorem : (φs : List 𝓕.FieldOp) →
ofFieldOpList φs = ∑ (φsΛ : WickContraction φs.length),
φsΛ.sign • φsΛ.staticContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
| [] => static_wick_theorem_nil
ofFieldOpList φs = ∑ (φsΛ : WickContraction φs.length), φsΛ.staticWickTerm
| [] => by
simp only [ofFieldOpList, ofFieldOpListF_nil, map_one, List.length_nil]
rw [sum_WickContraction_nil]
simp
| φ :: φs => by
rw [ofFieldOpList_cons, static_wick_theorem φs]
rw [show (φ :: φs) = φs.insertIdx (⟨0, Nat.zero_lt_succ φs.length⟩ : Fin φs.length.succ) φ
@ -53,66 +48,8 @@ theorem static_wick_theorem : (φs : List 𝓕.FieldOp) →
rw [Finset.mul_sum]
apply Finset.sum_congr rfl
intro c _
trans (sign φs c • ↑c.staticContract * (ofFieldOp φ * normalOrder (ofFieldOpList [c]ᵘᶜ)))
· have ht := Subalgebra.mem_center_iff.mp (Subalgebra.smul_mem (Subalgebra.center _)
(c.staticContract).2 c.sign)
conv_rhs => rw [← mul_assoc, ← ht]
simp [mul_assoc]
rw [ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum]
rw [Finset.mul_sum]
rw [uncontractedFieldOpEquiv_list_sum]
refine Finset.sum_congr rfl (fun n _ => ?_)
match n with
| none =>
simp only [contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_none', one_mul, Algebra.smul_mul_assoc, Nat.succ_eq_add_one,
Fin.zero_eta, Fin.val_zero, List.insertIdx_zero, staticContract_insertAndContract_none,
insertAndContract_uncontractedList_none_zero]
erw [sign_insert_none_zero]
rfl
| some n =>
simp only [Algebra.smul_mul_assoc, Nat.succ_eq_add_one, Fin.zero_eta, Fin.val_zero,
List.insertIdx_zero]
rw [normalOrder_uncontracted_some]
simp only [← mul_assoc]
rw [← smul_mul_assoc]
conv_rhs => rw [← smul_mul_assoc]
congr 1
rw [staticConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt]
swap
· simp
rw [smul_smul]
by_cases hn : GradingCompliant φs c ∧ (𝓕|>ₛφ) = (𝓕|>ₛ φs[n.1])
· congr 1
swap
· have h1 := c.staticContract.2
rw [@Subalgebra.mem_center_iff] at h1
rw [h1]
erw [sign_insert_some]
rw [mul_assoc, mul_comm c.sign, ← mul_assoc]
rw [signInsertSome_mul_filter_contracted_of_not_lt]
simp only [instCommGroup.eq_1, Fin.zero_succAbove, Fin.not_lt_zero, Finset.filter_False,
ofFinset_empty, map_one, one_mul]
simp only [Fin.zero_succAbove, Fin.not_lt_zero, not_false_eq_true]
exact hn
· simp only [Fin.getElem_fin, not_and] at hn
by_cases h0 : ¬ GradingCompliant φs c
· rw [staticContract_of_not_gradingCompliant]
simp only [ZeroMemClass.coe_zero, zero_mul, smul_zero, instCommGroup.eq_1, mul_zero]
exact h0
· simp_all only [Finset.mem_univ, not_not, instCommGroup.eq_1, forall_const]
have h1 : contractStateAtIndex φ [c]ᵘᶜ
((uncontractedFieldOpEquiv φs c) (some n)) = 0 := by
simp only [contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply,
instCommGroup.eq_1, Fin.coe_cast, Fin.getElem_fin, smul_eq_zero]
right
simp only [uncontractedListGet, List.getElem_map,
uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem]
rw [superCommute_anPart_ofFieldOpF_diff_grade_zero]
exact hn
rw [h1]
simp
rw [mul_staticWickTerm_eq_sum]
rfl
end FieldOpAlgebra
end FieldSpecification

View file

@ -29,6 +29,12 @@ noncomputable section
def wickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : 𝓕.FieldOpAlgebra :=
φsΛ.sign • φsΛ.timeContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
@[simp]
lemma wickTerm_empty_nil :
wickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
rw [wickTerm]
simp [sign_empty]
/--
Let `φsΛ` be a Wick Contraction for `φs = φ₀φ₁…φₙ`. Then the wick-term of ` (φsΛ ↩Λ φ i none)`
@ -48,20 +54,20 @@ The proof of this result relies primarily on:
- `sign_insert_none` and `signInsertNone_eq_filterset` which are used to take account of
signs.
-/
lemma wickTerm_none_eq_wick_term_cons (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
lemma wickTerm_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(i : Fin φs.length.succ) (φsΛ : WickContraction φs.length) :
(φsΛ ↩Λ φ i none).wickTerm =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (Finset.univ.filter (fun k => i.succAbove k < i))⟩)
• (φsΛ.sign • φsΛ.timeContract * 𝓝(ofFieldOpList (φ :: [φsΛ]ᵘᶜ))) := by
rw [wickTerm]
by_cases hg : GradingCompliant φs φsΛ
· rw [normalOrder_uncontracted_none, sign_insert_none]
· rw [normalOrder_uncontracted_none, sign_insert_none _ _ _ _ hg]
simp only [Nat.succ_eq_add_one, timeContract_insertAndContract_none, instCommGroup.eq_1,
Algebra.mul_smul_comm, Algebra.smul_mul_assoc, smul_smul]
congr 1
rw [← mul_assoc]
congr 1
rw [signInsertNone_eq_filterset _ _ _ _ hg, ← map_mul]
rw [← map_mul]
congr
rw [ofFinset_union]
congr
@ -96,7 +102,7 @@ This lemma states that
for `c` equal to `c ↩Λ φ i (some k)` is equal to that for just `c`
mulitiplied by the exchange sign of `φ` and `φ₀φ₁…φᵢ₋₁`.
-/
lemma wickTerm_some_eq_wick_term_optionEraseZ (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
lemma wickTerm_insert_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(i : Fin φs.length.succ) (φsΛ : WickContraction φs.length) (k : φsΛ.uncontracted)
(hlt : ∀ (k : Fin φs.length), timeOrderRel φ φs[k])
(hn : ∀ (k : Fin φs.length), i.succAbove k < i → ¬ timeOrderRel φs[k] φ) :
@ -108,29 +114,27 @@ lemma wickTerm_some_eq_wick_term_optionEraseZ (φ : 𝓕.FieldOp) (φs : List
rw [wickTerm]
by_cases hg : GradingCompliant φs φsΛ ∧ (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[k.1])
· by_cases hk : i.succAbove k < i
· rw [WickContraction.timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt]
· rw [WickContraction.timeContract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt]
swap
· exact hn _ hk
rw [normalOrder_uncontracted_some, sign_insert_some]
simp only [instCommGroup.eq_1, smul_smul, Algebra.smul_mul_assoc]
congr 1
rw [mul_assoc, mul_comm (sign φs φsΛ), ← mul_assoc]
congr 1
exact signInsertSome_mul_filter_contracted_of_lt φ φs φsΛ i k hk hg
· rw [normalOrder_uncontracted_some, sign_insert_some_of_lt φ φs φsΛ i k hk hg]
simp only [instCommGroup.eq_1, smul_smul, Algebra.smul_mul_assoc]
congr 1
rw [mul_assoc, mul_assoc, mul_comm, mul_assoc, mul_assoc]
simp
· omega
· have hik : i.succAbove ↑k ≠ i := Fin.succAbove_ne i ↑k
rw [timeContract_insertAndContract_some_eq_mul_contractStateAtIndex_lt]
swap
· exact hlt _
rw [normalOrder_uncontracted_some]
rw [sign_insert_some]
simp only [instCommGroup.eq_1, smul_smul, Algebra.smul_mul_assoc]
congr 1
rw [mul_assoc, mul_comm (sign φs φsΛ), ← mul_assoc]
congr 1
exact signInsertSome_mul_filter_contracted_of_not_lt φ φs φsΛ i k hk hg
· rw [normalOrder_uncontracted_some]
rw [sign_insert_some_of_not_lt φ φs φsΛ i k hk hg]
simp only [instCommGroup.eq_1, smul_smul, Algebra.smul_mul_assoc]
congr 1
rw [mul_assoc, mul_assoc, mul_comm, mul_assoc, mul_assoc]
simp
· omega
· rw [timeConract_insertAndContract_some]
· rw [timeContract_insertAndContract_some]
simp only [Fin.getElem_fin, not_and] at hg
by_cases hg' : GradingCompliant φs φsΛ
· have hg := hg hg'
@ -172,7 +176,7 @@ The proof of this result primarily depends on
- `wick_term_none_eq_wick_term_cons`
- `wick_term_some_eq_wick_term_optionEraseZ`
-/
lemma wickTerm_cons_eq_sum_wick_term (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (i : Fin φs.length.succ)
lemma mul_wickTerm_eq_sum (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (i : Fin φs.length.succ)
(φsΛ : WickContraction φs.length) (hlt : ∀ (k : Fin φs.length), timeOrderRel φ φs[k])
(hn : ∀ (k : Fin φs.length), i.succAbove k < i → ¬timeOrderRel φs[k] φ) :
ofFieldOp φ * φsΛ.wickTerm =
@ -190,7 +194,7 @@ lemma wickTerm_cons_eq_sum_wick_term (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldO
funext n
match n with
| none =>
rw [wickTerm_none_eq_wick_term_cons]
rw [wickTerm_insert_none]
simp only [contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_none', one_mul, Algebra.smul_mul_assoc, instCommGroup.eq_1,
smul_smul]
@ -198,7 +202,7 @@ lemma wickTerm_cons_eq_sum_wick_term (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldO
rw [← mul_assoc, exchangeSign_mul_self]
simp
| some n =>
rw [wickTerm_some_eq_wick_term_optionEraseZ _ _ _ _ _
rw [wickTerm_insert_some _ _ _ _ _
(fun k => hlt k) (fun k a => hn k a)]
simp only [uncontractedFieldOpEquiv, Equiv.optionCongr_apply, Equiv.coe_trans, Option.map_some',
Function.comp_apply, finCongr_apply, Algebra.smul_mul_assoc, instCommGroup.eq_1, smul_smul]

View file

@ -40,28 +40,11 @@ open FieldStatistic
-/
/-- Wick's theorem for the empty list. -/
lemma wicks_theorem_nil :
𝓣(ofFieldOpList (𝓕 := 𝓕) []) = ∑ (nilΛ : WickContraction [].length),
(nilΛ.sign (𝓕 := 𝓕) • nilΛ.timeContract) * 𝓝(ofFieldOpList [nilΛ]ᵘᶜ) := by
rw [timeOrder_ofFieldOpList_nil]
simp only [map_one, List.length_nil, Algebra.smul_mul_assoc]
rw [sum_WickContraction_nil, uncontractedListGet, nil_zero_uncontractedList]
simp only [List.map_nil]
have h1 : ofFieldOpList (𝓕 := 𝓕) [] = ofCrAnFieldOpList [] := by
rw [ofFieldOpList, ofCrAnFieldOpList]
simp
rw [h1, normalOrder_ofCrAnFieldOpList]
simp only [sign, List.length_nil, empty, Finset.univ_eq_empty, instCommGroup.eq_1,
Fin.getElem_fin, Finset.prod_empty, WickContraction.timeContract, List.get_eq_getElem,
OneMemClass.coe_one, normalOrderSign_nil, normalOrderList_nil, one_smul, one_mul]
rfl
lemma wicks_theorem_congr {φs φs' : List 𝓕.FieldOp} (h : φs = φs') :
∑ (φsΛ : WickContraction φs.length), φsΛ.wickTerm
= ∑ (φs'Λ : WickContraction φs'.length), φs'Λ.wickTerm := by
subst h
simp
rfl
remark wicks_theorem_context := "
In perturbation quantum field theory, Wick's theorem allows
@ -93,7 +76,11 @@ The inductive step works as follows:
-/
theorem wicks_theorem : (φs : List 𝓕.FieldOp) → 𝓣(ofFieldOpList φs) =
∑ (φsΛ : WickContraction φs.length), φsΛ.wickTerm
| [] => wicks_theorem_nil
| [] => by
rw [timeOrder_ofFieldOpList_nil]
simp only [map_one, List.length_nil, Algebra.smul_mul_assoc]
rw [sum_WickContraction_nil]
simp
| φ :: φs => by
have ih := wicks_theorem (eraseMaxTimeField φ φs)
conv_lhs => rw [timeOrder_eq_maxTimeField_mul_finset, ih, Finset.mul_sum]
@ -106,7 +93,7 @@ theorem wicks_theorem : (φs : List 𝓕.FieldOp) → 𝓣(ofFieldOpList φs) =
conv_rhs => rw [insertLift_sum]
apply Finset.sum_congr rfl
intro c _
rw [Algebra.smul_mul_assoc, wickTerm_cons_eq_sum_wick_term
rw [Algebra.smul_mul_assoc, mul_wickTerm_eq_sum
(maxTimeField φ φs) (eraseMaxTimeField φ φs) (maxTimeFieldPosFin φ φs) c]
trans (1 : ) • ∑ k : Option { x // x ∈ c.uncontracted },
(c ↩Λ (maxTimeField φ φs) (maxTimeFieldPosFin φ φs) k).wickTerm

View file

@ -31,6 +31,7 @@ lemma timeOrder_ofFieldOpList_eqTimeOnly (φs : List 𝓕.FieldOp) :
rw [← e2.symm.sum_comp]
simp only [Equiv.symm_symm, Algebra.smul_mul_assoc, Fintype.sum_sum_type,
Equiv.sumCompl_apply_inl, Equiv.sumCompl_apply_inr, map_add, map_sum, map_smul, e2]
simp only [staticWickTerm, Algebra.smul_mul_assoc, map_smul]
conv_lhs =>
enter [2, 2, x]
rw [timeOrder_timeOrder_left]

View file

@ -29,7 +29,7 @@ def signFinset (c : WickContraction n) (i1 i2 : Fin n) : Finset (Fin n) :=
(c.getDual? i = none ∀ (h : (c.getDual? i).isSome), i1 < (c.getDual? i).get h))
/-- Given a Wick contraction `φsΛ` associated with a list of states `φs`
the sign associated with `φsΛ` is sign corresponding to the number
the sign associated with `φsΛ` is the sign corresponding to the number
of fermionic-fermionic exchanges one must do to put elements in contracted pairs
of `φsΛ` next to each other. -/
def sign (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) : :=

View file

@ -87,8 +87,8 @@ def signInsertNone (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickCo
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
else 1
lemma sign_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) :
lemma sign_insert_none_eq_signInsertNone_mul_sign (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
(φsΛ ↩Λ φ i none).sign = (φsΛ.signInsertNone φ φs i) * φsΛ.sign := by
rw [sign]
rw [signInsertNone, sign, ← Finset.prod_mul_distrib]
@ -159,7 +159,7 @@ lemma signInsertNone_eq_prod_prod (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
lemma sign_insert_none_zero (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) : (φsΛ ↩Λ φ 0 none).sign = φsΛ.sign := by
rw [sign_insert_none]
rw [sign_insert_none_eq_signInsertNone_mul_sign]
simp [signInsertNone]
lemma signInsertNone_eq_prod_getDual?_Some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
@ -243,4 +243,16 @@ lemma signInsertNone_eq_filterset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
simp [h]
· exact hG
/-- For `φsΛ` a grading compliant Wick contraction, and `i : Fin φs.length.succ` we have
`(φsΛ ↩Λ φ i none).sign = s * φsΛ.sign`
where `s` is the sign got by moving `φ` through the elements of `φ₀…φᵢ₋₁` which
are contracted. -/
lemma sign_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
(φsΛ ↩Λ φ i none).sign = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, Finset.univ.filter
(fun x => (φsΛ.getDual? x).isSome ∧ i.succAbove x < i)⟩) * φsΛ.sign := by
rw [sign_insert_none_eq_signInsertNone_mul_sign]
rw [signInsertNone_eq_filterset]
exact hG
end WickContraction

View file

@ -9,6 +9,8 @@ import HepLean.PerturbationTheory.WickContraction.Sign.Basic
# Sign on inserting and contracting
The main results of this file are `sign_insert_some_of_lt` and `sign_insert_some_of_not_lt` which
write the sign of `(φsΛ ↩Λ φ i (some k)).sign` in terms of the sign of `φsΛ`.
-/
open FieldSpecification
@ -852,4 +854,51 @@ lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.FieldOp) (φs :
forall_const]
exact fun h => lt_of_le_of_ne h (Fin.succAbove_ne i ((φsΛ.getDual? j).get hj))
/--
For `k < i`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ`,
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
- the sign of `φsΛ`.
The proof of this result involves a careful consideration of the contributions of different
fields in `φs` to the sign of `φsΛ ↩Λ φ i (some k)`.
-/
lemma sign_insert_some_of_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
(hk : i.succAbove k < i) (hg : GradingCompliant φs φsΛ ∧ (𝓕 |>ₛ φ) = 𝓕 |>ₛ φs[k.1]) :
(φsΛ ↩Λ φ i (some k)).sign =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, φsΛ.uncontracted.filter (fun x => x ≤ ↑k)⟩)
* 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, Finset.univ.filter (fun x => i.succAbove x < i)⟩)
* φsΛ.sign := by
rw [sign_insert_some,
← signInsertSome_mul_filter_contracted_of_lt φ φs φsΛ i k hk hg]
rw [← mul_assoc]
congr 1
rw [mul_comm, ← mul_assoc]
simp
/--
For `i ≤ k`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ₋₁`,
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
- the sign of `φsΛ`.
The proof of this result involves a careful consideration of the contributions of different
fields in `φs` to the sign of `φsΛ ↩Λ φ i (some k)`.
-/
lemma sign_insert_some_of_not_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
(hk : ¬ i.succAbove k < i) (hg : GradingCompliant φs φsΛ ∧ (𝓕 |>ₛ φ) = 𝓕 |>ₛ φs[k.1]) :
(φsΛ ↩Λ φ i (some k)).sign =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, φsΛ.uncontracted.filter (fun x => x < ↑k)⟩)
* 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, Finset.univ.filter (fun x => i.succAbove x < i)⟩) *
φsΛ.sign := by
rw [sign_insert_some,
← signInsertSome_mul_filter_contracted_of_not_lt φ φs φsΛ i k hk hg]
rw [← mul_assoc]
congr 1
rw [mul_comm, ← mul_assoc]
simp
end WickContraction

View file

@ -66,7 +66,7 @@ lemma staticContract_insertAndContract_some
open FieldStatistic
lemma staticConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
lemma staticContract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
(hik : i < i.succAbove k) :

View file

@ -50,7 +50,7 @@ lemma timeContract_insertAndContract_none (φ : 𝓕.FieldOp) (φs : List 𝓕.F
This follows from the fact that `(φsΛ ↩Λ φ i (some j))` has one more contracted pair than `φsΛ`,
corresponding to `φ` contracted with `φⱼ`. The order depends on whether we insert `φ` before
or after `φⱼ`. -/
lemma timeConract_insertAndContract_some
lemma timeContract_insertAndContract_some
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
(φsΛ ↩Λ φ i (some j)).timeContract =
@ -85,7 +85,7 @@ lemma timeContract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x < k))⟩)
• (contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedFieldOpEquiv φs φsΛ) (some k)) *
φsΛ.timeContract) := by
rw [timeConract_insertAndContract_some]
rw [timeContract_insertAndContract_some]
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
@ -110,7 +110,7 @@ lemma timeContract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
simp only [exchangeSign_mul_self]
· exact ht
lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt
lemma timeContract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
(ht : ¬ 𝓕.timeOrderRel φs[k.1] φ) (hik : ¬ i < i.succAbove k) :
@ -118,7 +118,7 @@ lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x ≤ k))⟩)
• (contractStateAtIndex φ [φsΛ]ᵘᶜ
((uncontractedFieldOpEquiv φs φsΛ) (some k)) * φsΛ.timeContract) := by
rw [timeConract_insertAndContract_some]
rw [timeContract_insertAndContract_some]
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,

View file

@ -162,6 +162,8 @@ def perturbationTheory : Note where
.h2 "Definition",
.name `WickContraction,
.name `WickContraction.GradingCompliant,
.h2 "Cardinality",
.name `WickContraction.card_eq_cardFun,
.h2 "Constructors",
.p "There are a number of ways to construct a Wick contraction from
other Wick contractions or single contractions.",
@ -171,12 +173,12 @@ def perturbationTheory : Note where
.h2 "Sign",
.name `WickContraction.sign,
.name `WickContraction.join_sign,
.name `WickContraction.signInsertNone_eq_filterset,
.name `WickContraction.signInsertSome_mul_filter_contracted_of_not_lt,
.h2 "Cardinality",
.name `WickContraction.card_eq_cardFun,
.name `WickContraction.sign_insert_none,
.name `WickContraction.sign_insert_some_of_not_lt,
.name `WickContraction.sign_insert_some_of_lt,
.h1 "Time and static contractions",
.h1 "Useful results",
.h1 "Wick terms",
.name `WickContraction.wickTerm,
.h1 "The three Wick's theorems",
.name `FieldSpecification.wicks_theorem,
.name `FieldSpecification.FieldOpAlgebra.static_wick_theorem,