refactor: More notes on Wick's thoerem
This commit is contained in:
parent
20df8ece6c
commit
034f6c8c91
6 changed files with 115 additions and 77 deletions
|
@ -32,8 +32,19 @@ def fieldOpIdealSet : Set (FieldOpFreeAlgebra 𝓕) :=
|
|||
∨ (∃ (φ φ' : 𝓕.CrAnFieldOp) (_ : ¬ (𝓕 |>ₛ φ) = (𝓕 |>ₛ φ')),
|
||||
x = [ofCrAnOpF φ, ofCrAnOpF φ']ₛca)}
|
||||
|
||||
/-- The algebra spanned by cr and an parts of fields, with appropriate super-commutors
|
||||
set to zero. -/
|
||||
/-- 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.
|
||||
The algebra `𝓕.FieldOpAlgebra` is the most general (in the correct sense) algebra
|
||||
satisfying these properties.
|
||||
-/
|
||||
abbrev FieldOpAlgebra : Type := (TwoSidedIdeal.span 𝓕.fieldOpIdealSet).ringCon.Quotient
|
||||
|
||||
namespace FieldOpAlgebra
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue