Merge pull request #259 from HEPLean/PerturbationTheory

feat: Update Wick contraction and string
This commit is contained in:
Joseph Tooby-Smith 2024-12-03 16:35:30 +00:00 committed by GitHub
commit 6b42cdd4e6
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
8 changed files with 174 additions and 199 deletions

View file

@ -103,6 +103,7 @@ import HepLean.Mathematics.Fin
import HepLean.Mathematics.LinearMaps
import HepLean.Mathematics.PiTensorProduct
import HepLean.Mathematics.SO3.Basic
import HepLean.Mathematics.SuperAlgebra.Basic
import HepLean.Meta.AllFilePaths
import HepLean.Meta.Basic
import HepLean.Meta.Informal

View file

@ -0,0 +1,31 @@
/-
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
-/
import Mathlib.RingTheory.GradedAlgebra.Basic
import HepLean.Meta.Informal
/-!
# Super Algebras
A super algebra is a special type of graded algebra.
It is used in physics to model the commutator of fermionic operators among themselves,
aswell as among bosonic operators.
-/
informal_definition SuperAlgebra where
math :≈ "A super algebra is a graded algebra A with a ℤ₂ grading."
physics :≈ "A super algebra is used to model the commutator of fermionic operators among
themselves, aswell as among bosonic operators."
ref :≈ "https://en.wikipedia.org/wiki/Superalgebra"
namespace SuperAlgebra
informal_definition superCommuator where
math :≈ "The commutator which for `a ∈ Aᵢ` and `b ∈ Aⱼ` is defined as `ab - (-1)^(i * j) ba`."
deps :≈ [``SuperAlgebra]
end SuperAlgebra

View file

@ -35,12 +35,12 @@ informal_definition FeynmanDiagram where
informal_definition _root_.Wick.Contract.toFeynmanDiagram where
math :≈ "The Feynman diagram constructed from a complete Wick contraction."
deps :≈ [``TwoComplexScalar.WickContract, ``FeynmanDiagram]
deps :≈ [``Wick.WickContract, ``FeynmanDiagram]
informal_lemma _root_.Wick.Contract.toFeynmanDiagram_surj where
math :≈ "The map from Wick contractions to Feynman diagrams is surjective."
physics :≈ "Every Feynman digram corresponds to some Wick contraction."
deps :≈ [``TwoComplexScalar.WickContract, ``FeynmanDiagram]
deps :≈ [``Wick.WickContract, ``FeynmanDiagram]
informal_definition FeynmanDiagram.toSimpleGraph where
math :≈ "The simple graph underlying a Feynman diagram."
@ -53,7 +53,7 @@ informal_definition FeynmanDiagram.IsConnected where
informal_definition _root_.Wick.Contract.toFeynmanDiagram_isConnected_iff where
math :≈ "The Feynman diagram corresponding to a Wick contraction is connected
if and only if the Wick contraction is connected."
deps :≈ [``TwoComplexScalar.WickContract.IsConnected, ``FeynmanDiagram.IsConnected]
deps :≈ [``Wick.WickContract.IsConnected, ``FeynmanDiagram.IsConnected]
/-! TODO: Define an equivalence relation on Wick contracts related to the their underlying tensors
been equal after permutation. Show that two Wick contractions are equal under this

View file

