/- 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