feat: Add definition of Number of Loops

This commit is contained in:
jstoobysmith 2024-06-19 13:02:09 -04:00
parent 277436c347
commit 65cbef4bd2
3 changed files with 207 additions and 53 deletions

View file

@ -348,6 +348,12 @@ def 𝓱𝓔 : Type := F.𝓱𝓔To𝓔𝓥.left
/-- The object in Over P.HalfEdgeLabel generated by a Feynman diagram. -/
def 𝓱𝓔𝓞 : Over P.HalfEdgeLabel := P.toHalfEdge.obj F.𝓱𝓔To𝓔𝓥
/-- The map `F.𝓱𝓔 → F.𝓔` as an object in `Over F.𝓔 `. -/
def 𝓱𝓔To𝓔 : Over F.𝓔 := P.toEdge.obj F.𝓱𝓔To𝓔𝓥
/-- The map `F.𝓱𝓔 → F.𝓥` as an object in `Over F.𝓥 `. -/
def 𝓱𝓔To𝓥 : Over F.𝓥 := P.toVertex.obj F.𝓱𝓔To𝓔𝓥
/-!
## Making a Feynman diagram
@ -414,6 +420,13 @@ class IsFiniteDiagram (F : FeynmanDiagram P) where
/-- The type of half-edges is decidable. -/
𝓱𝓔DecidableEq : DecidableEq F.𝓱𝓔
instance {𝓔 𝓥 𝓱𝓔 : Type} [h1 : Fintype 𝓔] [h2 : DecidableEq 𝓔]
[h3 : Fintype 𝓥] [h4 : DecidableEq 𝓥] [h5 : Fintype 𝓱𝓔] [h6 : DecidableEq 𝓱𝓔]
(𝓔𝓞 : 𝓔 → P.EdgeLabel) (𝓥𝓞 : 𝓥 → P.VertexLabel)
(𝓱𝓔To𝓔𝓥 : 𝓱𝓔 → P.HalfEdgeLabel × 𝓔 × 𝓥) (C : Cond 𝓔𝓞 𝓥𝓞 𝓱𝓔To𝓔𝓥) :
IsFiniteDiagram (mk' 𝓔𝓞 𝓥𝓞 𝓱𝓔To𝓔𝓥 C) :=
⟨h1, h2, h3, h4, h5, h6⟩
instance {F : FeynmanDiagram P} [IsFiniteDiagram F] : Fintype F.𝓔 :=
IsFiniteDiagram.𝓔Fintype
@ -432,6 +445,9 @@ instance {F : FeynmanDiagram P} [IsFiniteDiagram F] : Fintype F.𝓱𝓔 :=
instance {F : FeynmanDiagram P} [IsFiniteDiagram F] : DecidableEq F.𝓱𝓔 :=
IsFiniteDiagram.𝓱𝓔DecidableEq
instance {F : FeynmanDiagram P} [IsFiniteDiagram F] (i : F.𝓱𝓔) (j : F.𝓔) :
Decidable (F.𝓱𝓔To𝓔.hom i = j) := IsFiniteDiagram.𝓔DecidableEq (F.𝓱𝓔To𝓔.hom i) j
instance fintypeProdHalfEdgeLabel𝓔𝓥 {F : FeynmanDiagram P} [IsFinitePreFeynmanRule P]
[IsFiniteDiagram F] : DecidableEq (P.HalfEdgeLabel × F.𝓔 × F.𝓥) :=
instDecidableEqProd

View file

@ -3,7 +3,7 @@ Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.FeynmanDiagrams.Basic
import HepLean.FeynmanDiagrams.Momentum
/-!
# Feynman diagrams in Phi^4 theory
@ -65,6 +65,35 @@ instance : IsFinitePreFeynmanRule phi4PreFeynmanRules where
match v with
| 0 => instDecidableEqFin _
/-!
## The figure eight diagram
This section provides an example of the use of Feynman diagrams in HepLean.
-/
section Example
/-- The figure eight Feynman diagram. -/
abbrev figureEight : FeynmanDiagram phi4PreFeynmanRules :=
mk'
![0, 0] -- edges
![1] -- one internal vertex
![⟨0, 0, 0⟩, ⟨0, 0, 0⟩, ⟨0, 1, 0⟩, ⟨0, 1, 0⟩] -- four half-edges
(by decide) -- the condition to form a Feynman diagram.
/-- `figureEight` is connected. We can get this from
`#eval Connected figureEight`. -/
lemma figureEight_connected : Connected figureEight := by decide
/-- The symmetry factor of `figureEight` is 8. We can get this from
`#eval symmetryFactor figureEight`. -/
lemma figureEight_symmetryFactor : symmetryFactor figureEight = 8 := by decide
lemma figureEight_halfEdgeToEdgeIntMatrix :
figureEight.halfEdgeToEdgeIntMatrix = !![1, 0; 1, 0 ; 0, 1; 0, 1] := by decide
end Example
end PhiFour

View file

@ -6,6 +6,11 @@ Authors: Joseph Tooby-Smith
import HepLean.FeynmanDiagrams.Basic
import Mathlib.Data.Real.Basic
import Mathlib.Algebra.Category.ModuleCat.Basic
import Mathlib.LinearAlgebra.StdBasis
import Mathlib.LinearAlgebra.Matrix.ToLin
import Mathlib.Data.Matrix.Rank
import Mathlib.Algebra.DirectSum.Module
import Mathlib.LinearAlgebra.SesquilinearForm
/-!
# Momentum in Feynman diagrams
@ -15,72 +20,176 @@ and constrain the momentums based conservation at each vertex and edge.
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`.
-/
namespace FeynmanDiagram
open CategoryTheory
open PreFeynmanRule
variable {P : PreFeynmanRule} (F : FeynmanDiagram P)
variable (d : )
/-- The momentum space for a `d`-dimensional field theory for a single particle.
TODO: Move this definition, and define it as a four-vector. -/
def SingleMomentumSpace : Type := Fin d →
instance : AddCommGroup (SingleMomentumSpace d) := Pi.addCommGroup
instance : Module (SingleMomentumSpace d) := Pi.module _ _ _
/-- The type which asociates to each half-edge a `d`-dimensional vector.
This is to be interpreted as the momentum associated to that half-edge flowing from the
corresponding `edge` to the corresponding `vertex`. So all momentums flow into vertices. -/
def FullMomentumSpace : Type := F.𝓱𝓔 → Fin d →
instance : AddCommGroup (F.FullMomentumSpace d) := Pi.addCommGroup
instance : Module (F.FullMomentumSpace d) := Pi.module _ _ _
/-- The linear map taking a half-edge to its momentum.
(defined as flowing from the `edge` to the vertex.) -/
def toHalfEdgeMomentum (i : F.𝓱𝓔) : F.FullMomentumSpace d →ₗ[] SingleMomentumSpace d where
toFun x := x i
map_add' x y := by rfl
map_smul' c x := by rfl
namespace Hom
variable {F G : FeynmanDiagram P}
variable (f : F ⟶ G)
/-- The linear map induced by a morphism of Feynman diagrams. -/
def toLinearMap : G.FullMomentumSpace d →ₗ[] F.FullMomentumSpace d where
toFun x := x ∘ f.𝓱𝓔
map_add' x y := by rfl
map_smul' c x := by rfl
end Hom
/-- The contravariant functor from Feynman diagrams to Modules over ``. -/
noncomputable def funcFullMomentumSpace : FeynmanDiagram P ⥤ (ModuleCat )ᵒᵖ where
obj F := Opposite.op $ ModuleCat.of (F.FullMomentumSpace d)
map f := Opposite.op $ Hom.toLinearMap d f
variable {P : PreFeynmanRule} (F : FeynmanDiagram P) [IsFiniteDiagram F]
/-!
## Edge constraints
There is a linear map from `F.FullMomentumSpace` to `F.EdgeMomentumSpace`, induced
by the constraints at each edge.
## Vector spaces associated with momenta in Feynman diagrams.
We impose the constraint that we live in the kernal of this linear map.
We define the vector space associated with momenta carried by half-edges,
outflowing momenta of edges, and inflowing momenta of vertices.
A similar result is true for the vertex constraints.
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]
/-- 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.𝓔
instance : AddCommGroup F.EdgeMomenta := Pi.addCommGroup
instance : Module F.EdgeMomenta := Pi.module _ _ _
/-- 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.𝓥
instance : AddCommGroup F.VertexMomenta := Pi.addCommGroup
instance : Module F.VertexMomenta := Pi.module _ _ _
/-- 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
instance (i : Fin 2) : AddCommGroup (EdgeVertexMomentaMap F i) :=
match i with
| 0 => instAddCommGroupEdgeMomenta F
| 1 => instAddCommGroupVertexMomenta F
instance (i : Fin 2) : Module (EdgeVertexMomentaMap F i) :=
match i with
| 0 => instModuleRealEdgeMomenta F
| 1 => instModuleRealVertexMomenta F
/-- The direct sum of `EdgeMomenta` and `VertexMomenta`.-/
def EdgeVertexMomenta : Type := DirectSum (Fin 2) (EdgeVertexMomentaMap F)
instance : AddCommGroup F.EdgeVertexMomenta := DirectSum.instAddCommGroup _
instance : Module F.EdgeVertexMomenta := DirectSum.instModule
/-!
## Linear maps between the vector spaces.
We define various maps into `F.HalfEdgeMomenta`.
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.
-/
/-- 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
end FeynmanDiagram