Skip to content

Commit

Permalink
Merge pull request #299 from HEPLean/FieldOpAlgebra
Browse files Browse the repository at this point in the history
feat: Remove state algebra, add time order CrAnAlgebra
  • Loading branch information
jstoobysmith authored Jan 27, 2025
2 parents 0dfb2b7 + ec2e1e7 commit acfd0ed
Show file tree
Hide file tree
Showing 20 changed files with 496 additions and 293 deletions.
3 changes: 1 addition & 2 deletions HepLean.lean
Original file line number Diff line number Diff line change
Expand Up @@ -124,13 +124,12 @@ import HepLean.Meta.TransverseTactics
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.Basic
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.NormalOrder
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.SuperCommute
import HepLean.PerturbationTheory.Algebras.CrAnAlgebra.TimeOrder
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.Basic
import HepLean.PerturbationTheory.Algebras.FieldOpAlgebra.NormalOrder
import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.Basic
import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.NormalOrder
import HepLean.PerturbationTheory.Algebras.ProtoOperatorAlgebra.TimeContraction
import HepLean.PerturbationTheory.Algebras.StateAlgebra.Basic
import HepLean.PerturbationTheory.Algebras.StateAlgebra.TimeOrder
import HepLean.PerturbationTheory.CreateAnnihilate
import HepLean.PerturbationTheory.FeynmanDiagrams.Basic
import HepLean.PerturbationTheory.FeynmanDiagrams.Instances.ComplexScalar
Expand Down
53 changes: 16 additions & 37 deletions HepLean/PerturbationTheory/Algebras/CrAnAlgebra/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -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.Algebras.StateAlgebra.Basic
import HepLean.PerturbationTheory.FieldSpecification.CrAnStates
import HepLean.PerturbationTheory.FieldSpecification.CrAnSection
/-!
Expand Down Expand Up @@ -43,8 +43,6 @@ abbrev CrAnAlgebra (𝓕 : FieldSpecification) : Type := FreeAlgebra ℂ 𝓕.Cr

namespace CrAnAlgebra

open StateAlgebra

/-- Maps a creation and annihlation state to the creation and annihlation free-algebra. -/
def ofCrAnState (φ : 𝓕.CrAnStates) : CrAnAlgebra 𝓕 :=
FreeAlgebra.ι ℂ φ
Expand All @@ -71,15 +69,6 @@ lemma ofCrAnList_singleton (φ : 𝓕.CrAnStates) :
def ofState (φ : 𝓕.States) : CrAnAlgebra 𝓕 :=
∑ (i : 𝓕.statesToCrAnType φ), ofCrAnState ⟨φ, i⟩

/-- The algebra map from the state free-algebra to the creation and annihlation free-algebra. -/
def ofStateAlgebra : StateAlgebra 𝓕 →ₐ[ℂ] CrAnAlgebra 𝓕 :=
FreeAlgebra.lift ℂ ofState

@[simp]
lemma ofStateAlgebra_ofState (φ : 𝓕.States) :
ofStateAlgebra (StateAlgebra.ofState φ) = ofState φ := by
simp [ofStateAlgebra, StateAlgebra.ofState]

