docs: Normal ordering

This commit is contained in:
jstoobysmith 2025-02-06 14:10:45 +00:00
parent fede0b7904
commit ce9668a3cd
7 changed files with 74 additions and 40 deletions

View file

@ -217,12 +217,19 @@ lemma anPart_mul_normalOrder_ofFieldOpList_eq_superCommute (φ : 𝓕.FieldOp)
-/
/--
For a `CrAnFieldOp` `φ` and a list of `CrAnFieldOp`s `φs`, the following is true
For a field specification `𝓕`, an element `φ` of `𝓕.CrAnFieldOp`, a list `φs` of `𝓕.CrAnFieldOp`,
the following relation holds
`[φ, 𝓝(φ₀…φₙ)]ₛ = ∑ i, 𝓢(φ, φ₀…φᵢ₋₁) • [φ, φᵢ]ₛ * 𝓝(φ₀…φᵢ₋₁φᵢ₊₁…φₙ)`.
The proof of this result ultimetly depends on
- `superCommuteF_ofCrAnListF_ofFieldOpListF_eq_sum`
- `normalOrderSign_eraseIdx`
The proof of this result ultimetly goes as follows
- The definition of `normalOrder` is used to rewrite `𝓝(φ₀…φₙ)` as a scalar multiple of
a `ofCrAnList φsn` where `φsn` is the normal ordering of `φ₀…φₙ`.
- `superCommuteF_ofCrAnListF_ofCrAnListF_eq_sum` is used to rewrite the super commutator of `φ`
(considered as a list with one lement) with
`ofCrAnList φsn` as a sum of supercommutors, one for each element of `φsn`.
- The fact that super-commutors are in the center of `𝓕.FieldOpAlgebra` is used to rearange terms.
- Properties of ordered lists, and `normalOrderSign_eraseIdx` is then used to complete the proof.
-/
lemma ofCrAnOp_superCommute_normalOrder_ofCrAnList_sum (φ : 𝓕.CrAnFieldOp)
(φs : List 𝓕.CrAnFieldOp) : [ofCrAnOp φ, 𝓝(ofCrAnList φs)]ₛ = ∑ n : Fin φs.length,
@ -335,11 +342,18 @@ noncomputable def contractStateAtIndex (φ : 𝓕.FieldOp) (φs : List 𝓕.Fiel
| some n => 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (φs.take n)) • [anPart φ, ofFieldOp φs[n]]ₛ
/--
For a field specification `𝓕`, the following relation holds in the algebra `𝓕.FieldOpAlgebra`,
For a field specification `𝓕`, a `φ` in `𝓕.FieldOp` and a list `φs` of `𝓕.FieldOp`
the following relation holds in the algebra `𝓕.FieldOpAlgebra`,
`φ * 𝓝(φ₀φ₁…φₙ) = 𝓝(φφ₀φ₁…φₙ) + ∑ i, (𝓢(φ,φ₀φ₁…φᵢ₋₁) • [anPartF φ, φᵢ]ₛ) * 𝓝(φ₀φ₁…φᵢ₋₁φᵢ₊₁…φₙ)`.
The proof of this ultimently depends on :
- `ofCrAnOp_superCommute_normalOrder_ofCrAnList_sum`
The proof of ultimetly goes as follows:
- `ofFieldOp_eq_crPart_add_anPart` is used to split `φ` into its creation and annihilation parts.
- The fact that `crPart φ * 𝓝(φ₀φ₁…φₙ) = 𝓝(crPart φ * φ₀φ₁…φₙ)` is used.
- The fact that `anPart φ * 𝓝(φ₀φ₁…φₙ)` is
`𝓢(φ, φ₀φ₁…φₙ) 𝓝(φ₀φ₁…φₙ) * anPart φ + [anPart φ, 𝓝(φ₀φ₁…φₙ)]` is used
- The fact that `𝓢(φ, φ₀φ₁…φₙ) 𝓝(φ₀φ₁…φₙ) * anPart φ = 𝓝(anPart φ * φ₀φ₁…φₙ)`
- The result `ofCrAnOp_superCommute_normalOrder_ofCrAnList_sum` is used
to expand `[anPart φ, 𝓝(φ₀φ₁…φₙ)]` as a sum.
-/
lemma ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) :
ofFieldOp φ * 𝓝(ofFieldOpList φs) =