/- 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.Algebras.OperatorAlgebra.NormalOrder import HepLean.PerturbationTheory.Algebras.StateAlgebra.TimeOrder /-! # Time contractions We define the state algebra of a field structure to be the free algebra generated by the states. -/ namespace FieldStruct variable {𝓕 : FieldStruct} open CrAnAlgebra noncomputable section namespace OperatorAlgebra variable (π“ž : 𝓕.OperatorAlgebra) open FieldStatistic /-- The time contraction of two States as an element of `π“ž.A` defined as their time ordering in the state algebra minus their normal ordering in the creation and annihlation algebra, both mapped to `π“ž.A`.. -/ def timeContract (Ο† ψ : 𝓕.States) : π“ž.A := π“ž.crAnF (ofStateAlgebra (StateAlgebra.timeOrder (StateAlgebra.ofState Ο† * StateAlgebra.ofState ψ)) - normalOrder (ofState Ο† * ofState ψ)) lemma timeContract_eq_smul (Ο† ψ : 𝓕.States) : π“ž.timeContract Ο† ψ = π“ž.crAnF (ofStateAlgebra (StateAlgebra.timeOrder (StateAlgebra.ofState Ο† * StateAlgebra.ofState ψ)) + (-1 : β„‚) β€’ normalOrder (ofState Ο† * ofState ψ)) := by rfl lemma timeContract_of_timeOrderRel (Ο† ψ : 𝓕.States) (h : timeOrderRel Ο† ψ) : π“ž.timeContract Ο† ψ = π“ž.crAnF (⟨anPart (StateAlgebra.ofState Ο†), ofState ΟˆβŸ©β‚›ca) := by conv_rhs => rw [ofState_eq_crPart_add_anPart] rw [map_add, map_add, crAnF_superCommute_anPart_anPart, superCommute_anPart_crPart] simp only [timeContract, instCommGroup.eq_1, Algebra.smul_mul_assoc, add_zero] rw [StateAlgebra.timeOrder_ofState_ofState_ordered h] rw [normalOrder_ofState_mul_ofState] rw [map_mul] simp only [ofStateAlgebra_ofState, instCommGroup.eq_1] rw [ofState_eq_crPart_add_anPart, ofState_eq_crPart_add_anPart] simp only [mul_add, add_mul] abel_nf lemma timeContract_of_not_timeOrderRel (Ο† ψ : 𝓕.States) (h : Β¬ timeOrderRel Ο† ψ) : π“ž.timeContract Ο† ψ = 𝓒(𝓕 |>β‚› Ο†, 𝓕 |>β‚› ψ) β€’ π“ž.timeContract ψ Ο† := by rw [timeContract_eq_smul] simp only [Int.reduceNeg, one_smul, map_add] rw [map_smul] rw [crAnF_normalOrder_ofState_ofState_swap] rw [StateAlgebra.timeOrder_ofState_ofState_not_ordered_eq_timeOrder h] rw [timeContract_eq_smul] simp only [FieldStatistic.instCommGroup.eq_1, map_smul, one_smul, map_add, smul_add] rw [smul_smul, smul_smul, mul_comm] lemma timeContract_mem_center (Ο† ψ : 𝓕.States) : π“ž.timeContract Ο† ψ ∈ Subalgebra.center β„‚ π“ž.A := by by_cases h : timeOrderRel Ο† ψ Β· rw [timeContract_of_timeOrderRel _ _ _ h] exact π“ž.crAnF_superCommute_anPart_ofState_mem_center _ _ Β· rw [timeContract_of_not_timeOrderRel _ _ _ h] refine Subalgebra.smul_mem (Subalgebra.center β„‚ π“ž.A) ?_ 𝓒(𝓕 |>β‚› Ο†, 𝓕 |>β‚› ψ) rw [timeContract_of_timeOrderRel] exact π“ž.crAnF_superCommute_anPart_ofState_mem_center _ _ have h1 := IsTotal.total (r := 𝓕.timeOrderRel) Ο† ψ simp_all lemma timeContract_zero_of_diff_grade (Ο† ψ : 𝓕.States) (h : (𝓕 |>β‚› Ο†) β‰  (𝓕 |>β‚› ψ)) : π“ž.timeContract Ο† ψ = 0 := by by_cases h1 : timeOrderRel Ο† ψ Β· rw [timeContract_of_timeOrderRel _ _ _ h1] rw [crAnF_superCommute_anPart_ofState_diff_grade_zero] exact h Β· rw [timeContract_of_not_timeOrderRel _ _ _ h1] rw [timeContract_of_timeOrderRel _ _ _] rw [crAnF_superCommute_anPart_ofState_diff_grade_zero] simp only [instCommGroup.eq_1, smul_zero] exact h.symm have ht := IsTotal.total (r := 𝓕.timeOrderRel) Ο† ψ simp_all end OperatorAlgebra end end FieldStruct