PhysLean/HepLean/PerturbationTheory/Algebras/FieldOpFreeAlgebra/NormTimeOrder.lean

48 lines
1.2 KiB
Text
Raw Normal View History

/-
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.FieldSpecification.TimeOrder
2025-02-03 11:54:08 +00:00
import HepLean.PerturbationTheory.Algebras.FieldOpFreeAlgebra.Basic
import HepLean.PerturbationTheory.Koszul.KoszulSign
/-!
2025-02-03 11:05:43 +00:00
# Norm-time Ordering in the FieldOpFreeAlgebra
-/
namespace FieldSpecification
variable {𝓕 : FieldSpecification}
open FieldStatistic
2025-02-03 11:05:43 +00:00
namespace FieldOpFreeAlgebra
noncomputable section
open HepLean.List
/-!
## Norm-time order
-/
2025-02-03 11:05:43 +00:00
/-- The normal-time ordering on `FieldOpFreeAlgebra`. -/
def normTimeOrder : FieldOpFreeAlgebra 𝓕 →ₗ[] FieldOpFreeAlgebra 𝓕 :=
Basis.constr ofCrAnListFBasis fun φs =>
normTimeOrderSign φs • ofCrAnListF (normTimeOrderList φs)
@[inherit_doc normTimeOrder]
2025-02-03 11:05:43 +00:00
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "𝓣𝓝ᶠ(" a ")" => normTimeOrder a
2025-02-03 11:28:14 +00:00
lemma normTimeOrder_ofCrAnListF (φs : List 𝓕.CrAnFieldOp) :
𝓣𝓝ᶠ(ofCrAnListF φs) = normTimeOrderSign φs • ofCrAnListF (normTimeOrderList φs) := by
rw [← ofListBasis_eq_ofList]
simp only [normTimeOrder, Basis.constr_basis]
end
2025-02-03 11:05:43 +00:00
end FieldOpFreeAlgebra
end FieldSpecification