@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.PerturbationTheory.Wick.Species
import HepLean.Mathematics.SuperAlgebra.Basic
/-!
# Operator algebra
@ -28,10 +29,10 @@ informal_definition WickAlgebra where
math :≈ "
Modifications of this may be needed.
A structure with the following data:
- A ℤ₂-graded algebra A.
- A map from `ψ : 𝓔 × SpaceTime → A` where 𝓔 are field colors.
- A map `ψc : 𝓔 × SpaceTime → A`.
- A map `ψd : 𝓔 × SpaceTime → A`.
- A super algebra A.
- A map from `ψ : S.𝓯 × SpaceTime → A` where S.𝓯 are field colors.
- A map `ψc : S.𝓯 × SpaceTime → A`.
- A map `ψd : S.𝓯 × SpaceTime → A`.
Subject to the conditions:
- The sum of `ψc` and `ψd` is `ψ`.
- All maps land on homogeneous elements.
@ -39,14 +40,15 @@ informal_definition WickAlgebra where
- The super-commutator of two fields is always in the
center of the algebra.
Asympotic states:
- `φc : 𝓔 × SpaceTime → A`. The creation asympotic state (incoming).
- `φd : 𝓔 × SpaceTime → A`. The destruction asympotic state (outgoing).
- `φc : S.𝓯 × SpaceTime → A`. The creation asympotic state (incoming).
- `φd : S.𝓯 × SpaceTime → A`. The destruction asympotic state (outgoing).
Subject to the conditions:
...
"
physics :≈ "This is defined to be an
abstraction of the notion of an operator algebra."
ref :≈ "https://physics.stackexchange.com/questions/24157/"
deps :≈ [``SuperAlgebra, ``SuperAlgebra.superCommuator]
informal_definition WickMonomial where
math :≈ "The type of elements of the Wick algebra which is a product of fields."
@ -86,18 +88,18 @@ informal_definition normalOrder where
end WickMonomial
informal_definition asymptoicContract where
math :≈ "Given two `i j : 𝓔 × SpaceTime`, the super-commutator [φd(i), ψ(j)]."
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the super-commutator [φd(i), ψ(j)]."
ref :≈ "See e.g. http://www.dylanjtemples.com:82/solutions/QFT_Solution_I-6.pdf"
informal_definition contractAsymptotic where
math :≈ "Given two `i j : 𝓔 × SpaceTime`, the super-commutator [ψ(i), φc(j)]."
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the super-commutator [ψ(i), φc(j)]."
informal_definition asymptoicContractAsymptotic where
math :≈ "Given two `i j : 𝓔 × SpaceTime`, the super-commutator
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the super-commutator
[φd(i), φc(j)]."
informal_definition contraction where
math :≈ "Given two `i j : 𝓔 × SpaceTime`, the element of WickAlgebra
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the element of WickAlgebra
defined by subtracting the normal ordering of `ψ i ψ j` from the time-ordering of
`ψ i ψ j`."
deps :≈ [``WickAlgebra, ``WickMonomial]

View file

