/- 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.FieldOpAlgebra.NormalOrder import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.TimeOrder /-! # Time contractions We define the state algebra of a field structure to be the free algebra generated by the states. -/ namespace FieldSpecification variable {๐“• : FieldSpecification} open CrAnAlgebra noncomputable section namespace FieldOpAlgebra 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) : ๐“•.FieldOpAlgebra := ๐“ฃ(ofFieldOp ฯ† * ofFieldOp ฯˆ) - ๐“(ofFieldOp ฯ† * ofFieldOp ฯˆ) lemma timeContract_eq_smul (ฯ† ฯˆ : ๐“•.States) : timeContract ฯ† ฯˆ = ๐“ฃ(ofFieldOp ฯ† * ofFieldOp ฯˆ) + (-1 : โ„‚) โ€ข ๐“(ofFieldOp ฯ† * ofFieldOp ฯˆ) := by rfl lemma timeContract_of_timeOrderRel (ฯ† ฯˆ : ๐“•.States) (h : timeOrderRel ฯ† ฯˆ) : timeContract ฯ† ฯˆ = [anPart ฯ†, ofFieldOp ฯˆ]โ‚› := by conv_rhs => rw [ofFieldOp_eq_crPart_add_anPart] rw [map_add, superCommute_anPart_anPart, superCommute_anPart_crPart] simp only [timeContract, instCommGroup.eq_1, Algebra.smul_mul_assoc, add_zero] rw [timeOrder_ofFieldOp_ofFieldOp_ordered h] rw [normalOrder_ofFieldOp_mul_ofFieldOp] simp only [instCommGroup.eq_1] rw [ofFieldOp_eq_crPart_add_anPart, ofFieldOp_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 [normalOrder_ofFieldOp_ofFieldOp_swap] rw [timeOrder_ofFieldOp_ofFieldOp_not_ordered_eq_timeOrder h] rw [timeContract_eq_smul] simp only [instCommGroup.eq_1, map_smul, map_add, smul_add] rw [smul_smul, smul_smul, mul_comm] lemma timeContract_mem_center (ฯ† ฯˆ : ๐“•.States) : timeContract ฯ† ฯˆ โˆˆ Subalgebra.center โ„‚ ๐“•.FieldOpAlgebra := by by_cases h : timeOrderRel ฯ† ฯˆ ยท rw [timeContract_of_timeOrderRel _ _ h] exact superCommute_anPart_ofFieldOp_mem_center ฯ† ฯˆ ยท rw [timeContract_of_not_timeOrderRel _ _ h] refine Subalgebra.smul_mem (Subalgebra.center โ„‚ _) ?_ ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯˆ) rw [timeContract_of_timeOrderRel] exact superCommute_anPart_ofFieldOp_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 [superCommute_anPart_ofState_diff_grade_zero] exact h ยท rw [timeContract_of_not_timeOrderRel _ _ h1] rw [timeContract_of_timeOrderRel _ _ _] rw [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 FieldOpAlgebra end end FieldSpecification