PhysLean/HepLean/FeynmanDiagrams/Instances/Phi4.lean
2024-06-18 11:40:36 -04:00

83 lines
2.1 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) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license.
Authors: Joseph Tooby-Smith
-/
import HepLean.FeynmanDiagrams.Basic
/-!
# Feynman diagrams in Phi^4 theory
The aim of this file is to start building up the theory of Feynman diagrams in the context of
Phi^4 theory.
-/
namespace PhiFour
open CategoryTheory
open FeynmanDiagram
open PreFeynmanRule
@[simps!]
def phi4PreFeynmanRules : PreFeynmanRule where
/- There is only 1 type of `half-edge`. -/
HalfEdgeLabel := Fin 1
/- There is only 1 type of `edge`. -/
EdgeLabel := Fin 1
/- There are two types of `vertex`, external `0` and internal `1`. -/
VertexLabel := Fin 2
edgeLabelMap x :=
match x with
| 0 => Over.mk ![0, 0]
vertexLabelMap x :=
match x with
| 0 => Over.mk ![0]
| 1 => Over.mk ![0, 0, 0, 0]
instance (a : ) : OfNat phi4PreFeynmanRules.EdgeLabel a where
ofNat := (a : Fin _)
instance (a : ) : OfNat phi4PreFeynmanRules.HalfEdgeLabel a where
ofNat := (a : Fin _)
instance (a : ) : OfNat phi4PreFeynmanRules.VertexLabel a where
ofNat := (a : Fin _)
instance : IsFinitePreFeynmanRule phi4PreFeynmanRules where
edgeLabelDecidable := instDecidableEqFin _
vertexLabelDecidable := instDecidableEqFin _
halfEdgeLabelDecidable := instDecidableEqFin _
vertexMapFintype := fun v =>
match v with
| 0 => Fin.fintype _
| 1 => Fin.fintype _
edgeMapFintype := fun v =>
match v with
| 0 => Fin.fintype _
vertexMapDecidable := fun v =>
match v with
| 0 => instDecidableEqFin _
| 1 => instDecidableEqFin _
edgeMapDecidable := fun v =>
match v with
| 0 => instDecidableEqFin _
def figureEight : FeynmanDiagram phi4PreFeynmanRules :=
mk' ![0, 0] ![1] ![⟨0, 0, 0⟩, ⟨0, 0, 0⟩, ⟨0, 1, 0⟩, ⟨0, 1, 0⟩] (by decide)
instance : IsFiniteDiagram figureEight where
𝓔Fintype := Fin.fintype _
𝓔DecidableEq := instDecidableEqFin _
𝓥Fintype := Fin.fintype _
𝓥DecidableEq := instDecidableEqFin _
𝓱𝓔Fintype := Fin.fintype _
𝓱𝓔DecidableEq := instDecidableEqFin _
#eval symmetryFactor figureEight
#eval Connected figureEight
end PhiFour