2025-02-03 11:28:14 +00:00
|
|
|
|
/-
|
|
|
|
|
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.Basic
|
|
|
|
|
import HepLean.PerturbationTheory.CreateAnnihilate
|
|
|
|
|
/-!
|
|
|
|
|
|
|
|
|
|
# Creation and annihilation states
|
|
|
|
|
|
|
|
|
|
Called `CrAnFieldOp` for short here.
|
|
|
|
|
|
|
|
|
|
Given a field specification, in addition to defining states
|
|
|
|
|
(see: `HepLean.PerturbationTheory.FieldSpecification.Basic`),
|
|
|
|
|
we can also define creation and annihilation states.
|
|
|
|
|
These are similar to states but come with an additional specification of whether they correspond to
|
|
|
|
|
creation or annihilation operators.
|
|
|
|
|
|
|
|
|
|
In particular we have the following creation and annihilation states for each field:
|
|
|
|
|
- Negative asymptotic states - with the implicit specification that it is a creation state.
|
|
|
|
|
- Position states with a creation specification.
|
|
|
|
|
- Position states with an annihilation specification.
|
|
|
|
|
- Positive asymptotic states - with the implicit specification that it is an annihilation state.
|
|
|
|
|
|
|
|
|
|
In this module in addition to defining `CrAnFieldOp` we also define some maps:
|
|
|
|
|
- The map `crAnFieldOpToFieldOp` takes a `CrAnFieldOp` to its state in `FieldOp`.
|
|
|
|
|
- The map `crAnFieldOpToCreateAnnihilate` takes a `CrAnFieldOp` to its corresponding
|
|
|
|
|
`CreateAnnihilate` value.
|
|
|
|
|
- The map `crAnStatistics` takes a `CrAnFieldOp` to its corresponding `FieldStatistic`
|
|
|
|
|
(bosonic or fermionic).
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
namespace FieldSpecification
|
|
|
|
|
variable (𝓕 : FieldSpecification)
|
|
|
|
|
|
2025-02-08 13:07:54 +00:00
|
|
|
|
/-- To each field operator the specification of the type of creation and annihilation parts.
|
2025-02-07 15:43:59 +00:00
|
|
|
|
For asymptotic states there is only one allowed part, whilst for position
|
2025-02-03 15:59:25 +00:00
|
|
|
|
field operator there is two. -/
|
2025-02-03 11:28:14 +00:00
|
|
|
|
def fieldOpToCrAnType : 𝓕.FieldOp → Type
|
|
|
|
|
| FieldOp.inAsymp _ => Unit
|
|
|
|
|
| FieldOp.position _ => CreateAnnihilate
|
|
|
|
|
| FieldOp.outAsymp _ => Unit
|
|
|
|
|
|
|
|
|
|
/-- The instance of a finite type on `𝓕.fieldOpToCreateAnnihilateType i`. -/
|
|
|
|
|
instance : ∀ i, Fintype (𝓕.fieldOpToCrAnType i) := fun i =>
|
|
|
|
|
match i with
|
|
|
|
|
| FieldOp.inAsymp _ => inferInstanceAs (Fintype Unit)
|
|
|
|
|
| FieldOp.position _ => inferInstanceAs (Fintype CreateAnnihilate)
|
|
|
|
|
| FieldOp.outAsymp _ => inferInstanceAs (Fintype Unit)
|
|
|
|
|
|
|
|
|
|
/-- The instance of a decidable equality on `𝓕.fieldOpToCreateAnnihilateType i`. -/
|
|
|
|
|
instance : ∀ i, DecidableEq (𝓕.fieldOpToCrAnType i) := fun i =>
|
|
|
|
|
match i with
|
|
|
|
|
| FieldOp.inAsymp _ => inferInstanceAs (DecidableEq Unit)
|
|
|
|
|
| FieldOp.position _ => inferInstanceAs (DecidableEq CreateAnnihilate)
|
|
|
|
|
| FieldOp.outAsymp _ => inferInstanceAs (DecidableEq Unit)
|
|
|
|
|
|
|
|
|
|
/-- The equivalence between `𝓕.fieldOpToCreateAnnihilateType i` and
|
|
|
|
|
`𝓕.fieldOpToCreateAnnihilateType j` from an equality `i = j`. -/
|
|
|
|
|
def fieldOpToCreateAnnihilateTypeCongr : {i j : 𝓕.FieldOp} → i = j →
|
|
|
|
|
𝓕.fieldOpToCrAnType i ≃ 𝓕.fieldOpToCrAnType j
|
|
|
|
|
| _, _, rfl => Equiv.refl _
|
|
|
|
|
|
2025-02-03 15:59:25 +00:00
|
|
|
|
/--
|
2025-02-10 10:21:57 +00:00
|
|
|
|
For a field specification `𝓕`, the (sigma) type `𝓕.CrAnFieldOp`
|
|
|
|
|
corresponds to the type of creation and annihilation parts of field operators.
|
|
|
|
|
It formally defined to consist of the following elements:
|
|
|
|
|
- for each in incoming asymptotic field operator `φ` in `𝓕.FieldOp` an element
|
|
|
|
|
written as `⟨φ, ()⟩` in `𝓕.CrAnFieldOp`, corresponding to the creation part of `φ`.
|
|
|
|
|
Here `φ` has no annihilation part. (Here `()` is the unique element of `Unit`.)
|
|
|
|
|
- for each position field operator `φ` in `𝓕.FieldOp` an element of `𝓕.CrAnFieldOp`
|
|
|
|
|
written as `⟨φ, .create⟩`, corresponding to the creation part of `φ`.
|
|
|
|
|
- for each position field operator `φ` in `𝓕.FieldOp` an element of `𝓕.CrAnFieldOp`
|
|
|
|
|
written as `⟨φ, .annihilate⟩`, corresponding to the annihilation part of `φ`.
|
|
|
|
|
- for each out outgoing asymptotic field operator `φ` in `𝓕.FieldOp` an element
|
|
|
|
|
written as `⟨φ, ()⟩` in `𝓕.CrAnFieldOp`, corresponding to the annihilation part of `φ`.
|
|
|
|
|
Here `φ` has no creation part. (Here `()` is the unique element of `Unit`.)
|
|
|
|
|
|
|
|
|
|
As some intuition, if `f` corresponds to a Weyl-fermion field, it would contribute
|
|
|
|
|
the following elements to `𝓕.CrAnFieldOp`
|
|
|
|
|
- an element corresponding to incoming asymptotic operators for each spin `s`: `a(p, s)`.
|
|
|
|
|
- an element corresponding to the creation parts of position operators for each each Lorentz
|
|
|
|
|
index `α`:
|
|
|
|
|
`∑ s, ∫ d^3p/(…) (x_α(p,s) a(p, s) e^{-i p x})`.
|
2025-02-10 10:51:44 +00:00
|
|
|
|
- an element corresponding to annihilation parts of position operator,
|
2025-02-10 10:21:57 +00:00
|
|
|
|
for each each Lorentz index `α`:
|
|
|
|
|
`∑ s, ∫ d^3p/(…) (y_α(p,s) a^†(p, s) e^{-i p x})`.
|
|
|
|
|
- an element corresponding to outgoing asymptotic operators for each spin `s`: `a^†(p, s)`.
|
|
|
|
|
|
2025-02-03 15:59:25 +00:00
|
|
|
|
-/
|
2025-02-03 11:28:14 +00:00
|
|
|
|
def CrAnFieldOp : Type := Σ (s : 𝓕.FieldOp), 𝓕.fieldOpToCrAnType s
|
|
|
|
|
|
2025-02-07 15:43:59 +00:00
|
|
|
|
/-- The map from creation and annihilation field operator to their underlying states. -/
|
2025-02-03 11:28:14 +00:00
|
|
|
|
def crAnFieldOpToFieldOp : 𝓕.CrAnFieldOp → 𝓕.FieldOp := Sigma.fst
|
|
|
|
|
|
|
|
|
|
@[simp]
|
|
|
|
|
lemma crAnFieldOpToFieldOp_prod (s : 𝓕.FieldOp) (t : 𝓕.fieldOpToCrAnType s) :
|
|
|
|
|
𝓕.crAnFieldOpToFieldOp ⟨s, t⟩ = s := rfl
|
|
|
|
|
|
2025-02-10 10:51:44 +00:00
|
|
|
|
/-- For a field specification `𝓕`, `𝓕.crAnFieldOpToCreateAnnihilate` is the map from
|
2025-02-10 10:21:57 +00:00
|
|
|
|
`𝓕.CrAnFieldOp` to `CreateAnnihilate` taking `φ` to `create` if
|
|
|
|
|
- `φ` corresponds to an incoming asymptotic field operator or the creation part of a position based
|
|
|
|
|
field operator.
|
|
|
|
|
|
|
|
|
|
otherwise it takes `φ` to `annihilate`.
|
|
|
|
|
-/
|
2025-02-03 11:28:14 +00:00
|
|
|
|
def crAnFieldOpToCreateAnnihilate : 𝓕.CrAnFieldOp → CreateAnnihilate
|
|
|
|
|
| ⟨FieldOp.inAsymp _, _⟩ => CreateAnnihilate.create
|
|
|
|
|
| ⟨FieldOp.position _, CreateAnnihilate.create⟩ => CreateAnnihilate.create
|
|
|
|
|
| ⟨FieldOp.position _, CreateAnnihilate.annihilate⟩ => CreateAnnihilate.annihilate
|
|
|
|
|
| ⟨FieldOp.outAsymp _, _⟩ => CreateAnnihilate.annihilate
|
|
|
|
|
|
2025-02-06 10:06:05 +00:00
|
|
|
|
/-- For a field specification `𝓕`, and an element `φ` in `𝓕.CrAnFieldOp`, the field
|
|
|
|
|
statistic `crAnStatistics φ` is defined to be the statistic associated with the field `𝓕.Field`
|
2025-02-10 10:21:57 +00:00
|
|
|
|
(or the `𝓕.FieldOp`) underlying `φ`.
|
2025-02-06 10:06:05 +00:00
|
|
|
|
|
|
|
|
|
The following notation is used in relation to `crAnStatistics`:
|
2025-02-06 13:31:59 +00:00
|
|
|
|
- For `φ` an element of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φ` is `crAnStatistics φ`.
|
|
|
|
|
- For `φs` a list of `𝓕.CrAnFieldOp`, `𝓕 |>ₛ φs` is the product of `crAnStatistics φ` over
|
2025-02-06 10:06:05 +00:00
|
|
|
|
the list `φs`.
|
|
|
|
|
-/
|
2025-02-03 11:28:14 +00:00
|
|
|
|
def crAnStatistics : 𝓕.CrAnFieldOp → FieldStatistic :=
|
2025-02-06 10:06:05 +00:00
|
|
|
|
𝓕.fieldOpStatistic ∘ 𝓕.crAnFieldOpToFieldOp
|
2025-02-03 11:28:14 +00:00
|
|
|
|
|
2025-02-06 10:06:05 +00:00
|
|
|
|
@[inherit_doc crAnStatistics]
|
2025-02-03 11:28:14 +00:00
|
|
|
|
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ =>
|
|
|
|
|
(crAnStatistics 𝓕) φ
|
|
|
|
|
|
2025-02-06 10:06:05 +00:00
|
|
|
|
@[inherit_doc crAnStatistics]
|
2025-02-03 11:28:14 +00:00
|
|
|
|
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
|
|
|
|
|
(crAnStatistics 𝓕) φ
|
|
|
|
|
|
2025-02-03 11:29:37 +00:00
|
|
|
|
/-- The `CreateAnnihilate` value of a `CrAnFieldOp`s, i.e. whether it is a creation or
|
2025-02-03 11:28:14 +00:00
|
|
|
|
annihilation operator. -/
|
|
|
|
|
scoped[FieldSpecification] infixl:80 "|>ᶜ" =>
|
|
|
|
|
crAnFieldOpToCreateAnnihilate
|
|
|
|
|
|
2025-02-10 10:21:57 +00:00
|
|
|
|
remark notation_remark := "When working with a field specification `𝓕` the
|
|
|
|
|
following notation will be used within doc-strings:
|
|
|
|
|
- when field statistics occur in exchange signs the `𝓕 |>ₛ _` may be dropped.
|
|
|
|
|
- lists of `FieldOp` or `CrAnFieldOp` `φs` may be written as `φ₀…φₙ`,
|
2025-02-06 10:06:05 +00:00
|
|
|
|
which should be interpreted within the context in which it appears.
|
2025-02-10 10:21:57 +00:00
|
|
|
|
- `φᶜ` may be used to indicate the creation part of an operator and
|
2025-02-06 10:47:34 +00:00
|
|
|
|
`φᵃ` to indicate the annihilation part of an operator.
|
2025-02-06 10:06:05 +00:00
|
|
|
|
|
2025-02-10 10:21:57 +00:00
|
|
|
|
Some examples of these notation-conventions are:
|
|
|
|
|
- `𝓢(φ, φs)` which corresponds to `𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs)`
|
|
|
|
|
- `φ₀…φᵢ₋₁φᵢ₊₁…φₙ` which corresponds to a (given) list `φs = φ₀…φₙ` with the element at the
|
|
|
|
|
`i`th position removed.
|
2025-02-06 10:06:05 +00:00
|
|
|
|
"
|
2025-02-06 05:27:33 +00:00
|
|
|
|
|
2025-02-03 11:28:14 +00:00
|
|
|
|
end FieldSpecification
|