PhysLean/HepLean/PerturbationTheory/FieldOpFreeAlgebra/NormTimeOrder.lean
2025-02-03 12:12:36 +00:00

47 lines
1.2 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
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
import HepLean.PerturbationTheory.FieldOpFreeAlgebra.Basic
import HepLean.PerturbationTheory.Koszul.KoszulSign
/-!
# Norm-time Ordering in the FieldOpFreeAlgebra
-/
namespace FieldSpecification
variable {𝓕 : FieldSpecification}
open FieldStatistic
namespace FieldOpFreeAlgebra
noncomputable section
open HepLean.List
/-!
## Norm-time order
-/
/-- The normal-time ordering on `FieldOpFreeAlgebra`. -/
def normTimeOrder : FieldOpFreeAlgebra 𝓕 →ₗ[] FieldOpFreeAlgebra 𝓕 :=
Basis.constr ofCrAnListFBasis fun φs =>
normTimeOrderSign φs • ofCrAnListF (normTimeOrderList φs)
@[inherit_doc normTimeOrder]
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "𝓣𝓝ᶠ(" a ")" => normTimeOrder a
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
end FieldOpFreeAlgebra
end FieldSpecification