PhysLean/HepLean/PerturbationTheory/FieldOpAlgebra/WicksTheoremNormal.lean

258 lines
11 KiB
Text
Raw Normal View History

2025-02-01 11:51:06 +00:00
/-
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
-/
2025-02-03 12:12:36 +00:00
import HepLean.PerturbationTheory.FieldOpAlgebra.StaticWickTheorem
2025-02-05 08:52:14 +00:00
import HepLean.PerturbationTheory.FieldOpAlgebra.WicksTheorem
import HepLean.PerturbationTheory.WickContraction.Sign.Join
import HepLean.PerturbationTheory.WickContraction.TimeCond
2025-02-01 11:51:06 +00:00
/-!
# Wick's theorem for normal ordered lists
-/
namespace FieldSpecification
variable {𝓕 : FieldSpecification}
2025-02-03 11:05:43 +00:00
open FieldOpFreeAlgebra
2025-02-01 11:51:06 +00:00
namespace FieldOpAlgebra
open WickContraction
open EqTimeOnly
2025-02-07 09:56:37 +00:00
/--
For a list `φs` of `𝓕.FieldOp`, then
`𝓣(φs) = ∑ φsΛ, φsΛ.1.wickTerm • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))`
where the sum is over all Wick contraction `φsΛ` which only have equal time contractions.
This result follows from
- `static_wick_theorem` to rewrite `𝓣(φs)` on the left hand side as a sum of
`𝓣(φsΛ.staticWickTerm)`.
2025-02-07 10:34:48 +00:00
- `EqTimeOnly.timeOrder_staticContract_of_not_mem` and `timeOrder_timeOrder_mid` to set to
those `𝓣(φsΛ.staticWickTerm)` for which `φsΛ` has a contracted pair which are not
equal time to zero.
- `staticContract_eq_timeContract_of_eqTimeOnly` to rewrite the static contract
2025-02-10 10:51:44 +00:00
in the remaining `𝓣(φsΛ.staticWickTerm)` as a time contract.
2025-02-07 09:56:37 +00:00
- `timeOrder_timeContract_mul_of_eqTimeOnly_left` to move the time contracts out of the time
ordering.
-/
2025-02-03 11:28:14 +00:00
lemma timeOrder_ofFieldOpList_eqTimeOnly (φs : List 𝓕.FieldOp) :
2025-02-05 08:52:14 +00:00
𝓣(ofFieldOpList φs) = ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs)}),
2025-02-03 06:13:13 +00:00
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
2025-02-01 11:51:06 +00:00
rw [static_wick_theorem φs]
2025-02-03 05:39:48 +00:00
let e2 : WickContraction φs.length ≃
{φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} ⊕
2025-02-03 06:13:13 +00:00
{φsΛ : WickContraction φs.length // ¬ φsΛ.EqTimeOnly} :=
2025-02-03 05:39:48 +00:00
(Equiv.sumCompl _).symm
2025-02-01 11:51:06 +00:00
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]
2025-02-05 10:01:48 +00:00
simp only [staticWickTerm, Algebra.smul_mul_assoc, map_smul]
2025-02-01 11:51:06 +00:00
conv_lhs =>
enter [2, 2, x]
rw [timeOrder_timeOrder_left]
rw [timeOrder_staticContract_of_not_mem _ x.2]
2025-02-03 05:39:48 +00:00
simp only [Finset.univ_eq_attach, zero_mul, map_zero, smul_zero, Finset.sum_const_zero, add_zero]
2025-02-01 11:51:06 +00:00
congr
funext x
2025-02-03 05:39:48 +00:00
rw [staticContract_eq_timeContract_of_eqTimeOnly]
rw [timeOrder_timeContract_mul_of_eqTimeOnly_left]
exact x.2
2025-02-01 11:51:06 +00:00
exact x.2
2025-02-03 11:28:14 +00:00
lemma timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.FieldOp) :
2025-02-07 09:56:37 +00:00
𝓣(ofFieldOpList φs) = 𝓣(𝓝(ofFieldOpList φs)) +
2025-02-03 05:39:48 +00:00
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
2025-02-03 06:13:13 +00:00
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
2025-02-03 05:39:48 +00:00
let e1 : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} ≃
{φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // φsΛ.1 = empty} ⊕
{φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // ¬ φsΛ.1 = empty} :=
2025-02-03 06:13:13 +00:00
(Equiv.sumCompl _).symm
2025-02-01 11:51:06 +00:00
rw [timeOrder_ofFieldOpList_eqTimeOnly, ← e1.symm.sum_comp]
2025-02-03 05:39:48 +00:00
simp only [Equiv.symm_symm, Algebra.smul_mul_assoc, Fintype.sum_sum_type,
Equiv.sumCompl_apply_inl, Equiv.sumCompl_apply_inr, ne_eq, e1]
2025-02-01 11:51:06 +00:00
congr 1
2025-02-03 06:13:13 +00:00
· let e2 : {φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // φsΛ.1 = empty } ≃
Unit := {
2025-02-01 11:51:06 +00:00
toFun := fun x => (), invFun := fun x => ⟨⟨empty, by simp⟩, rfl⟩,
left_inv a := by
ext
2025-02-03 06:13:13 +00:00
simp [a.2], right_inv a := by simp}
2025-02-01 11:51:06 +00:00
rw [← e2.symm.sum_comp]
simp [e2, sign_empty]
2025-02-03 05:39:48 +00:00
· let e2 : { φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} // ¬ φsΛ.1 = empty } ≃
{φsΛ // φsΛ.EqTimeOnly ∧ φsΛ ≠ empty} := {
2025-02-01 11:51:06 +00:00
toFun := fun x => ⟨x, ⟨x.1.2, x.2⟩⟩, invFun := fun x => ⟨⟨x.1, x.2.1⟩, x.2.2⟩,
left_inv a := by rfl, right_inv a := by rfl }
rw [← e2.symm.sum_comp]
rfl
2025-02-07 09:56:37 +00:00
/--
For a list `φs` of `𝓕.FieldOp`, then
`𝓣(𝓝(φs)) = 𝓣(φs) - ∑ φsΛ, φsΛ.1.wickTerm • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))`
where the sum is over all *non-empty* Wick contraction `φsΛ` which only
have equal time contractions.
This result follows directly from
- `timeOrder_ofFieldOpList_eqTimeOnly`
-/
2025-02-03 11:28:14 +00:00
lemma normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.FieldOp) :
2025-02-03 06:13:13 +00:00
𝓣(𝓝(ofFieldOpList φs)) = 𝓣(ofFieldOpList φs) -
2025-02-03 05:39:48 +00:00
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
2025-02-03 06:13:13 +00:00
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
2025-02-01 11:51:06 +00:00
rw [timeOrder_ofFieldOpList_eq_eqTimeOnly_empty]
simp
2025-02-07 09:56:37 +00:00
/--
For a list `φs` of `𝓕.FieldOp`, then `𝓣(φs)` is equal to the sum of
- `∑ φsΛ, φsΛ.wickTerm` where the sum is over all Wick contraction `φsΛ` which have
no contractions of equal time.
- `∑ φsΛ, φsΛ.sign • φsΛ.timeContract * (∑ φssucΛ, φssucΛ.wickTerm)`, where
2025-02-07 09:56:37 +00:00
the first sum is over all Wick contraction `φsΛ` which only have equal time contractions
and the second sum is over all Wick contraction `φssucΛ` of the uncontracted elements of `φsΛ`
which do not have any equal time contractions.
The proof of proceeds as follows
- `wicks_theorem` is used to rewrite `𝓣(φs)` as a sum over all Wick contractions.
- The sum over all Wick contractions is then split additively into two parts using based on having
or not having an equal time contractions.
- Using `join`, the sum `∑ φsΛ, _` over Wick contractions which do have equal time contractions
is split into two sums `∑ φsΛ, ∑ φsucΛ, _`, the first over non-zero elements
which only have equal time contractions and the second over Wick contractions `φsucΛ` of
`[φsΛ]ᵘᶜ` which do not have equal time contractions.
- `join_sign_timeContract` is then used to equate terms.
2025-02-07 09:56:37 +00:00
-/
lemma timeOrder_haveEqTime_split (φs : List 𝓕.FieldOp) :
2025-02-07 10:36:35 +00:00
𝓣(ofFieldOpList φs) = (∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
2025-02-03 06:13:13 +00:00
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
2025-02-07 10:36:35 +00:00
+ ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}), φsΛ.1.sign • φsΛ.1.timeContract *
2025-02-07 10:34:48 +00:00
(∑ φssucΛ : { φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ φssucΛ.HaveEqTime },
φssucΛ.1.wickTerm) := by
2025-02-01 11:51:06 +00:00
rw [wicks_theorem]
2025-02-05 08:52:14 +00:00
simp only [wickTerm]
2025-02-03 06:13:13 +00:00
let e1 : WickContraction φs.length ≃ {φsΛ // HaveEqTime φsΛ} ⊕ {φsΛ // ¬ HaveEqTime φsΛ} := by
2025-02-01 11:51:06 +00:00
exact (Equiv.sumCompl HaveEqTime).symm
rw [← e1.symm.sum_comp]
2025-02-03 05:39:48 +00:00
simp only [Equiv.symm_symm, Algebra.smul_mul_assoc, Fintype.sum_sum_type,
Equiv.sumCompl_apply_inl, Equiv.sumCompl_apply_inr, ne_eq, sub_left_inj, e1]
2025-02-01 11:51:06 +00:00
rw [add_comm]
2025-02-07 09:56:37 +00:00
congr 1
2025-02-01 11:51:06 +00:00
let f : WickContraction φs.length → 𝓕.FieldOpAlgebra := fun φsΛ =>
2025-02-07 09:56:37 +00:00
φsΛ.sign • (φsΛ.timeContract.1 * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ))
2025-02-01 11:51:06 +00:00
change ∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}), f φsΛ.1 = _
rw [sum_haveEqTime]
congr
funext φsΛ
2025-02-03 06:13:13 +00:00
simp only [f]
2025-02-01 11:51:06 +00:00
conv_lhs =>
enter [2, φsucΛ]
2025-02-07 09:56:37 +00:00
rw [← Algebra.smul_mul_assoc]
2025-02-01 11:51:06 +00:00
rw [join_sign_timeContract φsΛ.1 φsucΛ.1]
conv_lhs =>
enter [2, φsucΛ]
rw [mul_assoc]
2025-02-07 09:56:37 +00:00
rw [← Finset.mul_sum, ← Algebra.smul_mul_assoc]
2025-02-01 11:51:06 +00:00
congr
funext φsΛ'
2025-02-03 05:39:48 +00:00
simp only [ne_eq, Algebra.smul_mul_assoc]
2025-02-01 11:51:06 +00:00
congr 1
rw [@join_uncontractedListGet]
2025-02-03 11:28:14 +00:00
lemma normalOrder_timeOrder_ofFieldOpList_eq_not_haveEqTime_sub_inductive (φs : List 𝓕.FieldOp) :
2025-02-07 10:34:48 +00:00
𝓣(𝓝(ofFieldOpList φs)) =
(∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}), φsΛ.1.wickTerm)
2025-02-03 05:39:48 +00:00
+ ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
2025-02-01 11:51:06 +00:00
sign φs ↑φsΛ • (φsΛ.1).timeContract *
(∑ φssucΛ : { φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ φssucΛ.HaveEqTime },
2025-02-07 10:34:48 +00:00
φssucΛ.1.wickTerm - 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))) := by
2025-02-07 09:56:37 +00:00
rw [normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty,
timeOrder_haveEqTime_split]
2025-02-01 11:51:06 +00:00
rw [add_sub_assoc]
congr 1
2025-02-03 05:39:48 +00:00
simp only [ne_eq, Algebra.smul_mul_assoc]
2025-02-01 11:51:06 +00:00
rw [← Finset.sum_sub_distrib]
congr 1
funext x
2025-02-03 05:39:48 +00:00
simp only
2025-02-01 11:51:06 +00:00
rw [← smul_sub, ← mul_sub]
2025-02-03 06:13:13 +00:00
lemma wicks_theorem_normal_order_empty : 𝓣(𝓝(ofFieldOpList [])) =
2025-02-03 11:28:14 +00:00
∑ (φsΛ : {φsΛ : WickContraction ([] : List 𝓕.FieldOp).length // ¬ HaveEqTime φsΛ}),
2025-02-05 08:52:14 +00:00
φsΛ.1.wickTerm := by
simp only [wickTerm]
2025-02-03 11:28:14 +00:00
let e2 : {φsΛ : WickContraction ([] : List 𝓕.FieldOp).length // ¬ HaveEqTime φsΛ} ≃ Unit :=
2025-02-01 11:51:06 +00:00
{
toFun := fun x => (),
invFun := fun x => ⟨empty, by simp⟩,
left_inv := by
intro a
2025-02-03 05:39:48 +00:00
simp only [List.length_nil]
2025-02-01 11:51:06 +00:00
apply Subtype.eq
apply Subtype.eq
2025-02-03 05:39:48 +00:00
simp only [empty]
2025-02-01 11:51:06 +00:00
ext i
2025-02-03 05:39:48 +00:00
simp only [Finset.not_mem_empty, false_iff]
2025-02-01 11:51:06 +00:00
by_contra hn
have h2 := a.1.2.1 i hn
rw [@Finset.card_eq_two] at h2
obtain ⟨a, b, ha, hb, hab⟩ := h2
exact Fin.elim0 a,
right_inv := by intro a; simp}
rw [← e2.symm.sum_comp]
2025-02-03 05:39:48 +00:00
simp only [Finset.univ_unique, PUnit.default_eq_unit, List.length_nil, Equiv.coe_fn_symm_mk,
sign_empty, timeContract_empty, OneMemClass.coe_one, one_smul, uncontractedListGet_empty,
one_mul, Finset.sum_const, Finset.card_singleton, e2]
have h1' : ofFieldOpList (𝓕 := 𝓕) [] = ofCrAnList [] := by rfl
2025-02-01 11:51:06 +00:00
rw [h1']
rw [normalOrder_ofCrAnList]
2025-02-03 05:39:48 +00:00
simp only [normalOrderSign_nil, normalOrderList_nil, one_smul]
rw [ofCrAnList, timeOrder_eq_ι_timeOrderF]
rw [timeOrderF_ofCrAnListF]
2025-02-01 11:51:06 +00:00
simp
2025-02-07 09:56:37 +00:00
/--For a list `φs` of `𝓕.FieldOp`, the normal-ordered version of Wick's theorem states that
`𝓣(𝓝(φs)) = ∑ φsΛ, φsΛ.wickTerm`
where the sum is over all Wick contraction `φsΛ` in which no two contracted elements
have the same time.
2025-02-07 10:34:48 +00:00
The proof of proceeds by induction on `φs`, with the base case `[]` holding by following
2025-02-07 09:56:37 +00:00
through definitions. and the inductive case holding as a result of
- `timeOrder_haveEqTime_split`
- `normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty`
- and the induction hypothesis on `𝓣(𝓝([φsΛ.1]ᵘᶜ))` for contractions `φsΛ` of `φs` which only
have equal time contractions and are non-empty.
2025-02-03 10:47:18 +00:00
-/
2025-02-03 11:28:14 +00:00
theorem wicks_theorem_normal_order : (φs : List 𝓕.FieldOp) →
2025-02-05 08:52:14 +00:00
𝓣(𝓝(ofFieldOpList φs)) =
∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}), φsΛ.1.wickTerm
2025-02-01 11:51:06 +00:00
| [] => wicks_theorem_normal_order_empty
| φ :: φs => by
rw [normalOrder_timeOrder_ofFieldOpList_eq_not_haveEqTime_sub_inductive]
2025-02-03 06:13:13 +00:00
simp only [Algebra.smul_mul_assoc, ne_eq, add_right_eq_self]
2025-02-01 11:51:06 +00:00
apply Finset.sum_eq_zero
intro φsΛ hφsΛ
simp only [smul_eq_zero]
right
have ih := wicks_theorem_normal_order [φsΛ.1]ᵘᶜ
rw [ih]
2025-02-05 08:52:14 +00:00
simp [wickTerm]
2025-02-01 11:51:06 +00:00
termination_by φs => φs.length
decreasing_by
simp only [uncontractedListGet, List.length_cons, List.length_map, gt_iff_lt]
rw [uncontractedList_length_eq_card]
have hc := uncontracted_card_eq_iff φsΛ.1
2025-02-03 05:39:48 +00:00
simp only [List.length_cons, φsΛ.2.2, iff_false] at hc
2025-02-01 11:51:06 +00:00
have hc' := uncontracted_card_le φsΛ.1
2025-02-03 05:39:48 +00:00
simp_all only [Algebra.smul_mul_assoc, List.length_cons, Finset.mem_univ, gt_iff_lt]
2025-02-01 11:51:06 +00:00
omega
end FieldOpAlgebra
end FieldSpecification