refactor: Rename CrAnAlgebra
This commit is contained in:
parent
9a5676e134
commit
b0735a1e13
16 changed files with 214 additions and 214 deletions
|
@ -0,0 +1,47 @@
|
|||
/-
|
||||
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.Algebras.FieldOpFreeAlgebra.SuperCommute
|
||||
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 ofCrAnListBasis ℂ fun φs =>
|
||||
normTimeOrderSign φs • ofCrAnList (normTimeOrderList φs)
|
||||
|
||||
@[inherit_doc normTimeOrder]
|
||||
scoped[FieldSpecification.FieldOpFreeAlgebra] notation "𝓣𝓝ᶠ(" a ")" => normTimeOrder a
|
||||
|
||||
lemma normTimeOrder_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
𝓣𝓝ᶠ(ofCrAnList φs) = normTimeOrderSign φs • ofCrAnList (normTimeOrderList φs) := by
|
||||
rw [← ofListBasis_eq_ofList]
|
||||
simp only [normTimeOrder, Basis.constr_basis]
|
||||
|
||||
end
|
||||
|
||||
end FieldOpFreeAlgebra
|
||||
|
||||
end FieldSpecification
|
Loading…
Add table
Add a link
Reference in a new issue