/- 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.CrAnFieldOp import HepLean.PerturbationTheory.FieldSpecification.CrAnSection /-! # Creation and annihlation free-algebra This module defines the creation and annihilation algebra for a field structure. The creation and annihilation algebra extends from the state algebra by adding information about whether a state is a creation or annihilation operator. The algebra is spanned by lists of creation/annihilation states. The main structures defined in this module are: * `FieldOpFreeAlgebra` - The creation and annihilation algebra * `ofCrAnOpF` - Maps a creation/annihilation state to the algebra * `ofCrAnListF` - Maps a list of creation/annihilation states to the algebra * `ofFieldOpF` - Maps a state to a sum of creation and annihilation operators * `crPartF` - The creation part of a state in the algebra * `anPartF` - The annihilation part of a state in the algebra * `superCommuteF` - The super commutator on the algebra The key lemmas show how these operators interact, particularly focusing on the super commutation relations between creation and annihilation operators. -/ namespace FieldSpecification variable {๐“• : FieldSpecification} /-- For a field specification `๐“•`, the algebra `๐“•.FieldOpFreeAlgebra` is the algebra generated by creation and annihilation parts of field operators defined in `๐“•.CrAnFieldOp`. It represents the algebra containing all possible products and linear combinations of creation and annihilation parts of field operators, without imposing any conditions. -/ abbrev FieldOpFreeAlgebra (๐“• : FieldSpecification) : Type := FreeAlgebra โ„‚ ๐“•.CrAnFieldOp namespace FieldOpFreeAlgebra /-- Maps a creation and annihlation state to the creation and annihlation free-algebra. -/ def ofCrAnOpF (ฯ† : ๐“•.CrAnFieldOp) : FieldOpFreeAlgebra ๐“• := FreeAlgebra.ฮน โ„‚ ฯ† /-- Maps a list creation and annihlation state to the creation and annihlation free-algebra by taking their product. -/ def ofCrAnListF (ฯ†s : List ๐“•.CrAnFieldOp) : FieldOpFreeAlgebra ๐“• := (List.map ofCrAnOpF ฯ†s).prod @[simp] lemma ofCrAnListF_nil : ofCrAnListF ([] : List ๐“•.CrAnFieldOp) = 1 := rfl lemma ofCrAnListF_cons (ฯ† : ๐“•.CrAnFieldOp) (ฯ†s : List ๐“•.CrAnFieldOp) : ofCrAnListF (ฯ† :: ฯ†s) = ofCrAnOpF ฯ† * ofCrAnListF ฯ†s := rfl lemma ofCrAnListF_append (ฯ†s ฯ†s' : List ๐“•.CrAnFieldOp) : ofCrAnListF (ฯ†s ++ ฯ†s') = ofCrAnListF ฯ†s * ofCrAnListF ฯ†s' := by simp [ofCrAnListF, List.map_append] lemma ofCrAnListF_singleton (ฯ† : ๐“•.CrAnFieldOp) : ofCrAnListF [ฯ†] = ofCrAnOpF ฯ† := by simp [ofCrAnListF] /-- Maps a state to the sum of creation and annihilation operators in creation and annihilation free-algebra. -/ def ofFieldOpF (ฯ† : ๐“•.FieldOp) : FieldOpFreeAlgebra ๐“• := โˆ‘ (i : ๐“•.fieldOpToCrAnType ฯ†), ofCrAnOpF โŸจฯ†, iโŸฉ /-- Maps a list of states to the creation and annihilation free-algebra by taking the product of their sums of creation and annihlation operators. Roughly `[ฯ†1, ฯ†2]` gets sent to `(ฯ†1แถœ+ ฯ†1แตƒ) * (ฯ†2แถœ+ ฯ†2แตƒ)` etc. -/ def ofFieldOpListF (ฯ†s : List ๐“•.FieldOp) : FieldOpFreeAlgebra ๐“• := (List.map ofFieldOpF ฯ†s).prod /-- Coercion from `List ๐“•.FieldOp` to `FieldOpFreeAlgebra ๐“•` through `ofFieldOpListF`. -/ instance : Coe (List ๐“•.FieldOp) (FieldOpFreeAlgebra ๐“•) := โŸจofFieldOpListFโŸฉ @[simp] lemma ofFieldOpListF_nil : ofFieldOpListF ([] : List ๐“•.FieldOp) = 1 := rfl lemma ofFieldOpListF_cons (ฯ† : ๐“•.FieldOp) (ฯ†s : List ๐“•.FieldOp) : ofFieldOpListF (ฯ† :: ฯ†s) = ofFieldOpF ฯ† * ofFieldOpListF ฯ†s := rfl lemma ofFieldOpListF_singleton (ฯ† : ๐“•.FieldOp) : ofFieldOpListF [ฯ†] = ofFieldOpF ฯ† := by simp [ofFieldOpListF] lemma ofFieldOpListF_append (ฯ†s ฯ†s' : List ๐“•.FieldOp) : ofFieldOpListF (ฯ†s ++ ฯ†s') = ofFieldOpListF ฯ†s * ofFieldOpListF ฯ†s' := by dsimp only [ofFieldOpListF] rw [List.map_append, List.prod_append] lemma ofFieldOpListF_sum (ฯ†s : List ๐“•.FieldOp) : ofFieldOpListF ฯ†s = โˆ‘ (s : CrAnSection ฯ†s), ofCrAnListF s.1 := by induction ฯ†s with | nil => simp | cons ฯ† ฯ†s ih => rw [CrAnSection.sum_cons] dsimp only [CrAnSection.cons, ofCrAnListF_cons] conv_rhs => enter [2, x] rw [โ† Finset.mul_sum] rw [โ† Finset.sum_mul, ofFieldOpListF_cons, โ† ih] rfl /-! ## Creation and annihilation parts of a state -/ /-- The algebra map taking an element of the free-state algbra to the part of it in the creation and annihlation free algebra spanned by creation operators. -/ def crPartF : ๐“•.FieldOp โ†’ ๐“•.FieldOpFreeAlgebra := fun ฯ† => match ฯ† with | FieldOp.inAsymp ฯ† => ofCrAnOpF โŸจFieldOp.inAsymp ฯ†, ()โŸฉ | FieldOp.position ฯ† => ofCrAnOpF โŸจFieldOp.position ฯ†, CreateAnnihilate.createโŸฉ | FieldOp.outAsymp _ => 0 @[simp] lemma crPartF_negAsymp (ฯ† : ๐“•.Fields ร— Lorentz.Contr 4) : crPartF (FieldOp.inAsymp ฯ†) = ofCrAnOpF โŸจFieldOp.inAsymp ฯ†, ()โŸฉ := by simp [crPartF] @[simp] lemma crPartF_position (ฯ† : ๐“•.Fields ร— SpaceTime) : crPartF (FieldOp.position ฯ†) = ofCrAnOpF โŸจFieldOp.position ฯ†, CreateAnnihilate.createโŸฉ := by simp [crPartF] @[simp] lemma crPartF_posAsymp (ฯ† : ๐“•.Fields ร— Lorentz.Contr 4) : crPartF (FieldOp.outAsymp ฯ†) = 0 := by simp [crPartF] /-- The algebra map taking an element of the free-state algbra to the part of it in the creation and annihilation free algebra spanned by annihilation operators. -/ def anPartF : ๐“•.FieldOp โ†’ ๐“•.FieldOpFreeAlgebra := fun ฯ† => match ฯ† with | FieldOp.inAsymp _ => 0 | FieldOp.position ฯ† => ofCrAnOpF โŸจFieldOp.position ฯ†, CreateAnnihilate.annihilateโŸฉ | FieldOp.outAsymp ฯ† => ofCrAnOpF โŸจFieldOp.outAsymp ฯ†, ()โŸฉ @[simp] lemma anPartF_negAsymp (ฯ† : ๐“•.Fields ร— Lorentz.Contr 4) : anPartF (FieldOp.inAsymp ฯ†) = 0 := by simp [anPartF] @[simp] lemma anPartF_position (ฯ† : ๐“•.Fields ร— SpaceTime) : anPartF (FieldOp.position ฯ†) = ofCrAnOpF โŸจFieldOp.position ฯ†, CreateAnnihilate.annihilateโŸฉ := by simp [anPartF] @[simp] lemma anPartF_posAsymp (ฯ† : ๐“•.Fields ร— Lorentz.Contr 4) : anPartF (FieldOp.outAsymp ฯ†) = ofCrAnOpF โŸจFieldOp.outAsymp ฯ†, ()โŸฉ := by simp [anPartF] lemma ofFieldOpF_eq_crPartF_add_anPartF (ฯ† : ๐“•.FieldOp) : ofFieldOpF ฯ† = crPartF ฯ† + anPartF ฯ† := by rw [ofFieldOpF] cases ฯ† with | inAsymp ฯ† => simp [fieldOpToCrAnType] | position ฯ† => simp [fieldOpToCrAnType, CreateAnnihilate.sum_eq] | outAsymp ฯ† => simp [fieldOpToCrAnType] /-! ## The basis of the creation and annihlation free-algebra. -/ /-- The basis of the free creation and annihilation algebra formed by lists of CrAnFieldOp. -/ noncomputable def ofCrAnListFBasis : Basis (List ๐“•.CrAnFieldOp) โ„‚ (FieldOpFreeAlgebra ๐“•) where repr := FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv @[simp] lemma ofListBasis_eq_ofList (ฯ†s : List ๐“•.CrAnFieldOp) : ofCrAnListFBasis ฯ†s = ofCrAnListF ฯ†s := by simp only [ofCrAnListFBasis, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply, Basis.coe_ofRepr, AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_apply, AlgEquiv.ofAlgHom_symm_apply, ofCrAnListF] erw [MonoidAlgebra.lift_apply] simp only [zero_smul, Finsupp.sum_single_index, one_smul] rw [@FreeMonoid.lift_apply] match ฯ†s with | [] => rfl | ฯ† :: ฯ†s => erw [List.map_cons] /-! ## Some useful multi-linear maps. -/ /-- The bi-linear map associated with multiplication in `FieldOpFreeAlgebra`. -/ noncomputable def mulLinearMap : FieldOpFreeAlgebra ๐“• โ†’โ‚—[โ„‚] FieldOpFreeAlgebra ๐“• โ†’โ‚—[โ„‚] FieldOpFreeAlgebra ๐“• where toFun a := { toFun := fun b => a * b, map_add' := fun c d => by simp [mul_add] map_smul' := by simp} map_add' := fun a b => by ext c simp [add_mul] map_smul' := by intros ext c simp [smul_mul'] lemma mulLinearMap_apply (a b : FieldOpFreeAlgebra ๐“•) : mulLinearMap a b = a * b := rfl /-- The linear map associated with scalar-multiplication in `FieldOpFreeAlgebra`. -/ noncomputable def smulLinearMap (c : โ„‚) : FieldOpFreeAlgebra ๐“• โ†’โ‚—[โ„‚] FieldOpFreeAlgebra ๐“• where toFun a := c โ€ข a map_add' := by simp map_smul' m x := by simp [smul_smul, RingHom.id_apply, NonUnitalNormedCommRing.mul_comm] end FieldOpFreeAlgebra end FieldSpecification