PhysLean/HepLean/PerturbationTheory/FeynmanDiagrams/Instances/ComplexScalar.lean

70 lines
2.1 KiB
Text
Raw Normal View History

2024-06-18 11:40:36 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
2024-07-12 16:39:44 -04:00
Released under Apache 2.0 license as described in the file LICENSE.
2024-06-18 11:40:36 -04:00
Authors: Joseph Tooby-Smith
-/
2024-12-02 16:26:19 +00:00
import HepLean.PerturbationTheory.FeynmanDiagrams.Basic
2024-06-18 11:40:36 -04:00
/-!
# Feynman diagrams in a complex scalar field theory
-/
namespace PhiFour
open CategoryTheory
open FeynmanDiagram
open PreFeynmanRule
2024-06-18 13:07:49 -04:00
/-- The pre-Feynman rules for a complex scalar theory. -/
2024-06-18 11:40:36 -04:00
@[simps!]
def complexScalarFeynmanRules : PreFeynmanRule where
/- There is 2 types of `half-edge`. -/
HalfEdgeLabel := Fin 2
/- There is only 1 type of `edge`. -/
EdgeLabel := Fin 1
/- There are two types of `vertex`, two external `0` and internal `1`. -/
VertexLabel := Fin 3
edgeLabelMap x :=
match x with
| 0 => Over.mk ![0, 1]
vertexLabelMap x :=
match x with
| 0 => Over.mk ![0]
| 1 => Over.mk ![1]
| 2 => Over.mk ![0, 0, 1, 1]
2024-11-12 05:49:43 +00:00
/-- An instance allowing us to use integers for edge labels for complex scalar theory. -/
2024-06-18 11:40:36 -04:00
instance (a : ) : OfNat complexScalarFeynmanRules.EdgeLabel a where
ofNat := (a : Fin _)
2024-11-12 05:49:43 +00:00
/-- An instance allowing us to use integers for half-edge labels for complex scalar theory. -/
2024-06-18 11:40:36 -04:00
instance (a : ) : OfNat complexScalarFeynmanRules.HalfEdgeLabel a where
ofNat := (a : Fin _)
2024-11-12 05:49:43 +00:00
/-- An instance allowing us to use integers for vertex labels for complex scalar theory. -/
2024-06-18 11:40:36 -04:00
instance (a : ) : OfNat complexScalarFeynmanRules.VertexLabel a where
ofNat := (a : Fin _)
2024-11-12 05:49:43 +00:00
/-- The pre feynman rules for complex scalars are finite. -/
2024-06-18 11:40:36 -04:00
instance : IsFinitePreFeynmanRule complexScalarFeynmanRules where
2024-07-12 11:23:02 -04:00
edgeLabelDecidable := instDecidableEqFin _
2024-06-18 11:40:36 -04:00
vertexLabelDecidable := instDecidableEqFin _
halfEdgeLabelDecidable := instDecidableEqFin _
vertexMapFintype := fun v =>
match v with
| 0 => Fin.fintype _
2024-07-12 11:23:02 -04:00
| 1 => Fin.fintype _
| 2 => Fin.fintype _
2024-06-18 11:40:36 -04:00
edgeMapFintype := fun v =>
match v with
| 0 => Fin.fintype _
vertexMapDecidable := fun v =>
match v with
| 0 => instDecidableEqFin _
2024-07-12 11:23:02 -04:00
| 1 => instDecidableEqFin _
| 2 => instDecidableEqFin _
2024-06-18 11:40:36 -04:00
edgeMapDecidable := fun v =>
match v with
| 0 => instDecidableEqFin _
end PhiFour