@ -11,30 +11,28 @@ import Mathlib.Logic.Equiv.Fin
# Wick Contract
Currently this file is only for an example of Wick contracts, correpsonding to a
theory with two complex scalar fields. The concepts will however generalize.
## Further reading
- https://www.imperial.ac.uk/media/imperial-college/research-centres-and-groups/theoretical-physics/msc/current/qft/handouts/qftwickstheorem.pdf
-/
namespace TwoComplexScalar
namespace Wick
variable {S : Species}
/-- A Wick contraction for a Wick string is a series of pairs `i` and `j` of indices
to be contracted, subject to ordering and subject to the condition that they can
be contracted. -/
inductive WickContract : {ni : } → {i : Fin ni → 𝓔} → {n : } → {c : Fin n → 𝓔} →
{no : } → {o : Fin no → 𝓔} →
inductive WickContract : {ni : } → {i : Fin ni → S.𝓯} → {n : } → {c : Fin n → S.𝓯} →
{no : } → {o : Fin no → S.𝓯} →
(str : WickString i c o final) →
{k : } → (b1 : Fin k → Fin n) → (b2 : Fin k → Fin n) → Type where
| string {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final} : WickContract str Fin.elim0 Fin.elim0
| contr {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final} {k : }
| string {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯}
{str : WickString i c o final} : WickContract str Fin.elim0 Fin.elim0
| contr {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final} {k : }
{b1 : Fin k → Fin n} {b2 : Fin k → Fin n} : (i : Fin n) →
(j : Fin n) → (h : c j = ξ (c i)) →
(j : Fin n) → (h : c j = S.ξ (c i)) →
(hilej : i < j) → (hb1 : ∀ r, b1 r < i) → (hb2i : ∀ r, b2 r ≠ i) → (hb2j : ∀ r, b2 r ≠ j) →
(w : WickContract str b1 b2) →
WickContract str (Fin.snoc b1 i) (Fin.snoc b2 j)
@ -42,15 +40,15 @@ inductive WickContract : {ni : } → {i : Fin ni → 𝓔} → {n : } →
namespace WickContract
/-- The number of nodes of a Wick contraction. -/
def size {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final} {k : } {b1 b2 : Fin k → Fin n} :
def size {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final} {k : } {b1 b2 : Fin k → Fin n} :
WickContract str b1 b2 → := fun
| string => 0
| contr _ _ _ _ _ _ _ w => w.size + 1
/-- The number of nodes in a wick contraction tree is the same as `k`. -/
lemma size_eq_k {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final} {k : } {b1 b2 : Fin k → Fin n} :
lemma size_eq_k {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final} {k : } {b1 b2 : Fin k → Fin n} :
(w : WickContract str b1 b2) → w.size = k := fun
| string => rfl
| contr _ _ _ _ _ _ _ w => by
@ -58,16 +56,16 @@ lemma size_eq_k {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
/-- The map giving the vertices on the left-hand-side of a contraction. -/
@[nolint unusedArguments]
def boundFst {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def boundFst {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} :
WickContract str b1 b2 → Fin k → Fin n := fun _ => b1
@[simp]
lemma boundFst_contr_castSucc {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundFst_contr_castSucc {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (i j : Fin n)
(h : c j = ξ (c i))
(h : c j = S.ξ (c i))
(hilej : i < j)
(hb1 : ∀ r, b1 r < i)
(hb2i : ∀ r, b2 r ≠ i)
@ -77,10 +75,10 @@ lemma boundFst_contr_castSucc {ni : } {i : Fin ni → 𝓔} {n : } {c : Fi
simp only [boundFst, Fin.snoc_castSucc]
@[simp]
lemma boundFst_contr_last {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundFst_contr_last {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (i j : Fin n)
(h : c j = ξ (c i))
(h : c j = S.ξ (c i))
(hilej : i < j)
(hb1 : ∀ r, b1 r < i)
(hb2i : ∀ r, b2 r ≠ i)
@ -89,8 +87,8 @@ lemma boundFst_contr_last {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
(contr i j h hilej hb1 hb2i hb2j w).boundFst (Fin.last k) = i := by
simp only [boundFst, Fin.snoc_last]
lemma boundFst_strictMono {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundFst_strictMono {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} : (w : WickContract str b1 b2) → StrictMono w.boundFst := fun
| string => fun k => Fin.elim0 k
| contr i j _ _ hb1 _ _ w => by
@ -119,16 +117,16 @@ lemma boundFst_strictMono {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
/-- The map giving the vertices on the right-hand-side of a contraction. -/
@[nolint unusedArguments]
def boundSnd {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def boundSnd {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} :
WickContract str b1 b2 → Fin k → Fin n := fun _ => b2
@[simp]
lemma boundSnd_contr_castSucc {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundSnd_contr_castSucc {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (i j : Fin n)
(h : c j = ξ (c i))
(h : c j = S.ξ (c i))
(hilej : i < j)
(hb1 : ∀ r, b1 r < i)
(hb2i : ∀ r, b2 r ≠ i)
@ -138,10 +136,10 @@ lemma boundSnd_contr_castSucc {ni : } {i : Fin ni → 𝓔} {n : } {c : Fi
simp only [boundSnd, Fin.snoc_castSucc]
@[simp]
lemma boundSnd_contr_last {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundSnd_contr_last {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (i j : Fin n)
(h : c j = ξ (c i))
(h : c j = S.ξ (c i))
(hilej : i < j)
(hb1 : ∀ r, b1 r < i)
(hb2i : ∀ r, b2 r ≠ i)
@ -150,8 +148,8 @@ lemma boundSnd_contr_last {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
(contr i j h hilej hb1 hb2i hb2j w).boundSnd (Fin.last k) = j := by
simp only [boundSnd, Fin.snoc_last]
lemma boundSnd_injective {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundSnd_injective {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} :
(w : WickContract str b1 b2) → Function.Injective w.boundSnd := fun
| string => by
@ -179,10 +177,10 @@ lemma boundSnd_injective {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
· subst hs
rfl
lemma color_boundSnd_eq_dual_boundFst {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma color_boundSnd_eq_dual_boundFst {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} :
(w : WickContract str b1 b2) → (i : Fin k) → c (w.boundSnd i) = ξ (c (w.boundFst i)) := fun
(w : WickContract str b1 b2) → (i : Fin k) → c (w.boundSnd i) = S.ξ (c (w.boundFst i)) := fun
| string => fun i => Fin.elim0 i
| contr i j hij hilej hi _ _ w => fun r => by
rcases Fin.eq_castSucc_or_eq_last r with hr | hr
@ -192,8 +190,8 @@ lemma color_boundSnd_eq_dual_boundFst {ni : } {i : Fin ni → 𝓔} {n : }
· subst hr
simpa using hij
lemma boundFst_lt_boundSnd {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundFst_lt_boundSnd {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} : (w : WickContract str b1 b2) → (i : Fin k) →
w.boundFst i < w.boundSnd i := fun
| string => fun i => Fin.elim0 i
@ -206,8 +204,8 @@ lemma boundFst_lt_boundSnd {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
simp only [boundFst_contr_last, boundSnd_contr_last]
exact hilej
lemma boundFst_neq_boundSnd {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma boundFst_neq_boundSnd {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} :
(w : WickContract str b1 b2) → (r1 r2 : Fin k) → b1 r1 ≠ b2 r2 := fun
| string => fun i => Fin.elim0 i
@ -233,8 +231,8 @@ lemma boundFst_neq_boundSnd {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin
/-- Casts a Wick contraction from `WickContract str b1 b2` to `WickContract str b1' b2'` with a
proof that `b1 = b1'` and `b2 = b2'`, and that they are defined from the same `k = k'`. -/
def castMaps {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def castMaps {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k k' : } {b1 b2 : Fin k → Fin n} {b1' b2' : Fin k' → Fin n}
(hk : k = k')
(hb1 : b1 = b1' ∘ Fin.cast hk) (hb2 : b2 = b2' ∘ Fin.cast hk) (w : WickContract str b1 b2) :
@ -242,17 +240,17 @@ def castMaps {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
cast (by subst hk; rfl) (hb2 ▸ hb1 ▸ w)
@[simp]
lemma castMaps_rfl {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma castMaps_rfl {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (w : WickContract str b1 b2) :
castMaps rfl rfl rfl w = w := rfl
lemma mem_snoc' {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma mem_snoc' {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1' b2' : Fin k → Fin n} :
(w : WickContract str b1' b2') →
{k' : } → (hk' : k'.succ = k) →
(b1 b2 : Fin k' → Fin n) → (i j : Fin n) → (h : c j = ξ (c i)) →
(b1 b2 : Fin k' → Fin n) → (i j : Fin n) → (h : c j = S.ξ (c i)) →
(hilej : i < j) → (hb1 : ∀ r, b1 r < i) → (hb2i : ∀ r, b2 r ≠ i) → (hb2j : ∀ r, b2 r ≠ j) →
(hb1' : Fin.snoc b1 i = b1' ∘ Fin.cast hk') →
(hb2' : Fin.snoc b2 j = b2' ∘ Fin.cast hk') →
@ -291,17 +289,17 @@ lemma mem_snoc' {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
subst hb1'' hb2'' hi hj
simp
lemma mem_snoc {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma mem_snoc {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(i j : Fin n) (h : c j = ξ (c i)) (hilej : i < j) (hb1 : ∀ r, b1 r < i)
(i j : Fin n) (h : c j = S.ξ (c i)) (hilej : i < j) (hb1 : ∀ r, b1 r < i)
(hb2i : ∀ r, b2 r ≠ i) (hb2j : ∀ r, b2 r ≠ j)
(w : WickContract str (Fin.snoc b1 i) (Fin.snoc b2 j)) :
∃ (w' : WickContract str b1 b2), w = contr i j h hilej hb1 hb2i hb2j w' := by
exact mem_snoc' w rfl b1 b2 i j h hilej hb1 hb2i hb2j rfl rfl
lemma is_subsingleton {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma is_subsingleton {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} :
Subsingleton (WickContract str b1 b2) := Subsingleton.intro fun w1 w2 => by
induction k with
@ -330,10 +328,10 @@ lemma eq_snoc_castSucc {k n : } (b1 : Fin k.succ → Fin n) :
/-- The construction of a Wick contraction from maps `b1 b2 : Fin k → Fin n`, with the former
giving the first index to be contracted, and the latter the second index. These
maps must satisfy a series of conditions. -/
def fromMaps {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def fromMaps {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } (b1 b2 : Fin k → Fin n)
(hi : ∀ i, c (b2 i) = ξ (c (b1 i)))
(hi : ∀ i, c (b2 i) = S.ξ (c (b1 i)))
(hb1ltb2 : ∀ i, b1 i < b2 i)
(hb1 : StrictMono b1)
(hb1neb2 : ∀ r1 r2, b1 r1 ≠ b2 r2)
@ -361,8 +359,8 @@ def fromMaps {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
/-- Given a Wick contraction with `k.succ` contractions, returns the Wick contraction with
`k` contractions by dropping the last contraction (defined by the first index contracted). -/
def dropLast {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def dropLast {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k.succ → Fin n}
(w : WickContract str b1 b2) : WickContract str (b1 ∘ Fin.castSucc) (b2 ∘ Fin.castSucc) :=
fromMaps (b1 ∘ Fin.castSucc) (b2 ∘ Fin.castSucc)
@ -372,16 +370,16 @@ def dropLast {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
(fun r1 r2 => boundFst_neq_boundSnd w r1.castSucc r2.castSucc)
(Function.Injective.comp w.boundSnd_injective (Fin.castSucc_injective k))
lemma eq_from_maps {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma eq_from_maps {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) :
w = fromMaps w.boundFst w.boundSnd w.color_boundSnd_eq_dual_boundFst
w.boundFst_lt_boundSnd w.boundFst_strictMono w.boundFst_neq_boundSnd
w.boundSnd_injective := is_subsingleton.allEq w _
lemma eq_dropLast_contr {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma eq_dropLast_contr {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k.succ → Fin n} (w : WickContract str b1 b2) :
w = castMaps rfl (eq_snoc_castSucc b1).symm (eq_snoc_castSucc b2).symm
(contr (b1 (Fin.last k)) (b2 (Fin.last k))
@ -395,14 +393,14 @@ lemma eq_dropLast_contr {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
rfl
/-- Wick contractions of a given Wick string with `k` different contractions. -/
def Level {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} (str : WickString i c o final) (k : ) : Type :=
def Level {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} (str : WickString i c o final) (k : ) : Type :=
Σ (b1 : Fin k → Fin n) (b2 : Fin k → Fin n), WickContract str b1 b2
/-- There is a finite number of Wick contractions with no contractions. In particular,
this is just the original Wick string. -/
instance levelZeroFintype {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} (str : WickString i c o final) :
instance levelZeroFintype {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} (str : WickString i c o final) :
Fintype (Level str 0) where
elems := {⟨Fin.elim0, Fin.elim0, WickContract.string⟩}
complete := by
@ -416,15 +414,15 @@ instance levelZeroFintype {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
rw [is_subsingleton.allEq w string]
/-- The pairs of additional indices which can be contracted given a Wick contraction. -/
structure ContrPair {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
structure ContrPair {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) where
/-- The first index in the contraction pair. -/
i : Fin n
/-- The second index in the contraction pair. -/
j : Fin n
h : c j = ξ (c i)
h : c j = S.ξ (c i)
hilej : i < j
hb1 : ∀ r, b1 r < i
hb2i : ∀ r, b2 r ≠ i
@ -433,10 +431,10 @@ structure ContrPair {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n →
/-- The pairs of additional indices which can be contracted, given an existing wick contraction,
is equivalent to the a subtype of `Fin n × Fin n` defined by certain conditions equivalent
to the conditions appearing in `ContrPair`. -/
def contrPairEquivSubtype {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def contrPairEquivSubtype {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (w : WickContract str b1 b2) :
ContrPair w ≃ {x : Fin n × Fin n // c x.2 = ξ (c x.1) ∧ x.1 < x.2 ∧
ContrPair w ≃ {x : Fin n × Fin n // c x.2 = S.ξ (c x.1) ∧ x.1 < x.2 ∧
(∀ r, b1 r < x.1) ∧ (∀ r, b2 r ≠ x.1) ∧ (∀ r, b2 r ≠ x.2)} where
toFun cp := ⟨⟨cp.i, cp.j⟩, ⟨cp.h, cp.hilej, cp.hb1, cp.hb2i, cp.hb2j⟩⟩
invFun x :=
@ -453,8 +451,8 @@ def contrPairEquivSubtype {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
obtain ⟨left_3, right⟩ := right
simp_all only [ne_eq]
lemma heq_eq {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma heq_eq {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 b1' b2' : Fin k → Fin n}
(w : WickContract str b1 b2)
(w' : WickContract str b1' b2') (h1 : b1 = b1') (h2 : b2 = b2') : HEq w w':= by
@ -464,8 +462,8 @@ lemma heq_eq {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
/-- The equivalence between Wick contractions consisting of `k.succ` contractions and
those with `k` contractions paired with a suitable contraction pair. -/
def levelSuccEquiv {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} (str : WickString i c o final) (k : ) :
def levelSuccEquiv {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} (str : WickString i c o final) (k : ) :
Level str k.succ ≃ (w : Level str k) × ContrPair w.2.2 where
toFun w :=
match w with
@ -517,28 +515,28 @@ def levelSuccEquiv {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n →
/-- The sum of `boundFst` and `boundSnd`, giving on `Sum.inl k` the first index
in the `k`th contraction, and on `Sum.inr k` the second index in the `k`th contraction. -/
def bound {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def bound {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) : Fin k ⊕ Fin k → Fin n :=
Sum.elim w.boundFst w.boundSnd
/-- On `Sum.inl k` the map `bound` acts via `boundFst`. -/
@[simp]
lemma bound_inl {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma bound_inl {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) (i : Fin k) : w.bound (Sum.inl i) = w.boundFst i := rfl
/-- On `Sum.inr k` the map `bound` acts via `boundSnd`. -/
@[simp]
lemma bound_inr {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma bound_inr {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) (i : Fin k) : w.bound (Sum.inr i) = w.boundSnd i := rfl
lemma bound_injection {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma bound_injection {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) : Function.Injective w.bound := by
intro x y h
@ -556,8 +554,8 @@ lemma bound_injection {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n →
simp only [bound_inr, bound_inl] at h
exact False.elim (w.boundFst_neq_boundSnd y x h.symm)
lemma bound_le_total {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma bound_le_total {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) : 2 * k ≤ n := by
refine Fin.nonempty_embedding_iff.mp ⟨w.bound ∘ finSumFinEquiv.symm ∘ Fin.cast (Nat.two_mul k),
@ -568,23 +566,23 @@ lemma bound_le_total {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n →
/-- The list of fields (indexed by `Fin n`) in a Wick contraction which are not bound,
i.e. which do not appear in any contraction. -/
def unboundList {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def unboundList {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) : List (Fin n) :=
List.filter (fun i => decide (∀ r, w.bound r ≠ i)) (List.finRange n)
/-- THe list of field positions which are not contracted has no duplicates. -/
lemma unboundList_nodup {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma unboundList_nodup {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) : (w.unboundList).Nodup :=
List.Nodup.filter _ (List.nodup_finRange n)
/-- The length of the `unboundList` is equal to `n - 2 * k`. That is
the total number of fields minus the number of contracted fields. -/
lemma unboundList_length {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma unboundList_length {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (w : WickContract str b1 b2) :
w.unboundList.length = n - 2 * k := by
rw [← List.Nodup.dedup w.unboundList_nodup]
@ -610,16 +608,16 @@ lemma unboundList_length {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n
decide_eq_true_eq, Finset.mem_image, Finset.mem_univ, true_and, Sum.exists, not_or, not_exists]
exact bound_injection w
lemma unboundList_sorted {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
lemma unboundList_sorted {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n} (w : WickContract str b1 b2) :
List.Sorted (fun i j => i < j) w.unboundList :=
List.Pairwise.sublist (List.filter_sublist (List.finRange n)) (List.pairwise_lt_finRange n)
/-- The ordered embedding giving the fields which are not bound in a contraction. These
are the fields that will appear in a normal operator in Wick's theorem. -/
def unbound {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
{no : } {o : Fin no → 𝓔} {str : WickString i c o final}
def unbound {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯}
{no : } {o : Fin no → S.𝓯} {str : WickString i c o final}
{k : } {b1 b2 : Fin k → Fin n}
(w : WickContract str b1 b2) : Fin (n - 2 * k) ↪o Fin n where
toFun := w.unboundList.get ∘ Fin.cast w.unboundList_length.symm
@ -659,4 +657,4 @@ informal_definition IsOneParticleIrreducible where
end WickContract
end TwoComplexScalar
end Wick

View file

@ -22,8 +22,7 @@ namespace Wick
/-- The basic structure needed to write down Wick contractions for a theory and
calculate the corresponding Feynman diagrams.
WARNING: This definition is not yet complete.
-/
WARNING: This definition is not yet complete. -/
structure Species where
/-- The color of Field operators which appear in a theory.
One may wish to call these `half-edges`, however we restrict this terminology

View file

@ -4,13 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Meta.Informal
import HepLean.PerturbationTheory.Wick.Species
import Mathlib.Data.Fin.Tuple.Basic
/-!
# Wick strings
Currently this file is only for an example of Wick strings, correpsonding to a
theory with two complex scalar fields. The concepts will however generalize.
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.
@ -19,68 +17,9 @@ term in the ring of operators. This has yet to be implemented.
-/
namespace TwoComplexScalar
namespace Wick
/-- The colors of edges which one can associate with a vertex for a theory
with two complex scalar fields. -/
inductive 𝓔 where
/-- Corresponds to the first complex scalar field flowing out of a vertex. -/
| complexScalarOut₁ : 𝓔
/-- Corresponds to the first complex scalar field flowing into a vertex. -/
| complexScalarIn₁ : 𝓔
/-- Corresponds to the second complex scalar field flowing out of a vertex. -/
| complexScalarOut₂ : 𝓔
/-- Corresponds to the second complex scalar field flowing into a vertex. -/
| complexScalarIn₂ : 𝓔
/-- The map taking each color to it's dual, specifying how we can contract edges. -/
def ξ : 𝓔𝓔
| 𝓔.complexScalarOut₁ => 𝓔.complexScalarIn₁
| 𝓔.complexScalarIn₁ => 𝓔.complexScalarOut₁
| 𝓔.complexScalarOut₂ => 𝓔.complexScalarIn₂
| 𝓔.complexScalarIn₂ => 𝓔.complexScalarOut₂
/-- The function `ξ` is an involution. -/
lemma ξ_involutive : Function.Involutive ξ := by
intro x
match x with
| 𝓔.complexScalarOut₁ => rfl
| 𝓔.complexScalarIn₁ => rfl
| 𝓔.complexScalarOut₂ => rfl
| 𝓔.complexScalarIn₂ => rfl
/-- The vertices associated with two complex scalars.
We call this type, the type of vertex colors. -/
inductive 𝓥 where
| φ₁φ₁φ₂φ₂ : 𝓥
| φ₁φ₁φ₁φ₁ : 𝓥
| φ₂φ₂φ₂φ₂ : 𝓥
/-- To each vertex, the association of the number of edges. -/
@[nolint unusedArguments]
def 𝓥NoEdges : 𝓥 := fun _ => 4
/-- To each vertex, associates the indexing map of half-edges associated with that edge. -/
def 𝓥Edges (v : 𝓥) : Fin (𝓥NoEdges v) → 𝓔 :=
match v with
| 𝓥.φ₁φ₁φ₂φ₂ => fun i =>
match i with
| (0 : Fin 4)=> 𝓔.complexScalarOut₁
| (1 : Fin 4) => 𝓔.complexScalarIn₁
| (2 : Fin 4) => 𝓔.complexScalarOut₂
| (3 : Fin 4) => 𝓔.complexScalarIn₂
| 𝓥.φ₁φ₁φ₁φ₁ => fun i =>
match i with
| (0 : Fin 4)=> 𝓔.complexScalarOut₁
| (1 : Fin 4) => 𝓔.complexScalarIn₁
| (2 : Fin 4) => 𝓔.complexScalarOut₁
| (3 : Fin 4) => 𝓔.complexScalarIn₁
| 𝓥.φ₂φ₂φ₂φ₂ => fun i =>
match i with
| (0 : Fin 4)=> 𝓔.complexScalarOut₂
| (1 : Fin 4) => 𝓔.complexScalarIn₂
| (2 : Fin 4) => 𝓔.complexScalarOut₂
| (3 : Fin 4) => 𝓔.complexScalarIn₂
variable {S : Species}
/-- A helper function for `WickString`. It is used to seperate incoming, vertex, and
outgoing nodes. -/
@ -102,29 +41,29 @@ open WickStringLast
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.
-/
inductive WickString : {ni : } → (i : Fin ni → 𝓔) → {n : } → (c : Fin n → 𝓔) →
{no : } → (o : Fin no → 𝓔) → WickStringLast → Type where
inductive WickString : {ni : } → (i : Fin ni → S.𝓯) → {n : } → (c : Fin n → S.𝓯) →
{no : } → (o : Fin no → S.𝓯) → WickStringLast → Type where
| empty : WickString Fin.elim0 Fin.elim0 Fin.elim0 incoming
| incoming {n ni no : } {i : Fin ni → 𝓔} {c : Fin n → 𝓔}
{o : Fin no → 𝓔} (w : WickString i c o incoming) (e : 𝓔) :
| incoming {n ni no : } {i : Fin ni → S.𝓯} {c : Fin n → S.𝓯}
{o : Fin no → S.𝓯} (w : WickString i c o incoming) (e : S.𝓯) :
WickString (Fin.cons e i) (Fin.cons e c) o incoming
| endIncoming {n ni no : } {i : Fin ni → 𝓔} {c : Fin n → 𝓔}
{o : Fin no → 𝓔} (w : WickString i c o incoming) : WickString i c o vertex
| vertex {n ni no : } {i : Fin ni → 𝓔} {c : Fin n → 𝓔}
{o : Fin no → 𝓔} (w : WickString i c o vertex) (v : 𝓥) :
WickString i (Fin.append (𝓥Edges v) c) o vertex
| endVertex {n ni no : } {i : Fin ni → 𝓔} {c : Fin n → 𝓔}
{o : Fin no → 𝓔} (w : WickString i c o vertex) : WickString i c o outgoing
| outgoing {n ni no : } {i : Fin ni → 𝓔} {c : Fin n → 𝓔}
{o : Fin no → 𝓔} (w : WickString i c o outgoing) (e : 𝓔) :
| 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.𝓯) :
WickString i (Fin.cons e c) (Fin.cons e o) outgoing
| endOutgoing {n ni no : } {i : Fin ni → 𝓔} {c : Fin n → 𝓔}
{o : Fin no → 𝓔} (w : WickString i c o outgoing) : WickString i c o final
| 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
namespace WickString
/-- The number of nodes in a Wick string. This is used to help prove termination. -/
def size {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔} {no : } {o : Fin no → 𝓔}
def size {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯} {no : } {o : Fin no → S.𝓯}
{f : WickStringLast} : WickString i c o f → := fun
| empty => 0
| incoming w e => size w + 1
@ -135,7 +74,7 @@ def size {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔} {no :
| endOutgoing w => size w + 1
/-- The number of vertices in a Wick string. This does NOT include external vertices. -/
def numIntVertex {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔} {no : } {o : Fin no → 𝓔}
def numIntVertex {ni : } {i : Fin ni → S.𝓯} {n : } {c : Fin n → S.𝓯} {no : } {o : Fin no → S.𝓯}
{f : WickStringLast} : WickString i c o f → := fun
| empty => 0
| incoming w e => numIntVertex w
@ -146,8 +85,8 @@ def numIntVertex {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔}
| endOutgoing w => numIntVertex w
/-- The vertices present in a Wick string. This does NOT include external vertices. -/
def intVertex {ni : } {i : Fin ni → 𝓔} {n : } {c : Fin n → 𝓔} {no : } {o : Fin no → 𝓔}
{f : WickStringLast} : (w : WickString i c o f) → Fin w.numIntVertex → 𝓥 := fun
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
| empty => Fin.elim0
| incoming w e => intVertex w
| endIncoming w => intVertex w
@ -191,4 +130,4 @@ informal_lemma loopLevel_fintype where
end WickString
end TwoComplexScalar
end Wick

View file

@ -42,8 +42,13 @@ __BSM physics [🗂️](https://heplean.github.io/HepLean/docs/HepLean/BeyondThe
__Flavor physics [🗂️](https://heplean.github.io/HepLean/docs/HepLean/FlavorPhysics/CKMMatrix/Basic.html):__ Properties of the CKM matrix.
__Perturbation Theory [🗂️](https://heplean.github.io/HepLean/docs/HepLean/PerturbationTheory/Wick/Species.html) [🚧](https://heplean.github.io/HepLean/InformalGraph.html):__ Informal statements relating to Feynman diagrams, Wick contractions, Operator
algebras.
## Associated media and publications
- [📄](https://arxiv.org/abs/2405.08863) Joseph Tooby-Smith, __HepLean: Digitalising high energy physics__, arXiv:2405.08863
- [📄](https://arxiv.org/abs/2405.08863) Joseph Tooby-Smith,
__HepLean: Digitalising high energy physics__, Computer Physics Communications, Volume 308,
2025, 109457, ISSN 0010-4655, https://doi.org/10.1016/j.cpc.2024.109457. \[arXiv:2405.08863\]
- [📄](https://arxiv.org/abs/2411.07667) Joseph Tooby-Smith, __Formalization of physics index notation in Lean 4__, arXiv:2411.07667
- [💻](https://live.lean-lang.org/#code=import%20Mathlib.Tactic.Polyrith%20%0A%0Atheorem%20threeFamily%20(a%20b%20c%20%3A%20)%20(h%20%3A%20a%20%2B%20b%20%2B%20c%20%3D%200)%20(h3%20%3A%20a%20%5E%203%20%2B%20b%20%5E%203%20%2B%20c%20%5E%203%20%3D%200)%20%3A%20%0A%20%20%20%20a%20%3D%200%20%20b%20%3D%200%20%20c%20%3D%200%20%20%3A%3D%20by%20%0A%20%20have%20h1%20%3A%20c%20%3D%20-%20(a%20%2B%20b)%20%3A%3D%20by%20%0A%20%20%20%20linear_combination%20h%20%0A%20%20have%20h4%20%3A%20%203%20*%20a%20*%20b%20*%20c%20%3D%200%20%3A%3D%20by%20%0A%20%20%20%20rw%20%5B←%20h3%2C%20h1%5D%0A%20%20%20%20ring%20%0A%20%20simp%20at%20h4%20%0A%20%20exact%20or_assoc.mp%20h4%0A%20%20%0A) Example code snippet related to Anomaly cancellation conditions.
- [🎥](https://www.youtube.com/watch?v=W2cObnopqas) Seminar recording of "HepLean: Lean and high energy physics" by J. Tooby-Smith