2025-01-27 12:19:21 +00:00
|
|
|
|
/-
|
|
|
|
|
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 12:12:36 +00:00
|
|
|
|
import HepLean.PerturbationTheory.FieldOpFreeAlgebra.Basic
|
2025-01-27 12:19:21 +00:00
|
|
|
|
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
|
|
|
|
/-!
|
|
|
|
|
|
2025-02-03 11:05:43 +00:00
|
|
|
|
# Norm-time Ordering in the FieldOpFreeAlgebra
|
2025-01-27 12:19:21 +00:00
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
namespace FieldSpecification
|
|
|
|
|
variable {𝓕 : FieldSpecification}
|
|
|
|
|
open FieldStatistic
|
|
|
|
|
|
2025-02-03 11:05:43 +00:00
|
|
|
|
namespace FieldOpFreeAlgebra
|
2025-01-27 12:19:21 +00:00
|
|
|
|
|
|
|
|
|
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 𝓕 :=
|
2025-02-03 11:21:11 +00:00
|
|
|
|
Basis.constr ofCrAnListFBasis ℂ fun φs =>
|
|
|
|
|
normTimeOrderSign φs • ofCrAnListF (normTimeOrderList φs)
|
2025-01-27 12:19:21 +00:00
|
|
|
|
|
|
|
|
|
@[inherit_doc normTimeOrder]
|
2025-02-03 11:05:43 +00:00
|
|
|
|
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "𝓣𝓝ᶠ(" a ")" => normTimeOrder a
|
2025-01-27 12:19:21 +00:00
|
|
|
|
|
2025-02-03 11:28:14 +00:00
|
|
|
|
lemma normTimeOrder_ofCrAnListF (φs : List 𝓕.CrAnFieldOp) :
|
2025-02-03 11:21:11 +00:00
|
|
|
|
𝓣𝓝ᶠ(ofCrAnListF φs) = normTimeOrderSign φs • ofCrAnListF (normTimeOrderList φs) := by
|
2025-01-27 12:19:21 +00:00
|
|
|
|
rw [← ofListBasis_eq_ofList]
|
|
|
|
|
simp only [normTimeOrder, Basis.constr_basis]
|
|
|
|
|
|
|
|
|
|
end
|
|
|
|
|
|
2025-02-03 11:05:43 +00:00
|
|
|
|
end FieldOpFreeAlgebra
|
2025-01-27 12:19:21 +00:00
|
|
|
|
|
|
|
|
|
end FieldSpecification
|