refactor: More docs for Wick's theorems
This commit is contained in:
parent
4096010e70
commit
b30a49d7db
8 changed files with 38 additions and 38 deletions
|
@ -31,8 +31,8 @@ noncomputable section
|
|||
def staticWickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : 𝓕.FieldOpAlgebra :=
|
||||
φsΛ.sign • φsΛ.staticContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
|
||||
|
||||
/-- For the empty list `[]` of `𝓕.FieldOp`, the `staticWickTerm` of the empty Wick contraction
|
||||
`empty` of `[]` (its only Wick contraction) is `1`. -/
|
||||
/-- For the empty list `[]` of `𝓕.FieldOp`, the `staticWickTerm` of the Wick contraction
|
||||
corresponding to the empty set `∅` (the only Wick contraction of `[]`) is `1`. -/
|
||||
@[simp]
|
||||
lemma staticWickTerm_empty_nil :
|
||||
staticWickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
|
||||
|
@ -116,9 +116,9 @@ lemma staticWickTerm_insert_zero_some (φ : 𝓕.FieldOp) (φs : List 𝓕.Field
|
|||
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, the following relation
|
||||
holds
|
||||
|
||||
`φ * φsΛ.staticWickTerm = ∑ k, (φsΛ ↩Λ φ i k).wickTerm`
|
||||
`φ * φsΛ.staticWickTerm = ∑ k, (φsΛ ↩Λ φ 0 k).staticWickTerm`
|
||||
|
||||
where the sum is over all `k` in `Option φsΛ.uncontracted` (so either `none` or `some k`).
|
||||
where the sum is over all `k` in `Option φsΛ.uncontracted`, so `k` is either `none` or `some k`.
|
||||
|
||||
The proof of proceeds as follows:
|
||||
- `ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum` is used to expand `φ 𝓝([φsΛ]ᵘᶜ)` as
|
||||
|
|
|
@ -22,7 +22,7 @@ namespace FieldOpAlgebra
|
|||
/--
|
||||
For a list `φs` of `𝓕.FieldOp`, the static version of Wick's theorem states that
|
||||
|
||||
`φs =∑ φsΛ, φsΛ.staticWickTerm`
|
||||
`φs = ∑ φsΛ, φsΛ.staticWickTerm`
|
||||
|
||||
where the sum is over all Wick contraction `φsΛ`.
|
||||
|
||||
|
|
|
@ -24,7 +24,7 @@ namespace FieldOpAlgebra
|
|||
open FieldStatistic
|
||||
|
||||
/-- For a field specification `𝓕`, and `φ` and `ψ` elements of `𝓕.FieldOp`, the element of
|
||||
`𝓕.FieldOpAlgebra`, `timeContract φ ψ` is defined to be `𝓣(φψ)- 𝓝(φψ)`. -/
|
||||
`𝓕.FieldOpAlgebra`, `timeContract φ ψ` is defined to be `𝓣(φψ) - 𝓝(φψ)`. -/
|
||||
def timeContract (φ ψ : 𝓕.FieldOp) : 𝓕.FieldOpAlgebra :=
|
||||
𝓣(ofFieldOp φ * ofFieldOp ψ) - 𝓝(ofFieldOp φ * ofFieldOp ψ)
|
||||
|
||||
|
|
|
@ -33,8 +33,8 @@ noncomputable section
|
|||
def wickTerm {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : 𝓕.FieldOpAlgebra :=
|
||||
φsΛ.sign • φsΛ.timeContract * 𝓝(ofFieldOpList [φsΛ]ᵘᶜ)
|
||||
|
||||
/-- For the empty list `[]` of `𝓕.FieldOp`, the `wickTerm` of the empty Wick contraction
|
||||
`empty` of `[]` (its only Wick contraction) is `1`. -/
|
||||
/-- For the empty list `[]` of `𝓕.FieldOp`, the `wickTerm` of the Wick contraction
|
||||
corresponding to the empty set `∅` (the only Wick contraction of `[]`) is `1`. -/
|
||||
@[simp]
|
||||
lemma wickTerm_empty_nil :
|
||||
wickTerm (empty (n := ([] : List 𝓕.FieldOp).length)) = 1 := by
|
||||
|
@ -95,9 +95,9 @@ lemma wickTerm_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
|||
|
||||
/-- 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
|
||||
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`
|
||||
`(φsΛ ↩Λ φ i (some k)).staticWickTerm`
|
||||
is equal the product of
|
||||
- the sign `𝓢(φ, φ₀…φᵢ₋₁) `
|
||||
- the sign `φsΛ.sign`
|
||||
|
@ -175,17 +175,18 @@ lemma wickTerm_insert_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
|||
exact hg'
|
||||
|
||||
/--
|
||||
Let `φsΛ` be a Wick contraction for `φs = φ₀φ₁…φₙ`. Let `φ` be a field with time
|
||||
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
|
||||
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||
`𝓕.FieldOp`, and `i ≤ φs.length`
|
||||
such that all `𝓕.FieldOp` in `φ₀…φᵢ₋₁` have time strictly less then `φ` and
|
||||
`φ` has a time greater then or equal to all `FieldOp` in `φ₀…φₙ`, 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`).
|
||||
where the sum is over all `k` in `Option φsΛ.uncontracted`, so `k` is 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 `[anPart φ, φs[k]]ₛ`..
|
||||
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)
|
||||
|
|
|
@ -75,11 +75,11 @@ The inductive step works as follows:
|
|||
|
||||
For the LHS:
|
||||
1. `timeOrder_eq_maxTimeField_mul_finset` is used to write
|
||||
`𝓣(φ₀…φₙ)` as ` 𝓢(φᵢ,φ₀…φᵢ₋₁) • φᵢ * 𝓣(φ₀…φᵢ₋₁φᵢ₊₁φₙ)` where `φᵢ` is
|
||||
`𝓣(φ₀…φₙ)` as `𝓢(φᵢ,φ₀…φᵢ₋₁) • φᵢ * 𝓣(φ₀…φᵢ₋₁φᵢ₊₁φₙ)` where `φᵢ` is
|
||||
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
|
||||
3. This gives terms of the form `φᵢ * φsΛ.wickTerm` 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Λ`
|
||||
|
||||
|
|
|
@ -31,9 +31,10 @@ 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.
|
||||
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
|
||||
in the reminaing `𝓣(φsΛ.staticWickTerm)` as a time contract.
|
||||
- `timeOrder_timeContract_mul_of_eqTimeOnly_left` to move the time contracts out of the time
|
||||
ordering.
|
||||
-/
|
||||
|
@ -111,20 +112,20 @@ 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
|
||||
- `∑ φsΛ, φsΛ.sign • φsΛ.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.
|
||||
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.
|
||||
-/
|
||||
lemma timeOrder_haveEqTime_split (φs : List 𝓕.FieldOp) :
|
||||
𝓣(ofFieldOpList φs) = (∑ (φsΛ : {φsΛ : WickContraction φs.length // ¬ HaveEqTime φsΛ}),
|
||||
|
|
|
@ -97,7 +97,7 @@ lemma empty_mem {φs : List 𝓕.FieldOp} : empty (n := φs.length).EqTimeOnly :
|
|||
simp [empty]
|
||||
|
||||
/-- Let `φs` be a list of `𝓕.FieldOp` and `φsΛ` a `WickContraction` of `φs` with
|
||||
in which every contraction involves two `FieldOp`s that have the same time. Then
|
||||
in which every contraction involves two `𝓕FieldOp`s that have the same time, then
|
||||
`φsΛ.staticContract = φsΛ.timeContract`. -/
|
||||
lemma staticContract_eq_timeContract_of_eqTimeOnly (h : φsΛ.EqTimeOnly) :
|
||||
φsΛ.staticContract = φsΛ.timeContract := by
|
||||
|
@ -194,7 +194,7 @@ lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid {φs : List 𝓕.FieldOp}
|
|||
exact timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction φsΛ hl a b φsΛ.1.card rfl
|
||||
|
||||
/-- Let `φs` be a list of `𝓕.FieldOp`, `φsΛ` a `WickContraction` of `φs` with
|
||||
in which every contraction involves two `FieldOp`s that have the same time and
|
||||
in which every contraction involves two `𝓕.FieldOp`s that have the same time and
|
||||
`b` a general element in `𝓕.FieldOpAlgebra`. Then
|
||||
`𝓣(φsΛ.timeContract.1 * b) = φsΛ.timeContract.1 * 𝓣(b)`.
|
||||
|
||||
|
@ -248,7 +248,7 @@ lemma timeOrder_timeContract_of_not_eqTimeOnly {φs : List 𝓕.FieldOp}
|
|||
simp_all
|
||||
|
||||
/-- Let `φs` be a list of `𝓕.FieldOp` and `φsΛ` a `WickContraction` with
|
||||
at least one contraction between `FieldOp` that do not have the same time. Then
|
||||
at least one contraction between `𝓕.FieldOp` that do not have the same time. Then
|
||||
`𝓣(φsΛ.staticContract.1) = 0`. -/
|
||||
lemma timeOrder_staticContract_of_not_mem {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(hl : ¬ φsΛ.EqTimeOnly) : 𝓣(φsΛ.staticContract.1) = 0 := by
|
||||
|
|
|
@ -21,7 +21,7 @@ open FieldOpAlgebra
|
|||
|
||||
/-- For a list `φs` of `𝓕.FieldOp` and a Wick contraction `φsΛ` the
|
||||
element of the center of `𝓕.FieldOpAlgebra`, `φsΛ.timeContract` is defined as the product
|
||||
of `timeContract φs[j] φs[k]` over contracted pairs `{j, k}` (both indices of `φs`) in `φsΛ`
|
||||
of `timeContract φs[j] φs[k]` over contracted pairs `{j, k}` in `φsΛ`
|
||||
with `j < k`. -/
|
||||
noncomputable def timeContract {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
|
@ -86,7 +86,7 @@ open FieldStatistic
|
|||
- `[anPart φ, φs[k]]ₛ`
|
||||
- `φsΛ.timeContract`
|
||||
- two copies of the exchange sign of `φ` with the uncontracted fields in `φ₀…φₖ₋₁`.
|
||||
These two exchange signs cancle each other out but are included for convenience.
|
||||
These two exchange signs cancel each other out but are included for convenience.
|
||||
|
||||
The proof of this result ultimately a consequence of definitions and
|
||||
`timeContract_of_timeOrderRel`. -/
|
||||
|
@ -125,15 +125,13 @@ lemma timeContract_insert_some_of_lt
|
|||
|
||||
/-- 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` such that `k < i`, with the
|
||||
condition that `φs[k]` does not have has greater or equal time to `φ`, then
|
||||
condition that `φs[k]` does not have time greater or equal to `φ`, then
|
||||
`(φsΛ ↩Λ φ i (some k)).timeContract` is equal to the product of
|
||||
- `[anPart φ, φs[k]]ₛ`
|
||||
- `φsΛ.timeContract`
|
||||
- the exchange sign of `φ` with the uncontracted fields in `φ₀…φₖ₋₁`.
|
||||
- the exchange sign of `φ` with the uncontracted fields in `φ₀…φₖ`.
|
||||
|
||||
Most of the contributes to the exchange signs cancle.
|
||||
|
||||
The proof of this result ultimately a consequence of definitions and
|
||||
`timeContract_of_not_timeOrderRel_expand`. -/
|
||||
lemma timeContract_insert_some_of_not_lt
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue