refactor: Rename States to FieldOps
This commit is contained in:
parent
171e80fc04
commit
8f41de5785
36 changed files with 946 additions and 946 deletions
|
@ -3,7 +3,7 @@ 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.CrAnFieldOp
|
||||
import HepLean.PerturbationTheory.FieldSpecification.CrAnSection
|
||||
/-!
|
||||
|
||||
|
@ -35,63 +35,63 @@ namespace FieldSpecification
|
|||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
/-- The creation and annihlation free-algebra.
|
||||
The free algebra generated by `CrAnStates`,
|
||||
The free algebra generated by `CrAnFieldOp`,
|
||||
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 `FieldOpFreeAlgebra` is spanned by lists of `CrAnStates`. -/
|
||||
abbrev FieldOpFreeAlgebra (𝓕 : FieldSpecification) : Type := FreeAlgebra ℂ 𝓕.CrAnStates
|
||||
As a module `FieldOpFreeAlgebra` is spanned by lists of `CrAnFieldOp`. -/
|
||||
abbrev FieldOpFreeAlgebra (𝓕 : FieldSpecification) : Type := FreeAlgebra ℂ 𝓕.CrAnFieldOp
|
||||
|
||||
namespace FieldOpFreeAlgebra
|
||||
|
||||
/-- Maps a creation and annihlation state to the creation and annihlation free-algebra. -/
|
||||
def ofCrAnOpF (φ : 𝓕.CrAnStates) : FieldOpFreeAlgebra 𝓕 :=
|
||||
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 𝓕.CrAnStates) : FieldOpFreeAlgebra 𝓕 := (List.map ofCrAnOpF φs).prod
|
||||
def ofCrAnListF (φs : List 𝓕.CrAnFieldOp) : FieldOpFreeAlgebra 𝓕 := (List.map ofCrAnOpF φs).prod
|
||||
|
||||
@[simp]
|
||||
lemma ofCrAnListF_nil : ofCrAnListF ([] : List 𝓕.CrAnStates) = 1 := rfl
|
||||
lemma ofCrAnListF_nil : ofCrAnListF ([] : List 𝓕.CrAnFieldOp) = 1 := rfl
|
||||
|
||||
lemma ofCrAnListF_cons (φ : 𝓕.CrAnStates) (φs : List 𝓕.CrAnStates) :
|
||||
lemma ofCrAnListF_cons (φ : 𝓕.CrAnFieldOp) (φs : List 𝓕.CrAnFieldOp) :
|
||||
ofCrAnListF (φ :: φs) = ofCrAnOpF φ * ofCrAnListF φs := rfl
|
||||
|
||||
lemma ofCrAnListF_append (φs φs' : List 𝓕.CrAnStates) :
|
||||
lemma ofCrAnListF_append (φs φs' : List 𝓕.CrAnFieldOp) :
|
||||
ofCrAnListF (φs ++ φs') = ofCrAnListF φs * ofCrAnListF φs' := by
|
||||
simp [ofCrAnListF, List.map_append]
|
||||
|
||||
lemma ofCrAnListF_singleton (φ : 𝓕.CrAnStates) :
|
||||
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 (φ : 𝓕.States) : FieldOpFreeAlgebra 𝓕 :=
|
||||
∑ (i : 𝓕.statesToCrAnType φ), ofCrAnOpF ⟨φ, i⟩
|
||||
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 𝓕.States) : FieldOpFreeAlgebra 𝓕 := (List.map ofFieldOpF φs).prod
|
||||
def ofFieldOpListF (φs : List 𝓕.FieldOp) : FieldOpFreeAlgebra 𝓕 := (List.map ofFieldOpF φs).prod
|
||||
|
||||
/-- Coercion from `List 𝓕.States` to `FieldOpFreeAlgebra 𝓕` through `ofFieldOpListF`. -/
|
||||
instance : Coe (List 𝓕.States) (FieldOpFreeAlgebra 𝓕) := ⟨ofFieldOpListF⟩
|
||||
/-- Coercion from `List 𝓕.FieldOp` to `FieldOpFreeAlgebra 𝓕` through `ofFieldOpListF`. -/
|
||||
instance : Coe (List 𝓕.FieldOp) (FieldOpFreeAlgebra 𝓕) := ⟨ofFieldOpListF⟩
|
||||
|
||||
@[simp]
|
||||
lemma ofFieldOpListF_nil : ofFieldOpListF ([] : List 𝓕.States) = 1 := rfl
|
||||
lemma ofFieldOpListF_nil : ofFieldOpListF ([] : List 𝓕.FieldOp) = 1 := rfl
|
||||
|
||||
lemma ofFieldOpListF_cons (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
lemma ofFieldOpListF_cons (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) :
|
||||
ofFieldOpListF (φ :: φs) = ofFieldOpF φ * ofFieldOpListF φs := rfl
|
||||
|
||||
lemma ofFieldOpListF_singleton (φ : 𝓕.States) :
|
||||
lemma ofFieldOpListF_singleton (φ : 𝓕.FieldOp) :
|
||||
ofFieldOpListF [φ] = ofFieldOpF φ := by simp [ofFieldOpListF]
|
||||
|
||||
lemma ofFieldOpListF_append (φs φs' : List 𝓕.States) :
|
||||
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 𝓕.States) :
|
||||
lemma ofFieldOpListF_sum (φs : List 𝓕.FieldOp) :
|
||||
ofFieldOpListF φs = ∑ (s : CrAnSection φs), ofCrAnListF s.1 := by
|
||||
induction φs with
|
||||
| nil => simp
|
||||
|
@ -113,60 +113,60 @@ lemma ofFieldOpListF_sum (φs : List 𝓕.States) :
|
|||
/-- 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 → 𝓕.FieldOpFreeAlgebra := fun φ =>
|
||||
def crPartF : 𝓕.FieldOp → 𝓕.FieldOpFreeAlgebra := fun φ =>
|
||||
match φ with
|
||||
| States.inAsymp φ => ofCrAnOpF ⟨States.inAsymp φ, ()⟩
|
||||
| States.position φ => ofCrAnOpF ⟨States.position φ, CreateAnnihilate.create⟩
|
||||
| States.outAsymp _ => 0
|
||||
| FieldOp.inAsymp φ => ofCrAnOpF ⟨FieldOp.inAsymp φ, ()⟩
|
||||
| FieldOp.position φ => ofCrAnOpF ⟨FieldOp.position φ, CreateAnnihilate.create⟩
|
||||
| FieldOp.outAsymp _ => 0
|
||||
|
||||
@[simp]
|
||||
lemma crPartF_negAsymp (φ : 𝓕.IncomingAsymptotic) :
|
||||
crPartF (States.inAsymp φ) = ofCrAnOpF ⟨States.inAsymp φ, ()⟩ := by
|
||||
crPartF (FieldOp.inAsymp φ) = ofCrAnOpF ⟨FieldOp.inAsymp φ, ()⟩ := by
|
||||
simp [crPartF]
|
||||
|
||||
@[simp]
|
||||
lemma crPartF_position (φ : 𝓕.PositionStates) :
|
||||
crPartF (States.position φ) =
|
||||
ofCrAnOpF ⟨States.position φ, CreateAnnihilate.create⟩ := by
|
||||
lemma crPartF_position (φ : 𝓕.PositionFieldOp) :
|
||||
crPartF (FieldOp.position φ) =
|
||||
ofCrAnOpF ⟨FieldOp.position φ, CreateAnnihilate.create⟩ := by
|
||||
simp [crPartF]
|
||||
|
||||
@[simp]
|
||||
lemma crPartF_posAsymp (φ : 𝓕.OutgoingAsymptotic) :
|
||||
crPartF (States.outAsymp φ) = 0 := by
|
||||
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 : 𝓕.States → 𝓕.FieldOpFreeAlgebra := fun φ =>
|
||||
def anPartF : 𝓕.FieldOp → 𝓕.FieldOpFreeAlgebra := fun φ =>
|
||||
match φ with
|
||||
| States.inAsymp _ => 0
|
||||
| States.position φ => ofCrAnOpF ⟨States.position φ, CreateAnnihilate.annihilate⟩
|
||||
| States.outAsymp φ => ofCrAnOpF ⟨States.outAsymp φ, ()⟩
|
||||
| FieldOp.inAsymp _ => 0
|
||||
| FieldOp.position φ => ofCrAnOpF ⟨FieldOp.position φ, CreateAnnihilate.annihilate⟩
|
||||
| FieldOp.outAsymp φ => ofCrAnOpF ⟨FieldOp.outAsymp φ, ()⟩
|
||||
|
||||
@[simp]
|
||||
lemma anPartF_negAsymp (φ : 𝓕.IncomingAsymptotic) :
|
||||
anPartF (States.inAsymp φ) = 0 := by
|
||||
anPartF (FieldOp.inAsymp φ) = 0 := by
|
||||
simp [anPartF]
|
||||
|
||||
@[simp]
|
||||
lemma anPartF_position (φ : 𝓕.PositionStates) :
|
||||
anPartF (States.position φ) =
|
||||
ofCrAnOpF ⟨States.position φ, CreateAnnihilate.annihilate⟩ := by
|
||||
lemma anPartF_position (φ : 𝓕.PositionFieldOp) :
|
||||
anPartF (FieldOp.position φ) =
|
||||
ofCrAnOpF ⟨FieldOp.position φ, CreateAnnihilate.annihilate⟩ := by
|
||||
simp [anPartF]
|
||||
|
||||
@[simp]
|
||||
lemma anPartF_posAsymp (φ : 𝓕.OutgoingAsymptotic) :
|
||||
anPartF (States.outAsymp φ) = ofCrAnOpF ⟨States.outAsymp φ, ()⟩ := by
|
||||
anPartF (FieldOp.outAsymp φ) = ofCrAnOpF ⟨FieldOp.outAsymp φ, ()⟩ := by
|
||||
simp [anPartF]
|
||||
|
||||
lemma ofFieldOpF_eq_crPartF_add_anPartF (φ : 𝓕.States) :
|
||||
lemma ofFieldOpF_eq_crPartF_add_anPartF (φ : 𝓕.FieldOp) :
|
||||
ofFieldOpF φ = crPartF φ + anPartF φ := by
|
||||
rw [ofFieldOpF]
|
||||
cases φ with
|
||||
| inAsymp φ => simp [statesToCrAnType]
|
||||
| position φ => simp [statesToCrAnType, CreateAnnihilate.sum_eq]
|
||||
| outAsymp φ => simp [statesToCrAnType]
|
||||
| inAsymp φ => simp [fieldOpToCrAnType]
|
||||
| position φ => simp [fieldOpToCrAnType, CreateAnnihilate.sum_eq]
|
||||
| outAsymp φ => simp [fieldOpToCrAnType]
|
||||
|
||||
/-!
|
||||
|
||||
|
@ -174,12 +174,12 @@ lemma ofFieldOpF_eq_crPartF_add_anPartF (φ : 𝓕.States) :
|
|||
|
||||
-/
|
||||
|
||||
/-- The basis of the free creation and annihilation algebra formed by lists of CrAnStates. -/
|
||||
noncomputable def ofCrAnListFBasis : Basis (List 𝓕.CrAnStates) ℂ (FieldOpFreeAlgebra 𝓕) where
|
||||
/-- 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 𝓕.CrAnStates) :
|
||||
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,
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue