PhysLean/HepLean/PerturbationTheory/FieldSpecification/Basic.lean

140 lines
6.8 KiB
Text
Raw Normal View History

/-
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.Lorentz.RealVector.Basic
import HepLean.PerturbationTheory.FieldStatistics.ExchangeSign
import HepLean.SpaceTime.Basic
import HepLean.PerturbationTheory.FieldStatistics.OfFinset
2025-01-23 06:31:11 +00:00
import HepLean.Meta.Remark.Basic
/-!
# Field specification
In this module is the definition of a field specification.
A field specification is a structure consisting of a type of fields and a
the field statistics of each field.
2025-02-03 11:28:14 +00:00
From each field we can create three different types of `FieldOp`.
- Negative asymptotic states.
- Position states.
- Positive asymptotic states.
These states carry the same field statistic as the field they are derived from.
## Some references
- https://particle.physics.ucdavis.edu/modernsusy/slides/slideimages/spinorfeynrules.pdf
2025-01-23 06:31:11 +00:00
-/
2025-01-23 06:31:11 +00:00
2025-02-06 10:06:05 +00:00
/--
The structure `FieldSpecification` is defined to have the following content:
- A type `Field` whose elements are the constituent fields of the theory.
- For every field `f` in `Field`, a type `PositionLabel f` whose elements label the different
position operators associated with the field `f`. For example,
- For `f` a *real-scalar field*, `PositionLabel f` will have a unique element.
- For `f` a *complex-scalar field*, `PositionLabel f` will have two elements, one for the field
operator and one for its conjugate.
- For `f` a *Dirac fermion*, `PositionLabel f` will have eight elements, one for each Lorentz
index of the field and its conjugate.
- For `f` a *Weyl fermion*, `PositionLabel f` will have four elements, one for each Lorentz
index of the field and its conjugate.
- For every field `f` in `Field`, a type `AsymptoticLabel f` whose elements label the different
2025-02-10 10:21:57 +00:00
types of incoming asymptotic field operators associated with the
field `f` (this is also matches the types of outgoing asymptotic field operators).
For example,
2025-02-06 10:06:05 +00:00
- For `f` a *real-scalar field*, `AsymptoticLabel f` will have a unique element.
- For `f` a *complex-scalar field*, `AsymptoticLabel f` will have two elements, one for the
field operator and one for its conjugate.
- For `f` a *Dirac fermion*, `AsymptoticLabel f` will have four elements, two for each spin.
- For `f` a *Weyl fermion*, `AsymptoticLabel f` will have two elements, one for each spin.
2025-02-10 10:21:57 +00:00
- For each field `f` in `Field`, a field statistic `statistic f` which classifies `f` as either
2025-02-06 10:06:05 +00:00
`bosonic` or `fermionic`.
-/
structure FieldSpecification where
2025-02-06 10:06:05 +00:00
/-- A type whose elements are the constituent fields of the theory. -/
Field : Type
/-- For every field `f` in `Field`, the type `PositionLabel f` has elements that label the
different position operators associated with the field `f`. -/
PositionLabel : Field → Type
/-- For every field `f` in `Field`, the type `AsymptoticLabel f` has elements that label
the different asymptotic based field operators associated with the field `f`. -/
AsymptoticLabel : Field → Type
/-- For every field `f` in `Field`, the field statistic `statistic f` classifies `f` as either
`bosonic` or `fermionic`. -/
statistic : Field → FieldStatistic
namespace FieldSpecification
variable (𝓕 : FieldSpecification)
2025-02-10 10:21:57 +00:00
/-- For a field specification `𝓕`, the inductive type `𝓕.FieldOp` is defined
to contain the following elements:
- for every `f` in `𝓕.Field`, element of `e` of `AsymptoticLabel f` and `3`-momentum `p`, an
element labelled `inAsymp f e p` corresponding to an incoming asymptotic field operator
of the field `f`, of label `e` (e.g. specifying the spin), and momentum `p`.
- for every `f` in `𝓕.Field`, element of `e` of `PositionLabel f` and space-time position `x`, an
element labelled `position f e x` corresponding to a position field operator of the field `f`,
of label `e` (e.g. specifying the Lorentz index), and position `x`.
- for every `f` in `𝓕.Field`, element of `e` of `AsymptoticLabel f` and `3`-momentum `p`, an
element labelled `outAsymp f e p` corresponding to an outgoing asymptotic field operator of the
field `f`, of label `e` (e.g. specifying the spin), and momentum `p`.
2025-02-06 10:06:05 +00:00
2025-02-10 10:21:57 +00:00
As some intuition, if `f` corresponds to a Weyl-fermion field, then
- For `inAsymp f e p`, `e` would correspond to a spin `s`, and `inAsymp f e p` would, once
represented in the operator algebra,
be proportional to the creation operator `a(p, s)`.
- `position f e x`, `e` would correspond to a Lorentz index `a`, and `position f e x` would,
2025-02-10 10:21:57 +00:00
once represented in the operator algebra, be proportional to the operator
`∑ s, ∫ d³p/(…) (xₐ(p,s) a(p, s) e ^ (-i p x) + yₐ(p,s) a†(p, s) e ^ (-i p x))`.
2025-02-10 10:51:44 +00:00
- `outAsymp f e p`, `e` would correspond to a spin `s`, and `outAsymp f e p` would,
2025-02-10 10:21:57 +00:00
once represented in the operator algebra, be proportional to the
annihilation operator `a†(p, s)`.
2025-02-10 10:21:57 +00:00
This type contains all operators which are related to a field.
2025-02-03 15:59:25 +00:00
-/
2025-02-03 11:28:14 +00:00
inductive FieldOp (𝓕 : FieldSpecification) where
2025-02-06 10:06:05 +00:00
| inAsymp : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → ) → 𝓕.FieldOp
| position : (Σ f, 𝓕.PositionLabel f) × SpaceTime → 𝓕.FieldOp
| outAsymp : (Σ f, 𝓕.AsymptoticLabel f) × (Fin 3 → ) → 𝓕.FieldOp
2025-02-03 15:59:25 +00:00
/-- The bool on `FieldOp` which is true only for position field operator. -/
2025-02-03 11:28:14 +00:00
def statesIsPosition : 𝓕.FieldOp → Bool
| FieldOp.position _ => true
| _ => false
2025-01-27 11:26:02 +00:00
2025-02-06 10:06:05 +00:00
/-- For a field specification `𝓕`, `𝓕.fieldOpToField` is defined to take field operators
to their underlying field. -/
def fieldOpToField : 𝓕.FieldOp → 𝓕.Field
| FieldOp.inAsymp (f, _) => f.1
| FieldOp.position (f, _) => f.1
| FieldOp.outAsymp (f, _) => f.1
/-- For a field specification `𝓕`, and an element `φ` of `𝓕.FieldOp`.
The field statistic `fieldOpStatistic φ` is defined to be the statistic associated with
the field underlying `φ`.
The following notation is used in relation to `fieldOpStatistic`:
2025-02-06 13:31:59 +00:00
- For `φ` an element of `𝓕.FieldOp`, `𝓕 |>ₛ φ` is `fieldOpStatistic φ`.
- For `φs` a list of `𝓕.FieldOp`, `𝓕 |>ₛ φs` is the product of `fieldOpStatistic φ` over
2025-02-06 10:06:05 +00:00
the list `φs`.
2025-02-06 13:31:59 +00:00
- For a function `f : Fin n → 𝓕.FieldOp` and a finset `a` of `Fin n`, `𝓕 |>ₛ ⟨f, a⟩` is the
2025-02-06 10:06:05 +00:00
product of `fieldOpStatistic (f i)` for all `i ∈ a`. -/
2025-02-06 13:31:59 +00:00
def fieldOpStatistic : 𝓕.FieldOp → FieldStatistic := 𝓕.statistic ∘ 𝓕.fieldOpToField
2025-02-06 10:06:05 +00:00
@[inherit_doc fieldOpStatistic]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => fieldOpStatistic 𝓕 φ
@[inherit_doc fieldOpStatistic]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
2025-02-06 10:06:05 +00:00
(fieldOpStatistic 𝓕) φ
2025-02-06 10:06:05 +00:00
@[inherit_doc fieldOpStatistic]
scoped[FieldSpecification] notation 𝓕 "|>ₛ" "⟨" f ","a "⟩"=> FieldStatistic.ofFinset
2025-02-06 10:06:05 +00:00
(fieldOpStatistic 𝓕) f a
end FieldSpecification