refactor: Some properties of field specifications (#285)

* refactor: Fix field struct defn.

* rename: FieldStruct to FieldSpecification

* feat: Add examples of field specifications

* docs: Slight improvement of module docs
This commit is contained in:
Joseph Tooby-Smith 2025-01-21 06:11:47 +00:00 committed by GitHub
parent bb1db930b5
commit b5c987180a
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
29 changed files with 153 additions and 126 deletions

View file

@ -3,23 +3,18 @@ Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.PerturbationTheory.FieldStruct.NormalOrder
import HepLean.PerturbationTheory.FieldSpecification.NormalOrder
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute
import HepLean.PerturbationTheory.Koszul.KoszulSign
/-!
# Normal Ordering
# Normal Ordering in the CrAnAlgebra
The normal ordering puts all creation operators to the left and all annihilation operators to the
right. It acts on `CrAnStates` and defines a linear map from the `CrAnAlgebra` to itself.
The normal ordering satisfies a number of nice properties with relation to the operator
algebra 𝓞.A.
-/
namespace FieldStruct
variable {𝓕 : FieldStruct}
namespace FieldSpecification
variable {𝓕 : FieldSpecification}
open FieldStatistic
/-!
@ -552,4 +547,4 @@ end
end CrAnAlgebra
end FieldStruct
end FieldSpecification