2024-11-22 15:12:06 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
2024-12-03 09:51:09 +00:00
|
|
|
|
import HepLean.Meta.Informal
|
2024-12-03 15:12:54 +00:00
|
|
|
|
import HepLean.PerturbationTheory.Wick.Species
|
2024-12-03 09:51:09 +00:00
|
|
|
|
import Mathlib.Data.Fin.Tuple.Basic
|
2024-11-22 15:12:06 +00:00
|
|
|
|
/-!
|
|
|
|
|
# Wick strings
|
|
|
|
|
|
|
|
|
|
A wick string is defined to be a sequence of input fields,
|
|
|
|
|
followed by a squence of vertices, followed by a sequence of output fields.
|
|
|
|
|
|
|
|
|
|
A wick string can be combined with an appropriate map to spacetime to produce a specific
|
|
|
|
|
term in the ring of operators. This has yet to be implemented.
|
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
2024-12-03 15:12:54 +00:00
|
|
|
|
namespace Wick
|
|
|
|
|
|
|
|
|
|
variable {S : Species}
|
2024-11-22 15:12:06 +00:00
|
|
|
|
|
2024-11-22 16:02:26 +00:00
|
|
|
|
/-- A helper function for `WickString`. It is used to seperate incoming, vertex, and
|
|
|
|
|
outgoing nodes. -/
|
2024-11-22 15:12:06 +00:00
|
|
|
|
inductive WickStringLast where
|
|
|
|
|
| incoming : WickStringLast
|
|
|
|
|
| vertex : WickStringLast
|
|
|
|
|
| outgoing : WickStringLast
|
|
|
|
|
| final : WickStringLast
|
|
|
|
|
|
|
|
|
|
open WickStringLast
|
|
|
|
|
|
|
|
|
|
/-- A wick string is a representation of a string of fields from a theory.
|
2024-12-02 12:39:47 +00:00
|
|
|
|
The use of vertices in the Wick string
|
2024-11-29 06:35:33 +00:00
|
|
|
|
allows us to identify which fields have the same space-time coordinate.
|
|
|
|
|
|
|
|
|
|
Note: Fields are added to `c` from right to left - matching how we would write this on
|
2024-12-02 12:39:47 +00:00
|
|
|
|
pen and paper.
|
|
|
|
|
|
|
|
|
|
The incoming and outgoing fields should be viewed as asymptotic fields.
|
|
|
|
|
While the internal fields associated with vertices are fields at fixed space-time points.
|
|
|
|
|
-/
|
2024-12-03 15:12:54 +00:00
|
|
|
|
inductive WickString : {ni : ℕ} → (i : Fin ni → S.𝓯) → {n : ℕ} → (c : Fin n → S.𝓯) →
|
|
|
|
|
{no : ℕ} → (o : Fin no → S.𝓯) → WickStringLast → Type where
|
2024-11-26 09:05:55 +00:00
|
|
|
|
| empty : WickString Fin.elim0 Fin.elim0 Fin.elim0 incoming
|
2024-12-03 15:12:54 +00:00
|
|
|
|
| incoming {n ni no : ℕ} {i : Fin ni → S.𝓯} {c : Fin n → S.𝓯}
|
|
|
|
|
{o : Fin no → S.𝓯} (w : WickString i c o incoming) (e : S.𝓯) :
|
2024-11-26 09:05:55 +00:00
|
|
|
|
WickString (Fin.cons e i) (Fin.cons e c) o incoming
|
2024-12-03 15:12:54 +00:00
|
|
|
|
| endIncoming {n ni no : ℕ} {i : Fin ni → S.𝓯} {c : Fin n → S.𝓯}
|
|
|
|
|
{o : Fin no → S.𝓯} (w : WickString i c o incoming) : WickString i c o vertex
|
|
|
|
|
| vertex {n ni no : ℕ} {i : Fin ni → S.𝓯} {c : Fin n → S.𝓯}
|
|
|
|
|
{o : Fin no → S.𝓯} (w : WickString i c o vertex) (v : S.𝓘) :
|
|
|
|
|
WickString i (Fin.append (S.𝓘Fields v).2 c) o vertex
|
|
|
|
|
| endVertex {n ni no : ℕ} {i : Fin ni → S.𝓯} {c : Fin n → S.𝓯}
|
|
|
|
|
{o : Fin no → S.𝓯} (w : WickString i c o vertex) : WickString i c o outgoing
|
|
|
|
|
| outgoing {n ni no : ℕ} {i : Fin ni → S.𝓯} {c : Fin n → S.𝓯}
|
|
|
|
|
{o : Fin no → S.𝓯} (w : WickString i c o outgoing) (e : S.𝓯) :
|
2024-11-26 09:05:55 +00:00
|
|
|
|
WickString i (Fin.cons e c) (Fin.cons e o) outgoing
|
2024-12-03 15:12:54 +00:00
|
|
|
|
| endOutgoing {n ni no : ℕ} {i : Fin ni → S.𝓯} {c : Fin n → S.𝓯}
|
|
|
|
|
{o : Fin no → S.𝓯} (w : WickString i c o outgoing) : WickString i c o final
|
2024-11-26 09:05:55 +00:00
|
|
|
|
|
|
|
|
|
namespace WickString
|
|
|
|
|
|
|
|
|
|
/-- The number of nodes in a Wick string. This is used to help prove termination. -/
|
2024-12-03 15:12:54 +00:00
|
|
|
|
def size {ni : ℕ} {i : Fin ni → S.𝓯} {n : ℕ} {c : Fin n → S.𝓯} {no : ℕ} {o : Fin no → S.𝓯}
|
2024-11-29 06:38:52 +00:00
|
|
|
|
{f : WickStringLast} : WickString i c o f → ℕ := fun
|
2024-11-26 09:05:55 +00:00
|
|
|
|
| empty => 0
|
|
|
|
|
| incoming w e => size w + 1
|
|
|
|
|
| endIncoming w => size w + 1
|
|
|
|
|
| vertex w v => size w + 1
|
|
|
|
|
| endVertex w => size w + 1
|
|
|
|
|
| outgoing w e => size w + 1
|
|
|
|
|
| endOutgoing w => size w + 1
|
|
|
|
|
|
2024-11-29 16:20:54 +00:00
|
|
|
|
/-- The number of vertices in a Wick string. This does NOT include external vertices. -/
|
2024-12-03 15:12:54 +00:00
|
|
|
|
def numIntVertex {ni : ℕ} {i : Fin ni → S.𝓯} {n : ℕ} {c : Fin n → S.𝓯} {no : ℕ} {o : Fin no → S.𝓯}
|
2024-11-26 09:05:55 +00:00
|
|
|
|
{f : WickStringLast} : WickString i c o f → ℕ := fun
|
|
|
|
|
| empty => 0
|
2024-12-02 06:00:59 +00:00
|
|
|
|
| incoming w e => numIntVertex w
|
|
|
|
|
| endIncoming w => numIntVertex w
|
|
|
|
|
| vertex w v => numIntVertex w + 1
|
|
|
|
|
| endVertex w => numIntVertex w
|
|
|
|
|
| outgoing w e => numIntVertex w
|
|
|
|
|
| endOutgoing w => numIntVertex w
|
2024-11-26 09:05:55 +00:00
|
|
|
|
|
2024-11-29 16:20:54 +00:00
|
|
|
|
/-- The vertices present in a Wick string. This does NOT include external vertices. -/
|
2024-12-03 15:12:54 +00:00
|
|
|
|
def intVertex {ni : ℕ} {i : Fin ni → S.𝓯} {n : ℕ} {c : Fin n → S.𝓯} {no : ℕ} {o : Fin no → S.𝓯}
|
|
|
|
|
{f : WickStringLast} : (w : WickString i c o f) → Fin w.numIntVertex → S.𝓘 := fun
|
2024-11-29 06:35:33 +00:00
|
|
|
|
| empty => Fin.elim0
|
2024-12-02 06:00:59 +00:00
|
|
|
|
| incoming w e => intVertex w
|
|
|
|
|
| endIncoming w => intVertex w
|
|
|
|
|
| vertex w v => Fin.cons v (intVertex w)
|
|
|
|
|
| endVertex w => intVertex w
|
|
|
|
|
| outgoing w e => intVertex w
|
|
|
|
|
| endOutgoing w => intVertex w
|
|
|
|
|
|
|
|
|
|
informal_definition intExtVertex where
|
|
|
|
|
math :≈ "The vertices present in a Wick string, including external vertices."
|
|
|
|
|
deps :≈ [``WickString]
|
|
|
|
|
|
|
|
|
|
informal_definition fieldToIntExtVertex where
|
|
|
|
|
math :≈ "A function which takes a field and returns the internal or
|
|
|
|
|
external vertex it is associated with."
|
2024-11-29 16:20:54 +00:00
|
|
|
|
deps :≈ [``WickString]
|
|
|
|
|
|
2024-11-29 07:53:19 +00:00
|
|
|
|
informal_definition exponentialPrefactor where
|
|
|
|
|
math :≈ "The combinatorical prefactor from the expansion of the
|
|
|
|
|
exponential associated with a Wick string."
|
2024-12-02 06:00:59 +00:00
|
|
|
|
deps :≈ [``intVertex, ``WickString]
|
2024-11-29 07:53:19 +00:00
|
|
|
|
|
|
|
|
|
informal_definition vertexPrefactor where
|
|
|
|
|
math :≈ "The prefactor arising from the coefficent of vertices in the
|
|
|
|
|
Lagrangian. This should not take account of the exponential prefactor."
|
2024-12-02 06:00:59 +00:00
|
|
|
|
deps :≈ [``intVertex, ``WickString]
|
2024-11-29 07:53:19 +00:00
|
|
|
|
|
|
|
|
|
informal_definition minNoLoops where
|
|
|
|
|
math :≈ "The minimum number of loops a Feynman diagram based on a given Wick string can have.
|
|
|
|
|
There should be a lemma proving this statement."
|
|
|
|
|
deps :≈ [``WickString]
|
|
|
|
|
|
|
|
|
|
informal_definition LoopLevel where
|
|
|
|
|
math :≈ "The type of Wick strings for fixed input and output which may permit a Feynman diagram
|
|
|
|
|
which have a number of loops less than or equal to some number."
|
|
|
|
|
deps :≈ [``minNoLoops, ``WickString]
|
|
|
|
|
|
|
|
|
|
informal_lemma loopLevel_fintype where
|
|
|
|
|
math :≈ "The instance of a finite type on `LoopLevel`."
|
|
|
|
|
deps :≈ [``LoopLevel]
|
|
|
|
|
|
2024-11-26 09:05:55 +00:00
|
|
|
|
end WickString
|
2024-11-22 15:12:06 +00:00
|
|
|
|
|
2024-12-03 15:12:54 +00:00
|
|
|
|
end Wick
|