/- 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.CrAnStates 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: * `CrAnAlgebra` - The creation and annihilation algebra * `ofCrAnState` - Maps a creation/annihilation state to the algebra * `ofCrAnList` - Maps a list of creation/annihilation states to the algebra * `ofState` - 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} /-- The creation and annihlation free-algebra. The free algebra generated by `CrAnStates`, that is a position based states or assymptotic states with a specification of whether the state is a creation or annihlation state. As a module `CrAnAlgebra` is spanned by lists of `CrAnStates`. -/ abbrev CrAnAlgebra (๐“• : FieldSpecification) : Type := FreeAlgebra โ„‚ ๐“•.CrAnStates namespace CrAnAlgebra /-- Maps a creation and annihlation state to the creation and annihlation free-algebra. -/ def ofCrAnState (ฯ† : ๐“•.CrAnStates) : CrAnAlgebra ๐“• := FreeAlgebra.ฮน โ„‚ ฯ† /-- Maps a list creation and annihlation state to the creation and annihlation free-algebra by taking their product. -/ def ofCrAnList (ฯ†s : List ๐“•.CrAnStates) : CrAnAlgebra ๐“• := (List.map ofCrAnState ฯ†s).prod @[simp] lemma ofCrAnList_nil : ofCrAnList ([] : List ๐“•.CrAnStates) = 1 := rfl lemma ofCrAnList_cons (ฯ† : ๐“•.CrAnStates) (ฯ†s : List ๐“•.CrAnStates) : ofCrAnList (ฯ† :: ฯ†s) = ofCrAnState ฯ† * ofCrAnList ฯ†s := rfl lemma ofCrAnList_append (ฯ†s ฯ†s' : List ๐“•.CrAnStates) : ofCrAnList (ฯ†s ++ ฯ†s') = ofCrAnList ฯ†s * ofCrAnList ฯ†s' := by simp [ofCrAnList, List.map_append] lemma ofCrAnList_singleton (ฯ† : ๐“•.CrAnStates) : ofCrAnList [ฯ†] = ofCrAnState ฯ† := by simp [ofCrAnList] /-- Maps a state to the sum of creation and annihilation operators in creation and annihilation free-algebra. -/ def ofState (ฯ† : ๐“•.States) : CrAnAlgebra ๐“• := โˆ‘ (i : ๐“•.statesToCrAnType ฯ†), ofCrAnState โŸจฯ†, 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 ofStateList (ฯ†s : List ๐“•.States) : CrAnAlgebra ๐“• := (List.map ofState ฯ†s).prod /-- Coercion from `List ๐“•.States` to `CrAnAlgebra ๐“•` through `ofStateList`. -/ instance : Coe (List ๐“•.States) (CrAnAlgebra ๐“•) := โŸจofStateListโŸฉ @[simp] lemma ofStateList_nil : ofStateList ([] : List ๐“•.States) = 1 := rfl lemma ofStateList_cons (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : ofStateList (ฯ† :: ฯ†s) = ofState ฯ† * ofStateList ฯ†s := rfl lemma ofStateList_singleton (ฯ† : ๐“•.States) : ofStateList [ฯ†] = ofState ฯ† := by simp [ofStateList] lemma ofStateList_append (ฯ†s ฯ†s' : List ๐“•.States) : ofStateList (ฯ†s ++ ฯ†s') = ofStateList ฯ†s * ofStateList ฯ†s' := by dsimp only [ofStateList] rw [List.map_append, List.prod_append] lemma ofStateList_sum (ฯ†s : List ๐“•.States) : ofStateList ฯ†s = โˆ‘ (s : CrAnSection ฯ†s), ofCrAnList s.1 := by induction ฯ†s with | nil => simp | cons ฯ† ฯ†s ih => rw [CrAnSection.sum_cons] dsimp only [CrAnSection.cons, ofCrAnList_cons] conv_rhs => enter [2, x] rw [โ† Finset.mul_sum] rw [โ† Finset.sum_mul, ofStateList_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 : ๐“•.States โ†’ ๐“•.CrAnAlgebra := fun ฯ† => match ฯ† with | States.inAsymp ฯ† => ofCrAnState โŸจStates.inAsymp ฯ†, ()โŸฉ | States.position ฯ† => ofCrAnState โŸจStates.position ฯ†, CreateAnnihilate.createโŸฉ | States.outAsymp _ => 0 @[simp] lemma crPartF_negAsymp (ฯ† : ๐“•.IncomingAsymptotic) : crPartF (States.inAsymp ฯ†) = ofCrAnState โŸจStates.inAsymp ฯ†, ()โŸฉ := by simp [crPartF] @[simp] lemma crPartF_position (ฯ† : ๐“•.PositionStates) : crPartF (States.position ฯ†) = ofCrAnState โŸจStates.position ฯ†, CreateAnnihilate.createโŸฉ := by simp [crPartF] @[simp] lemma crPartF_posAsymp (ฯ† : ๐“•.OutgoingAsymptotic) : crPartF (States.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 : ๐“•.States โ†’ ๐“•.CrAnAlgebra := fun ฯ† => match ฯ† with | States.inAsymp _ => 0 | States.position ฯ† => ofCrAnState โŸจStates.position ฯ†, CreateAnnihilate.annihilateโŸฉ | States.outAsymp ฯ† => ofCrAnState โŸจStates.outAsymp ฯ†, ()โŸฉ @[simp] lemma anPartF_negAsymp (ฯ† : ๐“•.IncomingAsymptotic) : anPartF (States.inAsymp ฯ†) = 0 := by simp [anPartF] @[simp] lemma anPartF_position (ฯ† : ๐“•.PositionStates) : anPartF (States.position ฯ†) = ofCrAnState โŸจStates.position ฯ†, CreateAnnihilate.annihilateโŸฉ := by simp [anPartF] @[simp] lemma anPartF_posAsymp (ฯ† : ๐“•.OutgoingAsymptotic) : anPartF (States.outAsymp ฯ†) = ofCrAnState โŸจStates.outAsymp ฯ†, ()โŸฉ := by simp [anPartF] lemma ofState_eq_crPartF_add_anPartF (ฯ† : ๐“•.States) : ofState ฯ† = crPartF ฯ† + anPartF ฯ† := by rw [ofState] cases ฯ† with | inAsymp ฯ† => simp [statesToCrAnType] | position ฯ† => simp [statesToCrAnType, CreateAnnihilate.sum_eq] | outAsymp ฯ† => simp [statesToCrAnType] /-! ## The basis of the creation and annihlation free-algebra. -/ /-- The basis of the free creation and annihilation algebra formed by lists of CrAnStates. -/ noncomputable def ofCrAnListBasis : Basis (List ๐“•.CrAnStates) โ„‚ (CrAnAlgebra ๐“•) where repr := FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv @[simp] lemma ofListBasis_eq_ofList (ฯ†s : List ๐“•.CrAnStates) : ofCrAnListBasis ฯ†s = ofCrAnList ฯ†s := by simp only [ofCrAnListBasis, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply, Basis.coe_ofRepr, AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_apply, AlgEquiv.ofAlgHom_symm_apply, ofCrAnList] 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 `CrAnAlgebra`. -/ noncomputable def mulLinearMap : CrAnAlgebra ๐“• โ†’โ‚—[โ„‚] CrAnAlgebra ๐“• โ†’โ‚—[โ„‚] CrAnAlgebra ๐“• 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 : CrAnAlgebra ๐“•) : mulLinearMap a b = a * b := rfl /-- The linear map associated with scalar-multiplication in `CrAnAlgebra`. -/ noncomputable def smulLinearMap (c : โ„‚) : CrAnAlgebra ๐“• โ†’โ‚—[โ„‚] CrAnAlgebra ๐“• where toFun a := c โ€ข a map_add' := by simp map_smul' m x := by simp [smul_smul, RingHom.id_apply, NonUnitalNormedCommRing.mul_comm] end CrAnAlgebra end FieldSpecification