feat: Time dependent Wick theorem. (#274)
feat: Proof of the time-dependent Wick's theorem
This commit is contained in:
parent
4d43698b3c
commit
17f84b7153
53 changed files with 8563 additions and 3329 deletions
93
HepLean/PerturbationTheory/Algebras/StateAlgebra/Basic.lean
Normal file
93
HepLean/PerturbationTheory/Algebras/StateAlgebra/Basic.lean
Normal file
|
@ -0,0 +1,93 @@
|
|||
/-
|
||||
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.FieldStruct.CreateAnnihilate
|
||||
/-!
|
||||
|
||||
# State algebra
|
||||
|
||||
We define the state algebra of a field structure to be the free algebra
|
||||
generated by the states.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
|
||||
/-- The state free-algebra.
|
||||
The free algebra generated by `States`,
|
||||
that is a position based states or assymptotic states.
|
||||
As a module `StateAlgebra` is spanned by lists of `States`. -/
|
||||
abbrev StateAlgebra (𝓕 : FieldStruct) : Type := FreeAlgebra ℂ 𝓕.States
|
||||
|
||||
namespace StateAlgebra
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
/-- The element of the states free-algebra generated by a single state. -/
|
||||
def ofState (φ : 𝓕.States) : StateAlgebra 𝓕 :=
|
||||
FreeAlgebra.ι ℂ φ
|
||||
|
||||
/-- The element of the states free-algebra generated by a list of states. -/
|
||||
def ofList (φs : List 𝓕.States) : StateAlgebra 𝓕 :=
|
||||
(List.map ofState φs).prod
|
||||
|
||||
@[simp]
|
||||
lemma ofList_nil : ofList ([] : List 𝓕.States) = 1 := rfl
|
||||
|
||||
lemma ofList_singleton (φ : 𝓕.States) : ofList [φ] = ofState φ := by
|
||||
simp [ofList]
|
||||
|
||||
lemma ofList_append (φs ψs : List 𝓕.States) :
|
||||
ofList (φs ++ ψs) = ofList φs * ofList ψs := by
|
||||
rw [ofList, List.map_append, List.prod_append]
|
||||
rfl
|
||||
|
||||
lemma ofList_cons (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
ofList (φ :: φs) = ofState φ * ofList φs := rfl
|
||||
|
||||
/-- The basis of the free state algebra formed by lists of states. -/
|
||||
noncomputable def ofListBasis : Basis (List 𝓕.States) ℂ 𝓕.StateAlgebra where
|
||||
repr := FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
|
||||
|
||||
@[simp]
|
||||
lemma ofListBasis_eq_ofList (φs : List 𝓕.States) :
|
||||
ofListBasis φs = ofList φs := by
|
||||
simp only [ofListBasis, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
|
||||
Basis.coe_ofRepr, AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_apply,
|
||||
AlgEquiv.ofAlgHom_symm_apply, ofList]
|
||||
erw [MonoidAlgebra.lift_apply]
|
||||
simp only [zero_smul, Finsupp.sum_single_index, one_smul]
|
||||
rw [@FreeMonoid.lift_apply]
|
||||
simp only [List.prod]
|
||||
match φs with
|
||||
| [] => rfl
|
||||
| φ :: φs =>
|
||||
erw [List.map_cons]
|
||||
|
||||
/-!
|
||||
|
||||
## The super commutor on the state algebra.
|
||||
|
||||
-/
|
||||
|
||||
/-- The super commutor on the free state algebra. For two bosonic operators
|
||||
or a bosonic and fermionic operator this corresponds to the usual commutator
|
||||
whilst for two fermionic operators this corresponds to the anti-commutator. -/
|
||||
noncomputable def superCommute : 𝓕.StateAlgebra →ₗ[ℂ] 𝓕.StateAlgebra →ₗ[ℂ] 𝓕.StateAlgebra :=
|
||||
Basis.constr ofListBasis ℂ fun φs =>
|
||||
Basis.constr ofListBasis ℂ fun φs' =>
|
||||
ofList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofList (φs' ++ φs)
|
||||
|
||||
local notation "⟨" φs "," φs' "⟩ₛ" => superCommute φs φs'
|
||||
|
||||
lemma superCommute_ofList (φs φs' : List 𝓕.States) : ⟨ofList φs, ofList φs'⟩ₛ =
|
||||
ofList (φs ++ φs') - 𝓢(𝓕 |>ₛ φs, 𝓕 |>ₛ φs') • ofList (φs' ++ φs) := by
|
||||
rw [← ofListBasis_eq_ofList, ← ofListBasis_eq_ofList]
|
||||
simp only [superCommute, Basis.constr_basis]
|
||||
|
||||
end StateAlgebra
|
||||
|
||||
end FieldStruct
|
|
@ -0,0 +1,83 @@
|
|||
/-
|
||||
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.FieldStruct.TimeOrder
|
||||
import HepLean.PerturbationTheory.Koszul.KoszulSign
|
||||
/-!
|
||||
|
||||
# State algebra
|
||||
|
||||
We define the state algebra of a field structure to be the free algebra
|
||||
generated by the states.
|
||||
|
||||
-/
|
||||
|
||||
namespace FieldStruct
|
||||
variable {𝓕 : FieldStruct}
|
||||
noncomputable section
|
||||
|
||||
namespace StateAlgebra
|
||||
|
||||
open FieldStatistic
|
||||
|
||||
/-- The linear map on the free state algebra defined as the map taking
|
||||
a list of states to the time-ordered list of states multiplied by
|
||||
the sign corresponding to the number of fermionic-fermionic
|
||||
exchanges done in ordering. -/
|
||||
def timeOrder : StateAlgebra 𝓕 →ₗ[ℂ] StateAlgebra 𝓕 :=
|
||||
Basis.constr ofListBasis ℂ fun φs =>
|
||||
timeOrderSign φs • ofList (timeOrderList φs)
|
||||
|
||||
lemma timeOrder_ofList (φs : List 𝓕.States) :
|
||||
timeOrder (ofList φs) = timeOrderSign φs • ofList (timeOrderList φs) := by
|
||||
rw [← ofListBasis_eq_ofList]
|
||||
simp only [timeOrder, Basis.constr_basis]
|
||||
|
||||
lemma timeOrder_ofList_nil : timeOrder (𝓕 := 𝓕) (ofList []) = 1 := by
|
||||
rw [timeOrder_ofList]
|
||||
simp [timeOrderSign, Wick.koszulSign, timeOrderList]
|
||||
|
||||
@[simp]
|
||||
lemma timeOrder_ofList_singleton (φ : 𝓕.States) : timeOrder (ofList [φ]) = ofList [φ] := by
|
||||
simp [timeOrder_ofList, timeOrderSign, timeOrderList]
|
||||
|
||||
lemma timeOrder_ofState_ofState_ordered {φ ψ : 𝓕.States} (h : timeOrderRel φ ψ) :
|
||||
timeOrder (ofState φ * ofState ψ) = ofState φ * ofState ψ := by
|
||||
rw [← ofList_singleton, ← ofList_singleton, ← ofList_append, timeOrder_ofList]
|
||||
simp only [List.singleton_append]
|
||||
rw [timeOrderSign_pair_ordered h, timeOrderList_pair_ordered h]
|
||||
simp
|
||||
|
||||
lemma timeOrder_ofState_ofState_not_ordered {φ ψ : 𝓕.States} (h :¬ timeOrderRel φ ψ) :
|
||||
timeOrder (ofState φ * ofState ψ) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • ofState ψ * ofState φ := by
|
||||
rw [← ofList_singleton, ← ofList_singleton, ← ofList_append, timeOrder_ofList]
|
||||
simp only [List.singleton_append, instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
rw [timeOrderSign_pair_not_ordered h, timeOrderList_pair_not_ordered h]
|
||||
simp [← ofList_append]
|
||||
|
||||
lemma timeOrder_ofState_ofState_not_ordered_eq_timeOrder {φ ψ : 𝓕.States} (h :¬ timeOrderRel φ ψ) :
|
||||
timeOrder (ofState φ * ofState ψ) =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ψ) • timeOrder (ofState ψ * ofState φ) := by
|
||||
rw [timeOrder_ofState_ofState_not_ordered h]
|
||||
rw [timeOrder_ofState_ofState_ordered]
|
||||
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc]
|
||||
have hx := IsTotal.total (r := timeOrderRel) ψ φ
|
||||
simp_all
|
||||
|
||||
lemma timeOrder_eq_maxTimeField_mul (φ : 𝓕.States) (φs : List 𝓕.States) :
|
||||
timeOrder (ofList (φ :: φs)) =
|
||||
𝓢(𝓕 |>ₛ maxTimeField φ φs, 𝓕 |>ₛ (φ :: φs).take (maxTimeFieldPos φ φs)) •
|
||||
ofState (maxTimeField φ φs) * timeOrder (ofList (eraseMaxTimeField φ φs)) := by
|
||||
rw [timeOrder_ofList, timeOrderList_eq_maxTimeField_timeOrderList]
|
||||
rw [ofList_cons, timeOrder_ofList]
|
||||
simp only [instCommGroup.eq_1, Algebra.mul_smul_comm, Algebra.smul_mul_assoc, smul_smul]
|
||||
congr
|
||||
rw [timerOrderSign_of_eraseMaxTimeField, mul_assoc]
|
||||
simp
|
||||
|
||||
end StateAlgebra
|
||||
end
|
||||
end FieldStruct
|
Loading…
Add table
Add a link
Reference in a new issue