PhysLean/HepLean/PerturbationTheory/FieldSpecification/Basic.lean
Joseph Tooby-Smith b5c987180a
refactor: Some properties of field specifications (#285)
* refactor: Fix field struct defn.

* rename: FieldStruct to FieldSpecification

* feat: Add examples of field specifications

* docs: Slight improvement of module docs
2025-01-21 06:11:47 +00:00

73 lines
2.6 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
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
/-!
# 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.
From each field we can create three different types of `States`.
- Negative asymptotic states.
- Position states.
- Positive asymptotic states.
These states carry the same field statistic as the field they are derived from.
-/
/-- A field specification is a type of fields plus a specification of the
statistics (fermionic or bosonic) of each field. -/
structure FieldSpecification where
/-- The type of fields. This also includes anti-states. -/
Fields : Type
/-- The specification if a field is bosonic or fermionic. -/
statistics : Fields → FieldStatistic
namespace FieldSpecification
variable (𝓕 : FieldSpecification)
/-- Negative asymptotic states are specified by a field and a momentum. -/
def AsymptoticNegTime : Type := 𝓕.Fields × Lorentz.Contr 4
/-- Positive asymptotic states are specified by a field and a momentum. -/
def AsymptoticPosTime : Type := 𝓕.Fields × Lorentz.Contr 4
/-- States specified by a field and a space-time position. -/
def PositionStates : Type := 𝓕.Fields × SpaceTime
/-- The combination of asymptotic states and position states. -/
inductive States (𝓕 : FieldSpecification) where
| negAsymp : 𝓕.AsymptoticNegTime → 𝓕.States
| position : 𝓕.PositionStates → 𝓕.States
| posAsymp : 𝓕.AsymptoticPosTime → 𝓕.States
/-- Taking a state to its underlying field. -/
def statesToField : 𝓕.States → 𝓕.Fields
| States.negAsymp φ => φ.1
| States.position φ => φ.1
| States.posAsymp φ => φ.1
/-- The statistics associated to a state. -/
def statesStatistic : 𝓕.States → FieldStatistic := 𝓕.statistics ∘ 𝓕.statesToField
/-- The field statistics associated with a state. -/
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => statesStatistic 𝓕 φ
/-- The field statistics associated with a list states. -/
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
(statesStatistic 𝓕) φ
/-- The field statistic associated with a finite set-/
scoped[FieldSpecification] notation 𝓕 "|>ₛ" "⟨" f ","a "⟩"=> FieldStatistic.ofFinset
(statesStatistic 𝓕) f a
end FieldSpecification