docs: Notes for normal-ordered Wicks

This commit is contained in:
jstoobysmith 2025-02-07 09:56:37 +00:00
parent c9607c459f
commit d472604aec
15 changed files with 324 additions and 131 deletions

View file

@ -347,7 +347,7 @@ noncomputable def contractStateAtIndex (φ : 𝓕.FieldOp) (φs : List 𝓕.Fiel
/--
For a field specification `𝓕`, a `φ` in `𝓕.FieldOp` and a list `φs` of `𝓕.FieldOp`
the following relation holds in the algebra `𝓕.FieldOpAlgebra`,
`φ * 𝓝(φ₀φ₁…φₙ) = 𝓝(φφ₀φ₁…φₙ) + ∑ i, (𝓢(φ,φ₀φ₁…φᵢ₋₁) • [anPartF φ, φᵢ]ₛ) * 𝓝(φ₀φ₁…φᵢ₋₁φᵢ₊₁…φₙ)`.
`φ * 𝓝(φ₀φ₁…φₙ) = 𝓝(φφ₀φ₁…φₙ) + ∑ i, (𝓢(φ,φ₀φ₁…φᵢ₋₁) • [anPart φ, φᵢ]ₛ) * 𝓝(φ₀φ₁…φᵢ₋₁φᵢ₊₁…φₙ)`.
The proof of ultimetly goes as follows:
- `ofFieldOp_eq_crPart_add_anPart` is used to split `φ` into its creation and annihilation parts.

View file

@ -101,7 +101,7 @@ lemma normalOrder_uncontracted_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp
/--
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
`𝓕.FieldOp`, a `i ≤ φs.length` and a `k` in `φsΛ.uncontracted`, then
`𝓝([φsΛ ↩Λ φ i none]ᵘᶜ)` is equal to the normal ordering of `[φsΛ]ᵘᶜ` with the `FieldOp`
`𝓝([φsΛ ↩Λ φ i (some k)]ᵘᶜ)` is equal to the normal ordering of `[φsΛ]ᵘᶜ` with the `FieldOp`
corresponding to `k` removed.
The proof of this result ultimately a consequence of definitions.

View file

@ -22,12 +22,17 @@ 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Λ]ᵘᶜ)`. -/
/-- For a list `φs` of `𝓕.FieldOp`, and a Wick contraction `φsΛ` of `φs`, the element
of `𝓕.FieldOpAlgebra`, `φsΛ.staticWickTerm` is defined as
`φsΛ.sign • φsΛ.staticContract * 𝓝([φsΛ]ᵘᶜ)`.
This is term which appears in the static version Wick's theorem. -/
def staticWickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : 𝓕.FieldOpAlgebra :=
φsΛ.sign • φsΛ.staticContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
/-- The static Wick term for the empty contraction of the empty list is `1`. -/
/-- For the empty list `[]` of `𝓕.FieldOp`, the `staticWickTerm` of the empty Wick contraction
`empty` of `[]` (its only Wick contraction) is `1`. -/
@[simp]
lemma staticWickTerm_empty_nil :
staticWickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
@ -35,7 +40,9 @@ lemma staticWickTerm_empty_nil :
simp [sign, empty, staticContract]
/--
Let `φsΛ` be a Wick Contraction for `φs = φ₀φ₁…φₙ`. Then the following holds
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, and an element `φ` of
`𝓕.FieldOp`, the following relation holds
`(φsΛ ↩Λ φ 0 none).staticWickTerm = φsΛ.sign • φsΛ.staticWickTerm * 𝓝(φ :: [φsΛ]ᵘᶜ)`
The proof of this result relies on
@ -51,8 +58,9 @@ lemma staticWickTerm_insert_zero_none (φ : 𝓕.FieldOp) (φs : List 𝓕.Field
simp only [staticContract_insert_none, insertAndContract_uncontractedList_none_zero,
Algebra.smul_mul_assoc]
/-- Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Then`(φsΛ ↩Λ φ 0 (some k)).wickTerm`
is equal the product of
/-- For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
`𝓕.FieldOp`, and a `k` in `φsΛ.uncontracted`, `(φsΛ ↩Λ φ 0 (some k)).wickTerm` is equal
to the product of
- the sign `𝓢(φ, φ₀…φᵢ₋₁) `
- the sign `φsΛ.sign`
- `φsΛ.staticContract`
@ -60,9 +68,8 @@ is equal the product of
uncontracted fields in `φ₀…φₖ₋₁`
- the normal ordering `𝓝([φsΛ]ᵘᶜ.erase (uncontractedFieldOpEquiv φs φsΛ k))`.
The proof of this result relies on
- `staticContract_insert_some_of_lt` to rewrite static
contractions.
The proof of this result ultimitley relies on
- `staticContract_insert_some` to rewrite static contractions.
- `normalOrder_uncontracted_some` to rewrite normal orderings.
- `sign_insert_some_zero` to rewrite signs.
-/
@ -106,13 +113,16 @@ lemma staticWickTerm_insert_zero_some (φ : 𝓕.FieldOp) (φs : List 𝓕.Field
simp
/--
Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Then
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, the following relation
holds
`φ * φsΛ.staticWickTerm = ∑ k, (φsΛ ↩Λ φ i k).wickTerm`
where the sum is over all `k` in `Option φsΛ.uncontracted` (so either `none` or `some k`).
The proof of proceeds as follows:
- `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum` is used to expand `φ 𝓝([φsΛ]ᵘᶜ)` as
a sum over `k` in `Option φsΛ.uncontracted` of terms involving `[φ, φs[k]]` etc.
a sum over `k` in `Option φsΛ.uncontracted` of terms involving `[anPart φ, φs[k]]`.
- Then `staticWickTerm_insert_zero_none` and `staticWickTerm_insert_zero_some` are
used to equate terms.
-/

View file

@ -20,24 +20,41 @@ open FieldStatistic
namespace FieldOpAlgebra
/--
The static Wicks theorem states that
`φ₀…φₙ` is equal to
`∑ φsΛ, φsΛ.1.sign • φsΛ.1.staticContract * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)`
over all Wick contraction `φsΛ`.
This is compared to the ordinary Wicks theorem in which `staticContract` is replaced with
`timeContract`.
For a list `φs` of `𝓕.FieldOp`, the static version of Wick's theorem states that
`φs =∑ φsΛ, φsΛ.staticWickTerm`
where the sum is over all Wick contraction `φsΛ`.
The proof is via induction on `φs`.
- The base case `φs = []` is handled by `staticWickTerm_empty_nil`.
The proof is via induction on `φs`. The base case `φs = []` is handled by `static_wick_theorem_nil`.
The inductive step works as follows:
- The proof considers `φ₀…φₙ` as `φ₀(φ₁…φₙ)` and use the induction hypothesis on `φ₁…φₙ`.
- It also uses `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum`
For the LHS:
1. The proof considers `φ₀…φₙ` as `φ₀(φ₁…φₙ)` and use the induction hypothesis on `φ₁…φₙ`.
2. This gives terms of the form `φ * φsΛ.staticWickTerm` on which
`mul_staticWickTerm_eq_sum` is used where `φsΛ` is a Wick contraction of `φ₁…φₙ`,
to rewrite terms as a sum over optional uncontracted elements of `φsΛ`
On the LHS we now have a sum over Wick contractions `φsΛ` of `φ₁…φₙ` (from 1) and optional
uncontracted elements of `φsΛ` (from 2)
For the RHS:
1. The sum over Wick contractions of `φ₀…φₙ` on the RHS
is split via `insertLift_sum` into a sum over Wick contractions `φsΛ` of `φ₁…φₙ` and
sum over optional uncontracted elements of `φsΛ`.
Both side now are sums over the same thing and their terms equate by the nature of the
lemmas used.
-/
theorem static_wick_theorem : (φs : List 𝓕.FieldOp) →
ofFieldOpList φs = ∑ (φsΛ : WickContraction φs.length), φsΛ.staticWickTerm
| [] => by
simp only [ofFieldOpList, ofFieldOpListF_nil, map_one, List.length_nil]
rw [sum_WickContraction_nil]
simp
rw [staticWickTerm_empty_nil]
| φ :: φ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) φ

View file

@ -23,15 +23,19 @@ namespace FieldOpAlgebra
open FieldStatistic
/-- The time contraction of two FieldOp as an element of `𝓞.A` defined
as their time ordering in the state algebra minus their normal ordering in the
creation and annihlation algebra, both mapped to `𝓞.A`.. -/
/-- For a field specification `𝓕`, and `φ` and `ψ` elements of `𝓕.FieldOp`, the element of
`𝓕.FieldOpAlgebra`, `timeContract φ ψ` is defined to be `𝓣(φψ)- 𝓝(φψ)`. -/
def timeContract (φ ψ : 𝓕.FieldOp) : 𝓕.FieldOpAlgebra :=
𝓣(ofFieldOp φ * ofFieldOp ψ) - 𝓝(ofFieldOp φ * ofFieldOp ψ)
lemma timeContract_eq_smul (φ ψ : 𝓕.FieldOp) : timeContract φ ψ =
𝓣(ofFieldOp φ * ofFieldOp ψ) + (-1 : ) • 𝓝(ofFieldOp φ * ofFieldOp ψ) := by rfl
/-- For a field specification `𝓕`, and `φ` and `ψ` elements of `𝓕.FieldOp`, if
`φ` and `ψ` are time-ordered then
`timeContract φ ψ = [anPart φ, ofFieldOp ψ]ₛ`. -/
lemma timeContract_of_timeOrderRel (φ ψ : 𝓕.FieldOp) (h : timeOrderRel φ ψ) :
timeContract φ ψ = [anPart φ, ofFieldOp ψ]ₛ := by
conv_rhs =>
@ -55,6 +59,10 @@ lemma timeContract_of_not_timeOrderRel (φ ψ : 𝓕.FieldOp) (h : ¬ timeOrderR
simp only [instCommGroup.eq_1, map_smul, map_add, smul_add]
rw [smul_smul, smul_smul, mul_comm]
/-- For a field specification `𝓕`, and `φ` and `ψ` elements of `𝓕.FieldOp`, if
`φ` and `ψ` are not time-ordered then
`timeContract φ ψ = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • [anPart ψ, ofFieldOp φ]ₛ`. -/
lemma timeContract_of_not_timeOrderRel_expand (φ ψ : 𝓕.FieldOp) (h : ¬ timeOrderRel φ ψ) :
timeContract φ ψ = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • [anPart ψ, ofFieldOp φ]ₛ := by
rw [timeContract_of_not_timeOrderRel _ _ h]
@ -62,6 +70,8 @@ lemma timeContract_of_not_timeOrderRel_expand (φ ψ : 𝓕.FieldOp) (h : ¬ tim
have h1 := IsTotal.total (r := 𝓕.timeOrderRel) φ ψ
simp_all
/-- For a field specification `𝓕`, and `φ` and `ψ` elements of `𝓕.FieldOp`, then
`timeContract φ ψ` is in the center of `𝓕.FieldOpAlgebra`. -/
lemma timeContract_mem_center (φ ψ : 𝓕.FieldOp) :
timeContract φ ψ ∈ Subalgebra.center 𝓕.FieldOpAlgebra := by
by_cases h : timeOrderRel φ ψ

View file

@ -496,6 +496,8 @@ lemma timeOrder_superCommute_anPart_ofFieldOp_neq_time {φ ψ : 𝓕.FieldOp}
apply timeOrder_superCommute_neq_time
simp_all [crAnTimeOrderRel]
/-- For a field specification `𝓕`, and `a`, `b`, `c` in `𝓕.FieldOpAlgebra`, then
`𝓣(a * b * c) = 𝓣(a * 𝓣(b) * c)`. -/
lemma timeOrder_timeOrder_mid (a b c : 𝓕.FieldOpAlgebra) :
𝓣(a * b * c) = 𝓣(a * 𝓣(b) * c) := by
obtain ⟨a, rfl⟩ := ι_surjective a

View file

@ -24,12 +24,17 @@ open FieldOpAlgebra
open FieldStatistic
noncomputable section
/-- For a Wick contraction `φsΛ`, we define `wickTerm φsΛ` to be the element of `𝓕.FieldOpAlgebra`
given by `φsΛ.sign • φsΛ.timeContract * 𝓝([φsΛ]ᵘᶜ)`. -/
/-- For a list `φs` of `𝓕.FieldOp`, and a Wick contraction `φsΛ` of `φs`, the element
of `𝓕.FieldOpAlgebra`, `φsΛ.wickTerm` is defined as
`φsΛ.sign • φsΛ.timeContract * 𝓝([φsΛ]ᵘᶜ)`.
This is term which appears in the Wick's theorem. -/
def wickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : 𝓕.FieldOpAlgebra :=
φsΛ.sign • φsΛ.timeContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
/-- The Wick term for the empty contraction of the empty list is `1`. -/
/-- For the empty list `[]` of `𝓕.FieldOp`, the `wickTerm` of the empty Wick contraction
`empty` of `[]` (its only Wick contraction) is `1`. -/
@[simp]
lemma wickTerm_empty_nil :
wickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
@ -37,7 +42,9 @@ lemma wickTerm_empty_nil :
simp [sign_empty]
/--
Let `φsΛ` be a Wick Contraction for `φs = φ₀φ₁…φₙ`. Then the following holds
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
`𝓕.FieldOp`, and `i ≤ φs.length` the following relation holds
`(φsΛ ↩Λ φ i none).wickTerm = 𝓢(φ, φ₀…φᵢ₋₁) φsΛ.sign • φsΛ.timeContract * 𝓝(φ :: [φsΛ]ᵘᶜ)`
The proof of this result relies on
@ -86,16 +93,18 @@ lemma wickTerm_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
simp only [ZeroMemClass.coe_zero, zero_mul, smul_zero]
exact hg
/-- Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Let `φ` be a field with time
greater then or equal to all the fields in `φs`. Let `i` be a in `Fin φs.length.succ` such that
all files in `φ₀…φᵢ₋₁` have time strictly less then `φ`. Then`(φsΛ ↩Λ φ i (some k)).wickTerm`
/-- For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
`𝓕.FieldOp`, `i ≤ φs.length` and a `k` in `φsΛ.uncontracted`,
such that all `FieldOp` in `φ₀…φᵢ₋₁` have time strictly less then `φ` and
`φ` has a time greater then or equal to all `FieldOp` in `φ₀…φₙ`, then
`(φsΛ ↩Λ φ i (some k)).wickTerm`
is equal the product of
- the sign `𝓢(φ, φ₀…φᵢ₋₁) `
- the sign `φsΛ.sign`
- `φsΛ.timeContract`
- `s • [anPart φ, ofFieldOp φs[k]]ₛ` where `s` is the sign associated with moving `φ` through
uncontracted fields in `φ₀…φₖ₋₁`
- the normal ordering `𝓝([φsΛ]ᵘᶜ.erase (uncontractedFieldOpEquiv φs φsΛ k))`.
- the normal ordering `[φsΛ]ᵘᶜ` with the field corresonding to `k` removed.
The proof of this result relies on
- `timeContract_insert_some_of_not_lt`
@ -167,14 +176,16 @@ lemma wickTerm_insert_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
/--
Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Let `φ` be a field with time
greater then or equal to all the fields in `φs`. Let `i` be a in `Fin φs.length.succ` such that
all files in `φ₀…φᵢ₋₁` have time strictly less then `φ`. Then
greater then or equal to all the `FieldOp` in `φs`. Let `i ≤ φs.length` be such that
all `FieldOp` in `φ₀…φᵢ₋₁` have time strictly less then `φ`. Then
`φ * φsΛ.wickTerm = 𝓢(φ, φ₀…φᵢ₋₁) • ∑ k, (φsΛ ↩Λ φ i k).wickTerm`
where the sum is over all `k` in `Option φsΛ.uncontracted` (so either `none` or `some k`).
The proof of proceeds as follows:
- `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum` is used to expand `φ 𝓝([φsΛ]ᵘᶜ)` as
a sum over `k` in `Option φsΛ.uncontracted` of terms involving `[φ, φs[k]]` etc.
a sum over `k` in `Option φsΛ.uncontracted` of terms involving `[anPart φ, φs[k]]ₛ`..
- Then `wickTerm_insert_none` and `wickTerm_insert_some` are used to equate terms.
-/
lemma mul_wickTerm_eq_sum (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (i : Fin φs.length.succ)

View file

@ -61,18 +61,38 @@ remark wicks_theorem_context := "
The statement of these theorems for bosons is simplier then when fermions are involved, since
one does not have to worry about the minus-signs picked up on exchanging fields."
/-- Wick's theorem states that for a list of fields `φs = φ₀…φₙ`
`𝓣(φs) = ∑ φsΛ, (φsΛ.sign • φsΛ.timeContract) * 𝓝([φsΛ]ᵘᶜ)`
where the sum is over all Wick contractions `φsΛ` of `φs`.
/--
For a list `φs` of `𝓕.FieldOp`, Wick's theorem states that
`𝓣(φs) = ∑ φsΛ, φsΛ.wickTerm`
where the sum is over all Wick contraction `φsΛ`.
The proof is via induction on `φs`.
- The base case `φs = []` is handled by `wickTerm_empty_nil`.
The proof is via induction on `φs`. The base case `φs = []` is handled by `wicks_theorem_nil`.
The inductive step works as follows:
- The lemma `timeOrder_eq_maxTimeField_mul_finset` is used to write
For the LHS:
1. `timeOrder_eq_maxTimeField_mul_finset` is used to write
`𝓣(φ₀…φₙ)` as ` 𝓢(φᵢ,φ₀…φᵢ₋₁) • φᵢ * 𝓣(φ₀…φᵢ₋₁φᵢ₊₁φₙ)` where `φᵢ` is
the maximal time field in `φ₀…φₙ`.
- The induction hypothesis is used to expand `𝓣(φ₀…φᵢ₋₁φᵢ₊₁φₙ)` as a sum over Wick contractions of
`φ₀…φᵢ₋₁φᵢ₊₁φₙ`.
- Further the lemmas `wick_term_cons_eq_sum_wick_term` and `insertLift_sum` are used.
the maximal time field in `φ₀…φₙ`
2. The induction hypothesis is then used on `𝓣(φ₀…φᵢ₋₁φᵢ₊₁φₙ)` to expand it as a sum over
Wick contractions of `φ₀…φᵢ₋₁φᵢ₊₁φₙ`.
3. This gives terms of the form `φᵢ * φsΛ.timeContract` on which
`mul_wickTerm_eq_sum` is used where `φsΛ` is a Wick contraction of `φ₀…φᵢ₋₁φᵢ₊₁φ`,
to rewrite terms as a sum over optional uncontracted elements of `φsΛ`
On the LHS we now have a sum over Wick contractions `φsΛ` of `φ₀…φᵢ₋₁φᵢ₊₁φ` (from 2) and optional
uncontracted elements of `φsΛ` (from 3)
For the RHS:
1. The sum over Wick contractions of `φ₀…φₙ` on the RHS
is split via `insertLift_sum` into a sum over Wick contractions `φsΛ` of `φ₀…φᵢ₋₁φᵢ₊₁φ` and
sum over optional uncontracted elements of `φsΛ`.
Both side now are sums over the same thing and their terms equate by the nature of the
lemmas used.
-/
theorem wicks_theorem : (φs : List 𝓕.FieldOp) → 𝓣(ofFieldOpList φs) =
∑ (φsΛ : WickContraction φs.length), φsΛ.wickTerm
@ -80,7 +100,7 @@ theorem wicks_theorem : (φs : List 𝓕.FieldOp) → 𝓣(ofFieldOpList φs) =
rw [timeOrder_ofFieldOpList_nil]
simp only [map_one, List.length_nil, Algebra.smul_mul_assoc]
rw [sum_WickContraction_nil]
simp
simp only [wickTerm_empty_nil]
| φ :: φs => by
have ih := wicks_theorem (eraseMaxTimeField φ φs)
conv_lhs => rw [timeOrder_eq_maxTimeField_mul_finset, ih, Finset.mul_sum]

View file

@ -20,6 +20,23 @@ namespace FieldOpAlgebra
open WickContraction
open EqTimeOnly
/--
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)`.
- `EqTimeOnly.timeOrder_staticContract_of_not_mem` and `timeOrder_timeOrder_mid` to set to
zero those terms in which the contracted elements do not have equal time.
- `staticContract_eq_timeContract_of_eqTimeOnly` to rewrite the static contract as a time contract
for those terms which have equal time.
- `timeOrder_timeContract_mul_of_eqTimeOnly_left` to move the time contracts out of the time
ordering.
-/
lemma timeOrder_ofFieldOpList_eqTimeOnly (φs : List 𝓕.FieldOp) :
𝓣(ofFieldOpList φs) = ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs)}),
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
@ -44,8 +61,9 @@ lemma timeOrder_ofFieldOpList_eqTimeOnly (φs : List 𝓕.FieldOp) :
exact x.2
exact x.2
lemma timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.FieldOp) :
timeOrder (ofFieldOpList φs) = 𝓣(𝓝(ofFieldOpList φs)) +
𝓣(ofFieldOpList φs) = 𝓣(𝓝(ofFieldOpList φs)) +
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
let e1 : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly} ≃
@ -71,6 +89,17 @@ lemma timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.FieldOp) :
rw [← e2.symm.sum_comp]
rfl
/--
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`
-/
lemma normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.FieldOp) :
𝓣(𝓝(ofFieldOpList φs)) = 𝓣(ofFieldOpList φs) -
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
@ -78,14 +107,33 @@ lemma normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty (φs : List 𝓕.F
rw [timeOrder_ofFieldOpList_eq_eqTimeOnly_empty]
simp
lemma normalOrder_timeOrder_ofFieldOpList_eq_haveEqTime_sum_not_haveEqTime (φs : List 𝓕.FieldOp) :
𝓣(𝓝(ofFieldOpList φs)) = (∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
/--
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Λ, sign φs ↑φsΛ • (φsΛ.1).timeContract ∑ φssucΛ, φssucΛ.wickTerm`, where
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 this result relies on `wicks_theorem` 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 equal time contractions.
The sum over Wick contractions which do have equal time contractions is turned into two sums
one over the Wick contractions which only have equal time contractions and the other over the
uncontracted elements of the Wick contraction which do not have equal time contractions using
`join`.
The properties of `join_sign_timeContract` is then used to equate terms.
-/
lemma timeOrder_haveEqTime_split (φs : List 𝓕.FieldOp) :
𝓣(ofFieldOpList φs) = (∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
+ (∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}),
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))
- ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) := by
rw [normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty]
+ ∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
sign φs ↑φsΛ • (φsΛ.1).timeContract *
(∑ φssucΛ : { φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ φssucΛ.HaveEqTime },
sign [φsΛ.1]ᵘᶜ φssucΛ • (φssucΛ.1).timeContract * normalOrder (ofFieldOpList [φssucΛ.1]ᵘᶜ)) := by
rw [wicks_theorem]
simp only [wickTerm]
let e1 : WickContraction φs.length ≃ {φsΛ // HaveEqTime φsΛ} ⊕ {φsΛ // ¬ HaveEqTime φsΛ} := by
@ -94,17 +142,9 @@ lemma normalOrder_timeOrder_ofFieldOpList_eq_haveEqTime_sum_not_haveEqTime (φs
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]
rw [add_comm]
lemma haveEqTime_wick_sum_eq_split (φs : List 𝓕.FieldOp) :
(∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}),
φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)) =
∑ (φsΛ : {φsΛ // φsΛ.EqTimeOnly (φs := φs) ∧ φsΛ ≠ empty}),
(sign φs ↑φsΛ • (φsΛ.1).timeContract *
∑ φssucΛ : { φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬φssucΛ.HaveEqTime },
sign [φsΛ.1]ᵘᶜ φssucΛ •
(φssucΛ.1).timeContract * normalOrder (ofFieldOpList [φssucΛ.1]ᵘᶜ)) := by
congr 1
let f : WickContraction φs.length → 𝓕.FieldOpAlgebra := fun φsΛ =>
φsΛ.sign • φsΛ.timeContract.1 * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
φsΛ.sign • (φsΛ.timeContract.1 * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ))
change ∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}), f φsΛ.1 = _
rw [sum_haveEqTime]
congr
@ -112,12 +152,12 @@ lemma haveEqTime_wick_sum_eq_split (φs : List 𝓕.FieldOp) :
simp only [f]
conv_lhs =>
enter [2, φsucΛ]
enter [1]
rw [← Algebra.smul_mul_assoc]
rw [join_sign_timeContract φsΛ.1 φsucΛ.1]
conv_lhs =>
enter [2, φsucΛ]
rw [mul_assoc]
rw [← Finset.mul_sum]
rw [← Finset.mul_sum, ← Algebra.smul_mul_assoc]
congr
funext φsΛ'
simp only [ne_eq, Algebra.smul_mul_assoc]
@ -132,10 +172,10 @@ lemma normalOrder_timeOrder_ofFieldOpList_eq_not_haveEqTime_sub_inductive (φs :
(∑ φssucΛ : { φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ φssucΛ.HaveEqTime },
sign [φsΛ.1]ᵘᶜ φssucΛ • (φssucΛ.1).timeContract * normalOrder (ofFieldOpList [φssucΛ.1]ᵘᶜ) -
𝓣(𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ))) := by
rw [normalOrder_timeOrder_ofFieldOpList_eq_haveEqTime_sum_not_haveEqTime]
rw [normalOrder_timeOrder_ofFieldOpList_eq_eqTimeOnly_empty,
timeOrder_haveEqTime_split]
rw [add_sub_assoc]
congr 1
rw [haveEqTime_wick_sum_eq_split]
simp only [ne_eq, Algebra.smul_mul_assoc]
rw [← Finset.sum_sub_distrib]
congr 1
@ -177,12 +217,19 @@ lemma wicks_theorem_normal_order_empty : 𝓣(𝓝(ofFieldOpList [])) =
rw [timeOrderF_ofCrAnListF]
simp
/--
Wicks theorem for normal ordering followed by time-ordering, states that
`𝓣(𝓝(φ₀…φₙ))` is equal to
`∑ φsΛ, φsΛ.1.sign • φsΛ.1.timeContract.1 * 𝓝(ofFieldOpList [φsΛ.1]ᵘᶜ)`
over those Wick contraction `φsΛ` which do not have any equal time contractions.
This is compared to the ordinary Wicks theorem which sums over all Wick contractions.
/--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.
The proof of proceeds by induction on `φs`, with the base case `[]` holding by following
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.
-/
theorem wicks_theorem_normal_order : (φs : List 𝓕.FieldOp) →
𝓣(𝓝(ofFieldOpList φs)) =