/-- 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. -/
Expand All @@ -102,14 +91,6 @@ lemma ofStateList_append (φs φs' : List 𝓕.States) :
dsimp only [ofStateList]
rw [List.map_append, List.prod_append]

lemma ofStateAlgebra_ofList_eq_ofStateList : (φs : List 𝓕.States) →
ofStateAlgebra (ofList φs) = ofStateList φs
| [] => by simp [ofStateList]
| φ :: φs => by
rw [ofStateList_cons, StateAlgebra.ofList_cons, map_mul, ofStateAlgebra_ofState,
mul_eq_mul_left_iff]
exact .inl (ofStateAlgebra_ofList_eq_ofStateList φs)

lemma ofStateList_sum (φs : List 𝓕.States) :
ofStateList φs = ∑ (s : CrAnSection φs), ofCrAnList s.1 := by
induction φs with
Expand All @@ -132,57 +113,55 @@ lemma ofStateList_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 crPart : 𝓕.StateAlgebra →ₐ[ℂ] 𝓕.CrAnAlgebra :=
FreeAlgebra.lift ℂ fun φ =>
def crPart : 𝓕.States → 𝓕.CrAnAlgebra := fun φ =>
match φ with
| States.inAsymp φ => ofCrAnState ⟨States.inAsymp φ, ()⟩
| States.position φ => ofCrAnState ⟨States.position φ, CreateAnnihilate.create⟩
| States.outAsymp _ => 0

@[simp]
lemma crPart_negAsymp (φ : 𝓕.IncomingAsymptotic) :
crPart (StateAlgebra.ofState (States.inAsymp φ)) = ofCrAnState ⟨States.inAsymp φ, ()⟩ := by
simp [crPart, StateAlgebra.ofState]
crPart (States.inAsymp φ) = ofCrAnState ⟨States.inAsymp φ, ()⟩ := by
simp [crPart]

@[simp]
lemma crPart_position (φ : 𝓕.PositionStates) :
crPart (StateAlgebra.ofState (States.position φ)) =
crPart (States.position φ) =
ofCrAnState ⟨States.position φ, CreateAnnihilate.create⟩ := by
simp [crPart, StateAlgebra.ofState]
simp [crPart]

@[simp]
lemma crPart_posAsymp (φ : 𝓕.OutgoingAsymptotic) :
crPart (StateAlgebra.ofState (States.outAsymp φ)) = 0 := by
simp [crPart, StateAlgebra.ofState]
crPart (States.outAsymp φ) = 0 := by
simp [crPart]

/-- 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 anPart : 𝓕.StateAlgebra →ₐ[ℂ] 𝓕.CrAnAlgebra :=
FreeAlgebra.lift ℂ fun φ =>
def anPart : 𝓕.States → 𝓕.CrAnAlgebra := fun φ =>
match φ with
| States.inAsymp _ => 0
| States.position φ => ofCrAnState ⟨States.position φ, CreateAnnihilate.annihilate⟩
| States.outAsymp φ => ofCrAnState ⟨States.outAsymp φ, ()⟩

@[simp]
lemma anPart_negAsymp (φ : 𝓕.IncomingAsymptotic) :
anPart (StateAlgebra.ofState (States.inAsymp φ)) = 0 := by
simp [anPart, StateAlgebra.ofState]
anPart (States.inAsymp φ) = 0 := by
simp [anPart]

@[simp]
lemma anPart_position (φ : 𝓕.PositionStates) :
anPart (StateAlgebra.ofState (States.position φ)) =
anPart (States.position φ) =
ofCrAnState ⟨States.position φ, CreateAnnihilate.annihilate⟩ := by
simp [anPart, StateAlgebra.ofState]
simp [anPart]

@[simp]
lemma anPart_posAsymp (φ : 𝓕.OutgoingAsymptotic) :
anPart (StateAlgebra.ofState (States.outAsymp φ)) = ofCrAnState ⟨States.outAsymp φ, ()⟩ := by
simp [anPart, StateAlgebra.ofState]
anPart (States.outAsymp φ) = ofCrAnState ⟨States.outAsymp φ, ()⟩ := by
simp [anPart]

lemma ofState_eq_crPart_add_anPart (φ : 𝓕.States) :
ofState φ = crPart (StateAlgebra.ofState φ) + anPart (StateAlgebra.ofState φ) := by
ofState φ = crPart φ + anPart φ := by
rw [ofState]
cases φ with
| inAsymp φ => simp [statesToCrAnType]
Expand Down
72 changes: 36 additions & 36 deletions HepLean/PerturbationTheory/Algebras/CrAnAlgebra/NormalOrder.lean
Original file line number Diff line number Diff line change
Expand Up @@ -93,27 +93,27 @@ lemma normalOrder_mul_annihilate (φ : 𝓕.CrAnStates)
normalOrder_ofCrAnList_append_annihilate φ hφ]

