docs: Notes for normal-ordered Wicks
This commit is contained in:
parent
c9607c459f
commit
d472604aec
15 changed files with 324 additions and 131 deletions
|
@ -284,6 +284,14 @@ lemma insert_fin_eq_self (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
|||
use z
|
||||
rfl
|
||||
|
||||
/-- For a list `φs` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||
`𝓕.FieldOp`, a `i ≤ φs.length` a sum over
|
||||
Wick contractions of `φs` with `φ` inserted at `i` is equal to the sum over Wick contractions
|
||||
`φsΛ` of just `φs` and the sum over optional uncontracted elements of the `φsΛ`.
|
||||
|
||||
I.e. `∑ (φsΛ : WickContraction (φs.insertIdx i φ).length), f φsΛ` is equal to
|
||||
`∑ (φsΛ : WickContraction φs.length), ∑ (k : Option φsΛ.uncontracted), f (φsΛ ↩Λ φ i k) `.
|
||||
where `(φs.insertIdx i φ)` is `φs` with `φ` inserted at position `i`. -/
|
||||
lemma insertLift_sum (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) :
|
||||
∑ c, f c =
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue