doc: Edits to Wick theorem docs

This commit is contained in:
jstoobysmith 2025-02-10 10:21:57 +00:00
parent 26d2c24c83
commit 4096010e70
32 changed files with 255 additions and 134 deletions

View file

@ -33,15 +33,18 @@ def fieldOpIdealSet : Set (FieldOpFreeAlgebra 𝓕) :=
x = [ofCrAnOpF φ, ofCrAnOpF φ']ₛca)}
/-- For a field specification `𝓕`, the algebra `𝓕.FieldOpAlgebra` is defined as the quotient
of the free algebra `𝓕.FieldOpFreeAlgebra` by the ideal generated by elements of the form
- `[ofCrAnOpF φ1, [ofCrAnOpF φ2, ofCrAnOpF φ3]ₛca]ₛca`, which (with the other conditions)
gives us that super-commutors are in the center.
- `[ofCrAnOpF φc, ofCrAnOpF φc']ₛca` for `φc` and `φc'` creation operators. I.e two creation
operators always super-commute.
- `[ofCrAnOpF φa, ofCrAnOpF φa']ₛca` for `φa` and `φa'` annihilation operators. I.e two
annihilation operators always super-commute.
- `[ofCrAnOpF φ, ofCrAnOpF φ']ₛca` for `φ` and `φ'` field operators with different statistics.
I.e. Fermions super-commute with bosons. -/
of the free algebra `𝓕.FieldOpFreeAlgebra` by the ideal generated by
- `[ofCrAnOpF φc, ofCrAnOpF φc']ₛca` for `φc` and `φc'` field creation operators.
This corresponds to the condition that two creation operators always super-commute.
- `[ofCrAnOpF φa, ofCrAnOpF φa']ₛca` for `φa` and `φa'` field annihilation operators.
This corresponds to the condition that two annihilation operators always super-commute.
- `[ofCrAnOpF φ, ofCrAnOpF φ']ₛca` for `φ` and `φ'` operators with different statistics.
This corresponds to the condition that two operators with different statistics always
super-commute. In otherwords, fermions and bosons always super-commute.
- `[ofCrAnOpF φ1, [ofCrAnOpF φ2, ofCrAnOpF φ3]ₛca]ₛca`. This corresponds to the condition,
when combined with the conditions above, that the super-commutor is in the center of the
of the algebra.
-/
abbrev FieldOpAlgebra : Type := (TwoSidedIdeal.span 𝓕.fieldOpIdealSet).ringCon.Quotient
namespace FieldOpAlgebra
@ -457,13 +460,15 @@ lemma ι_eq_zero_iff_ι_bosonicProjF_fermonicProj_zero (x : FieldOpFreeAlgebra
-/
/-- For a field specification `𝓕` and an element `φ` of `𝓕.FieldOp`, the element of
/-- For a field specification `𝓕` and an element `φ` of `𝓕.FieldOp`,
`ofFieldOp φ` is defined as the element of
`𝓕.FieldOpAlgebra` given by `ι (ofFieldOpF φ)`. -/
def ofFieldOp (φ : 𝓕.FieldOp) : 𝓕.FieldOpAlgebra := ι (ofFieldOpF φ)
lemma ofFieldOp_eq_ι_ofFieldOpF (φ : 𝓕.FieldOp) : ofFieldOp φ = ι (ofFieldOpF φ) := rfl
/-- For a field specification `𝓕` and a list `φs` of `𝓕.FieldOp`, the element of
/-- For a field specification `𝓕` and a list `φs` of `𝓕.FieldOp`,
`ofFieldOpList φs` is defined as the element of
`𝓕.FieldOpAlgebra` given by `ι (ofFieldOpListF φ)`. -/
def ofFieldOpList (φs : List 𝓕.FieldOp) : 𝓕.FieldOpAlgebra := ι (ofFieldOpListF φs)
@ -487,7 +492,8 @@ lemma ofFieldOpList_singleton (φ : 𝓕.FieldOp) :
ofFieldOpList [φ] = ofFieldOp φ := by
simp only [ofFieldOpList, ofFieldOp, ofFieldOpListF_singleton]
/-- For a field specification `𝓕` and an element `φ` of `𝓕.CrAnFieldOp`, the element of
/-- For a field specification `𝓕` and an element `φ` of `𝓕.CrAnFieldOp`,
`ofCrAnOp φ` is defined as the element of
`𝓕.FieldOpAlgebra` given by `ι (ofCrAnOpF φ)`. -/
def ofCrAnOp (φ : 𝓕.CrAnFieldOp) : 𝓕.FieldOpAlgebra := ι (ofCrAnOpF φ)
@ -500,7 +506,8 @@ lemma ofFieldOp_eq_sum (φ : 𝓕.FieldOp) :
simp only [map_sum]
rfl
/-- For a field specification `𝓕` and a list `φs` of `𝓕.CrAnFieldOp`, the element of
/-- For a field specification `𝓕` and a list `φs` of `𝓕.CrAnFieldOp`,
`ofCrAnList φs` is defined as the element of
`𝓕.FieldOpAlgebra` given by `ι (ofCrAnListF φ)`. -/
def ofCrAnList (φs : List 𝓕.CrAnFieldOp) : 𝓕.FieldOpAlgebra := ι (ofCrAnListF φs)
@ -528,8 +535,10 @@ remark notation_drop := "In doc-strings we will often drop explicit applications
/-- For a field specification `𝓕`, and an element `φ` of `𝓕.FieldOp`, the
annihilation part of `𝓕.FieldOp` as an element of `𝓕.FieldOpAlgebra`.
If `φ` is an incoming asymptotic state this is zero by definition, otherwise
it is of the form `ofCrAnOp _`. -/
Thus for `φ`
- an incoming asymptotic state this is `0`.
- a position based state this is `ofCrAnOp ⟨φ, .create⟩`.
- an outgoing asymptotic state this is `ofCrAnOp ⟨φ, ()⟩`. -/
def anPart (φ : 𝓕.FieldOp) : 𝓕.FieldOpAlgebra := ι (anPartF φ)
lemma anPart_eq_ι_anPartF (φ : 𝓕.FieldOp) : anPart φ = ι (anPartF φ) := rfl
@ -552,8 +561,10 @@ lemma anPart_posAsymp (φ : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → )) :
/-- For a field specification `𝓕`, and an element `φ` of `𝓕.FieldOp`, the
creation part of `𝓕.FieldOp` as an element of `𝓕.FieldOpAlgebra`.
If `φ` is an outgoing asymptotic state this is zero by definition, otherwise
it is of the form `ofCrAnOp _`. -/
Thus for `φ`
- an incoming asymptotic state this is `ofCrAnOp ⟨φ, ()⟩`.
- a position based state this is `ofCrAnOp ⟨φ, .create⟩`.
- an outgoing asymptotic state this is `0`. -/
def crPart (φ : 𝓕.FieldOp) : 𝓕.FieldOpAlgebra := ι (crPartF φ)
lemma crPart_eq_ι_crPartF (φ : 𝓕.FieldOp) : crPart φ = ι (crPartF φ) := rfl