PhysLean/HepLean/FeynmanDiagrams/Instances/ComplexScalar.lean

66 lines
1.7 KiB
Text
Raw Normal View History

2024-06-18 11:40:36 -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
/-!
# 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]
instance (a : ) : OfNat complexScalarFeynmanRules.EdgeLabel a where
ofNat := (a : Fin _)
instance (a : ) : OfNat complexScalarFeynmanRules.HalfEdgeLabel a where
ofNat := (a : Fin _)
instance (a : ) : OfNat complexScalarFeynmanRules.VertexLabel a where
ofNat := (a : Fin _)
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