PhysLean/HepLean/FeynmanDiagrams/Momentum.lean

211 lines
6.6 KiB
Text
Raw Normal View History

2024-06-18 15:34:57 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.FeynmanDiagrams.Basic
2024-06-19 07:45:19 -04:00
import Mathlib.Data.Real.Basic
import Mathlib.Algebra.DirectSum.Module
import Mathlib.LinearAlgebra.SesquilinearForm
2024-06-25 07:06:32 -04:00
import Mathlib.LinearAlgebra.Dimension.Finrank
2024-06-18 15:34:57 -04:00
/-!
# Momentum in Feynman diagrams
The aim of this file is to associate with each half-edge of a Feynman diagram a momentum,
and constrain the momentums based conservation at each vertex and edge.
2024-06-19 07:45:19 -04:00
The number of loops of a Feynman diagram is related to the dimension of the resulting
vector space.
## TODO
- Prove lemmas that make the calcuation of the number of loops of a Feynman diagram easier.
## Note
This section is non-computable as we depend on the norm on `F.HalfEdgeMomenta`.
2024-06-18 15:34:57 -04:00
-/
namespace FeynmanDiagram
open CategoryTheory
open PreFeynmanRule
variable {P : PreFeynmanRule} (F : FeynmanDiagram P) [IsFiniteDiagram F]
/-!
## Vector spaces associated with momenta in Feynman diagrams.
We define the vector space associated with momenta carried by half-edges,
outflowing momenta of edges, and inflowing momenta of vertices.
We define the direct sum of the edge and vertex momentum spaces.
-/
/-- The type which assocaites to each half-edge a `1`-dimensional vector space.
Corresponding to that spanned by its momentum. -/
def HalfEdgeMomenta : Type := F.𝓱𝓔
instance : AddCommGroup F.HalfEdgeMomenta := Pi.addCommGroup
instance : Module F.HalfEdgeMomenta := Pi.module _ _ _
/-- An auxillary function used to define the Euclidean inner product on `F.HalfEdgeMomenta`. -/
def euclidInnerAux (x : F.HalfEdgeMomenta) : F.HalfEdgeMomenta →ₗ[] where
toFun y := ∑ i, (x i) * (y i)
map_add' z y :=
show (∑ i, (x i) * (z i + y i)) = (∑ i, x i * z i) + ∑ i, x i * (y i) by
simp only [mul_add, Finset.sum_add_distrib]
map_smul' c y :=
show (∑ i, x i * (c * y i)) = c * ∑ i, x i * y i by
rw [Finset.mul_sum]
refine Finset.sum_congr rfl (fun _ _ => by ring)
lemma euclidInnerAux_symm (x y : F.HalfEdgeMomenta) :
F.euclidInnerAux x y = F.euclidInnerAux y x := Finset.sum_congr rfl (fun _ _ => by ring)
/-- The Euclidean inner product on `F.HalfEdgeMomenta`. -/
def euclidInner : F.HalfEdgeMomenta →ₗ[] F.HalfEdgeMomenta →ₗ[] where
toFun x := F.euclidInnerAux x
map_add' x y := by
refine LinearMap.ext (fun z => ?_)
simp only [euclidInnerAux_symm, map_add, LinearMap.add_apply]
map_smul' c x := by
refine LinearMap.ext (fun z => ?_)
simp only [euclidInnerAux_symm, LinearMapClass.map_smul, smul_eq_mul, RingHom.id_apply,
LinearMap.smul_apply]
2024-06-18 15:34:57 -04:00
/-- The type which assocaites to each ege a `1`-dimensional vector space.
Corresponding to that spanned by its total outflowing momentum. -/
def EdgeMomenta : Type := F.𝓔
2024-06-18 15:34:57 -04:00
instance : AddCommGroup F.EdgeMomenta := Pi.addCommGroup
2024-06-19 07:45:19 -04:00
instance : Module F.EdgeMomenta := Pi.module _ _ _
2024-06-18 15:34:57 -04:00
/-- The type which assocaites to each ege a `1`-dimensional vector space.
Corresponding to that spanned by its total inflowing momentum. -/
def VertexMomenta : Type := F.𝓥
2024-06-18 15:34:57 -04:00
instance : AddCommGroup F.VertexMomenta := Pi.addCommGroup
2024-06-18 15:34:57 -04:00
instance : Module F.VertexMomenta := Pi.module _ _ _
2024-06-18 15:34:57 -04:00
/-- The map from `Fin 2` to `Type` landing on `EdgeMomenta` and `VertexMomenta`. -/
def EdgeVertexMomentaMap : Fin 2 → Type := fun i =>
match i with
| 0 => F.EdgeMomenta
| 1 => F.VertexMomenta
2024-06-18 15:34:57 -04:00
instance (i : Fin 2) : AddCommGroup (EdgeVertexMomentaMap F i) :=
match i with
| 0 => instAddCommGroupEdgeMomenta F
| 1 => instAddCommGroupVertexMomenta F
2024-06-18 15:34:57 -04:00
instance (i : Fin 2) : Module (EdgeVertexMomentaMap F i) :=
match i with
| 0 => instModuleRealEdgeMomenta F
| 1 => instModuleRealVertexMomenta F
2024-06-18 15:34:57 -04:00
/-- The direct sum of `EdgeMomenta` and `VertexMomenta`.-/
def EdgeVertexMomenta : Type := DirectSum (Fin 2) (EdgeVertexMomentaMap F)
2024-06-18 15:34:57 -04:00
instance : AddCommGroup F.EdgeVertexMomenta := DirectSum.instAddCommGroup _
2024-06-18 15:34:57 -04:00
instance : Module F.EdgeVertexMomenta := DirectSum.instModule
2024-06-18 15:34:57 -04:00
/-!
## Linear maps between the vector spaces.
2024-06-19 07:45:19 -04:00
We define various maps into `F.HalfEdgeMomenta`.
2024-06-19 07:45:19 -04:00
In particular, we define a map from `F.EdgeVertexMomenta` to `F.HalfEdgeMomenta`. This
map represents the space orthogonal (with respect to the standard Euclidean inner product)
to the allowed momenta of half-edges (up-to an offset determined by the
external momenta).
The number of loops of a diagram is defined as the number of half-edges minus
the rank of this matrix.
2024-06-19 07:45:19 -04:00
2024-06-18 15:34:57 -04:00
-/
2024-06-19 07:45:19 -04:00
/-- The linear map from `F.EdgeMomenta` to `F.HalfEdgeMomenta` induced by
the map `F.𝓱𝓔To𝓔.hom`. -/
def edgeToHalfEdgeMomenta : F.EdgeMomenta →ₗ[] F.HalfEdgeMomenta where
toFun x := x ∘ F.𝓱𝓔To𝓔.hom
map_add' _ _ := by rfl
map_smul' _ _ := by rfl
/-- The linear map from `F.VertexMomenta` to `F.HalfEdgeMomenta` induced by
the map `F.𝓱𝓔To𝓥.hom`. -/
def vertexToHalfEdgeMomenta : F.VertexMomenta →ₗ[] F.HalfEdgeMomenta where
toFun x := x ∘ F.𝓱𝓔To𝓥.hom
map_add' _ _ := rfl
map_smul' _ _ := rfl
/-- The linear map from `F.EdgeVertexMomenta` to `F.HalfEdgeMomenta` induced by
`F.edgeToHalfEdgeMomenta` and `F.vertexToHalfEdgeMomenta`. -/
def edgeVertexToHalfEdgeMomenta : F.EdgeVertexMomenta →ₗ[] F.HalfEdgeMomenta :=
DirectSum.toModule (Fin 2) F.HalfEdgeMomenta
(fun i => match i with | 0 => F.edgeToHalfEdgeMomenta | 1 => F.vertexToHalfEdgeMomenta)
/-!
## Submodules
We define submodules of `F.HalfEdgeMomenta` which correspond to
the orthogonal space to allowed momenta (up-to an offset), and the space of
allowed momenta.
-/
/-- The submodule of `F.HalfEdgeMomenta` corresponding to the range of
`F.edgeVertexToHalfEdgeMomenta`. -/
def orthogHalfEdgeMomenta : Submodule F.HalfEdgeMomenta :=
LinearMap.range F.edgeVertexToHalfEdgeMomenta
/-- The submodule of `F.HalfEdgeMomenta` corresponding to the allowed momenta. -/
def allowedHalfEdgeMomenta : Submodule F.HalfEdgeMomenta :=
Submodule.orthogonalBilin F.orthogHalfEdgeMomenta F.euclidInner
/-!
## Number of loops
We define the number of loops of a Feynman diagram as the dimension of the
allowed space of half-edge momenta.
-/
/-- The number of loops of a Feynman diagram. Defined as the dimension
of the space of allowed Half-loop momenta. -/
noncomputable def numberOfLoops : := FiniteDimensional.finrank F.allowedHalfEdgeMomenta
/-!
## Lemmas regarding `numberOfLoops`
We now give a series of lemmas which be used to help calculate the number of loops
for specific Feynman diagrams.
### TODO
- Complete this section.
-/
/-!
## Category theory
### TODO
- Complete this section.
-/
2024-06-19 07:45:19 -04:00
2024-06-18 15:34:57 -04:00
end FeynmanDiagram