lemma normalOrder_crPart_mul (φ : 𝓕.States) (a : CrAnAlgebra 𝓕) :
𝓝(crPart (StateAlgebra.ofState φ) * a) =
crPart (StateAlgebra.ofState φ) * 𝓝(a) := by
𝓝(crPart φ * a) =
crPart φ * 𝓝(a) := by
match φ with
| .inAsymp φ =>
rw [crPart, StateAlgebra.ofState, FreeAlgebra.lift_ι_apply]
rw [crPart]
exact normalOrder_create_mul ⟨States.inAsymp φ, ()⟩ rfl a
| .position φ =>
rw [crPart, StateAlgebra.ofState, FreeAlgebra.lift_ι_apply]
rw [crPart]
exact normalOrder_create_mul _ rfl _
| .outAsymp φ => simp

lemma normalOrder_mul_anPart (φ : 𝓕.States) (a : CrAnAlgebra 𝓕) :
𝓝(a * anPart (StateAlgebra.ofState φ)) =
𝓝(a) * anPart (StateAlgebra.ofState φ) := by
𝓝(a * anPart φ) =
𝓝(a) * anPart φ := by
match φ with
| .inAsymp φ => simp
| .position φ =>
rw [anPart, StateAlgebra.ofState, FreeAlgebra.lift_ι_apply]
rw [anPart]
exact normalOrder_mul_annihilate _ rfl _
| .outAsymp φ =>
rw [anPart, StateAlgebra.ofState, FreeAlgebra.lift_ι_apply]
rw [anPart]
refine normalOrder_mul_annihilate _ rfl _

/-!
Expand Down Expand Up @@ -182,9 +182,9 @@ lemma normalOrder_superCommute_annihilate_create (φc φa : 𝓕.CrAnStates)
exact Or.inr (normalOrder_superCommute_create_annihilate φc φa hφc hφa ..)

lemma normalOrder_swap_crPart_anPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
𝓝(a * (crPart (StateAlgebra.ofState φ)) * (anPart (StateAlgebra.ofState φ')) * b) =
𝓝(a * (crPart φ) * (anPart φ') * b) =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
𝓝(a * (anPart (StateAlgebra.ofState φ')) * (crPart (StateAlgebra.ofState φ)) * b) := by
𝓝(a * (anPart φ') * (crPart φ) * b) := by
match φ, φ' with
| _, .inAsymp φ' => simp
| .outAsymp φ, _ => simp
Expand Down Expand Up @@ -218,14 +218,14 @@ Using the results from above.
-/

lemma normalOrder_swap_anPart_crPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
𝓝(a * (anPart (StateAlgebra.ofState φ)) * (crPart (StateAlgebra.ofState φ')) * b) =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • 𝓝(a * (crPart (StateAlgebra.ofState φ')) *
(anPart (StateAlgebra.ofState φ)) * b) := by
𝓝(a * (anPart φ) * (crPart φ') * b) =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') • 𝓝(a * (crPart φ') *
(anPart φ) * b) := by
simp [normalOrder_swap_crPart_anPart, smul_smul]

lemma normalOrder_superCommute_crPart_anPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
𝓝(a * superCommute
(crPart (StateAlgebra.ofState φ)) (anPart (StateAlgebra.ofState φ')) * b) = 0 := by
(crPart φ) (anPart φ') * b) = 0 := by
match φ, φ' with
| _, .inAsymp φ' => simp
| .outAsymp φ', _ => simp
Expand All @@ -244,7 +244,7 @@ lemma normalOrder_superCommute_crPart_anPart (φ φ' : 𝓕.States) (a b : CrAnA

lemma normalOrder_superCommute_anPart_crPart (φ φ' : 𝓕.States) (a b : CrAnAlgebra 𝓕) :
𝓝(a * superCommute
(anPart (StateAlgebra.ofState φ)) (crPart (StateAlgebra.ofState φ')) * b) = 0 := by
(anPart φ) (crPart φ') * b) = 0 := by
match φ, φ' with
| .inAsymp φ', _ => simp
| _, .outAsymp φ' => simp
Expand All @@ -269,49 +269,49 @@ lemma normalOrder_superCommute_anPart_crPart (φ φ' : 𝓕.States) (a b : CrAnA

@[simp]
lemma normalOrder_crPart_mul_crPart (φ φ' : 𝓕.States) :
𝓝(crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ')) =
crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') := by
𝓝(crPart φ * crPart φ') =
crPart φ * crPart φ' := by
rw [normalOrder_crPart_mul]
conv_lhs => rw [← mul_one (crPart (StateAlgebra.ofState φ'))]
conv_lhs => rw [← mul_one (crPart φ')]
rw [normalOrder_crPart_mul, normalOrder_one]
simp

@[simp]
lemma normalOrder_anPart_mul_anPart (φ φ' : 𝓕.States) :
𝓝(anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ')) =
anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') := by
𝓝(anPart φ * anPart φ') =
anPart φ * anPart φ' := by
rw [normalOrder_mul_anPart]
conv_lhs => rw [← one_mul (anPart (StateAlgebra.ofState φ))]
conv_lhs => rw [← one_mul (anPart φ)]
rw [normalOrder_mul_anPart, normalOrder_one]
simp

@[simp]
lemma normalOrder_crPart_mul_anPart (φ φ' : 𝓕.States) :
𝓝(crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ')) =
crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') := by
𝓝(crPart φ * anPart φ') =
crPart φ * anPart φ' := by
rw [normalOrder_crPart_mul]
conv_lhs => rw [← one_mul (anPart (StateAlgebra.ofState φ'))]
conv_lhs => rw [← one_mul (anPart φ')]
rw [normalOrder_mul_anPart, normalOrder_one]
simp

@[simp]
lemma normalOrder_anPart_mul_crPart (φ φ' : 𝓕.States) :
𝓝(anPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ')) =
𝓝(anPart φ * crPart φ') =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
(crPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ)) := by
conv_lhs => rw [← one_mul (anPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ'))]
conv_lhs => rw [← mul_one (1 * (anPart (StateAlgebra.ofState φ) *
crPart (StateAlgebra.ofState φ')))]
(crPart φ' * anPart φ) := by
conv_lhs => rw [← one_mul (anPart φ * crPart φ')]
conv_lhs => rw [← mul_one (1 * (anPart φ *
crPart φ'))]
rw [← mul_assoc, normalOrder_swap_anPart_crPart]
simp

lemma normalOrder_ofState_mul_ofState (φ φ' : 𝓕.States) :
𝓝(ofState φ * ofState φ') =
crPart (StateAlgebra.ofState φ) * crPart (StateAlgebra.ofState φ') +
crPart φ * crPart φ' +
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φ') •
(crPart (StateAlgebra.ofState φ') * anPart (StateAlgebra.ofState φ)) +
crPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') +
anPart (StateAlgebra.ofState φ) * anPart (StateAlgebra.ofState φ') := by
(crPart φ' * anPart φ) +
crPart φ * anPart φ' +
anPart φ * anPart φ' := by
rw [ofState_eq_crPart_add_anPart, ofState_eq_crPart_add_anPart, mul_add, add_mul, add_mul]
simp only [map_add, normalOrder_crPart_mul_crPart, normalOrder_anPart_mul_crPart,
instCommGroup.eq_1, normalOrder_crPart_mul_anPart, normalOrder_anPart_mul_anPart]
Expand Down Expand Up @@ -497,9 +497,9 @@ lemma ofCrAnState_mul_normalOrder_ofStateList_eq_superCommute (φ : 𝓕.CrAnSta

lemma anPart_mul_normalOrder_ofStateList_eq_superCommute (φ : 𝓕.States)
(φs' : List 𝓕.States) :
anPart (StateAlgebra.ofState φ) * 𝓝(ofStateList φs') =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • 𝓝(ofStateList φs' * anPart (StateAlgebra.ofState φ))
+ [anPart (StateAlgebra.ofState φ), 𝓝(ofStateList φs')]ₛca := by
anPart φ * 𝓝(ofStateList φs') =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs') • 𝓝(ofStateList φs' * anPart φ)
+ [anPart φ, 𝓝(ofStateList φs')]ₛca := by
rw [normalOrder_mul_anPart]
match φ with
| .inAsymp φ => simp
Expand Down
Loading

0 comments on commit acfd0ed

Please sign in to comment.