refactor: Delete files and fix imports

This commit is contained in:
jstoobysmith 2024-12-19 11:41:30 +00:00
parent 3c8aaa4ec9
commit aad3afd3a7
10 changed files with 7 additions and 1768 deletions

View file

@ -1,63 +0,0 @@
/-
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 HepLean.PerturbationTheory.Wick.Contract
import HepLean.PerturbationTheory.Wick.Species
/-!
# Feynman diagrams
This file currently contains a lighter implmentation of Feynman digrams than can be found in
`HepLean.PerturbationTheory.FeynmanDiagrams.Basic`. Eventually this will superseed that file.
The implmentation here is done in conjunction with Wicks species etc.
-/
/-! TODO: Remove this namespace-/
namespace LightFeynman
informal_definition FeynmanDiagram where
math :≈ "
Let S be a WickSpecies
A Feynman diagram contains the following data:
- A type of vertices 𝓥 → S.𝓯 ⊕ S.𝓘.
- A type of edges ed : 𝓔 → S.𝓕.
- A type of half-edges 𝓱𝓔 with maps `e : 𝓱𝓔𝓔`, `v : 𝓱𝓔𝓥` and `f : 𝓱𝓔 → S.𝓯`
Subject to the following conditions:
- `𝓱𝓔` is a double cover of `𝓔` through `e`. That is,
for each edge `x : 𝓔` there exists an isomorphism between `i : Fin 2 → e⁻¹ 𝓱𝓔 {x}`.
- These isomorphisms must satisfy `⟦f(i(0))⟧ = ⟦f(i(1))⟧ = ed(e)` and `f(i(0)) = S.ξ (f(i(1)))`.
- For each vertex `ver : 𝓥` there exists an isomorphism between the object (roughly)
`(𝓘Fields v).2` and the pullback of `v` along `ver`."
deps :≈ [``Wick.Species]
informal_definition _root_.Wick.Contract.toFeynmanDiagram where
math :≈ "The Feynman diagram constructed from a complete Wick contraction."
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 :≈ [``Wick.WickContract, ``FeynmanDiagram]
informal_definition FeynmanDiagram.toSimpleGraph where
math :≈ "The simple graph underlying a Feynman diagram."
deps :≈ [``FeynmanDiagram]
informal_definition FeynmanDiagram.IsConnected where
math :≈ "A Feynman diagram is connected if its underlying simple graph is connected."
deps :≈ [``FeynmanDiagram]
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 :≈ [``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
equivalence relation if and only if they have the same Feynman diagram. First step
is to turn these statements into appropriate informal lemmas and definitions. -/
end LightFeynman

View file

@ -1,663 +0,0 @@
/-
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 HepLean.PerturbationTheory.Wick.Species
import HepLean.PerturbationTheory.Wick.Koszul
import HepLean.Lorentz.RealVector.Basic
import HepLean.Mathematics.Fin
import HepLean.SpaceTime.Basic
import HepLean.Mathematics.SuperAlgebra.Basic
import HepLean.Mathematics.List
import HepLean.Meta.Notes.Basic
import Init.Data.List.Sort.Basic
/-!
# Operator algebra
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.
We will formally define the operator ring, in terms of the fields present in the theory.
## Futher reading
- https://physics.stackexchange.com/questions/258718/ and links therein
- Ryan Thorngren (https://physics.stackexchange.com/users/10336/ryan-thorngren), Fermions,
different species and (anti-)commutation rules, URL (version: 2019-02-20) :
https://physics.stackexchange.com/q/461929
- Tong, https://www.damtp.cam.ac.uk/user/tong/qft/qft.pdf
-/
namespace Wick
note r"
<h2>Operator algebra</h2>
Given a Wick Species $S$, we can define the operator algebra of that theory.
The operator algebra is a super-algebra over the complex numbers, which acts on
the Hilbert space of the theory. A super-algebra is an algebra with a $\mathbb{Z}/2$ grading.
To do pertubation theory in a QFT we need a need some basic properties of the operator algebra,
$A$.
<br><br>
For every field $f ∈ \mathcal{f}$, we have a number of families of operators. For every
space-time point $x ∈ \mathbb{R}^4$, we have the operators $ψ(f, x)$ which we decomponse into
a creation and destruction part, $ψ_c(f, x)$ and $ψ_d(f, x)$ respectively. Thus
$ψ(f, x) = ψ_c(f, x) + ψ_d(f, x)$.
For each momentum $p$ we also have the asymptotic states $φ_c(f, p)$ and $φ_d(f, p)$.
<br><br>
If the field $f$ corresponds to a fermion, then all of these operators are homogeneous elements
in the non-identity part of $A$. Conversely, if the field $f$ corresponds to a boson, then all
of these operators are homogeneous elements in the module of $A$ corresponding to
$0 ∈ \mathbb{Z}/2$.
<br><br>
The super-commutator of any of the operators above is in the center of the algebra. Moreover,
the following super-commutators are zero:
<ul>
<li>$[ψ_c(f, x), ψ_c(g, y)] = 0$</li>
<li>$[ψ_d(f, x), ψ_d(g, y)] = 0$</li>
<li>$[φ_c(f, p), φ_c(g, q)] = 0$</li>
<li>$[φ_d(f, p), φ_d(g, q)] = 0$</li>
<li>$[φ_c(f, p), φ_d(g, q)] = 0$ for $f \neq \xi g$</li>
<li>$[φ_d(f, p), ψ_c(g, y)] = 0$ for $f \neq \xi g$</li>
<li>$[φ_c(f, p), ψ_d(g, y)] = 0$ for $f \neq \xi g$</li>
</ul>
<br>
This basic structure constitutes what we call a Wick Algebra:
"
/-- The abstract notion of a operator algebra containing all the necessary ingrediants
to do perturbation theory.
Warning: The definition here is not complete. -/
@[note_attr]
structure WickAlgebra (S : Species) where
/-- The underlying operator algebra. -/
A : Type
/-- The type `A` is a semiring. -/
[A_semiring : Semiring A]
/-- The type `A` is an algebra. -/
[A_algebra : Algebra A]
/-- Position based field operators. -/
ψ : S.𝓯 → SpaceTime → A
/-- Position based constructive operators. -/
ψc : S.𝓯 → SpaceTime → A
/-- Position based destructive operators. -/
ψd : S.𝓯 → SpaceTime → A
/-- Constructive asymptotic operators. -/
φc : S.𝓯 → Lorentz.Contr 3 → A
/-- Distructive asymptotic operators. -/
φd : S.𝓯 → Lorentz.Contr 3 → A
ψc_ψd : ∀ i x, ψc i x + ψd i x = ψ i x
/- Self comutators. -/
ψc_comm_ψc : ∀ i j x y, ψc i x * ψc j y + (S.commFactor i j) • ψc j y * ψc i x = 0
ψd_comm_ψd : ∀ i j x y, ψd i x * ψd j y + (S.commFactor i j) • ψd j y * ψd i x = 0
φc_comm_φc : ∀ i j x y, φc i x * φc j y + (S.commFactor i j) • φc j y * φc i x = 0
φd_comm_φd : ∀ i j x y, φd i x * φd j y + (S.commFactor i j) • φd j y * φd i x = 0
/- Cross comutators. -/
namespace WickAlgebra
variable {S : Species} (𝓞 : WickAlgebra S)
/-- The type `A` of a Wick algebra is a semi-ring. -/
instance : Semiring 𝓞.A := 𝓞.A_semiring
/-- The type `A` of a Wick algebra is an algebra. -/
instance : Algebra 𝓞.A := 𝓞.A_algebra
end WickAlgebra
namespace Species
variable (S : Species)
note r"
<h2>Order</h2>
Suppose we have a type $I$ with a order $r$, a map $g : I \to \mathbb{Z}/2$,
and a map $f : I \to A$ such that $f(i) \in A_{g(i)}$.
Consider the free algebra generated by $I$, which we will denote $A_I$.
The map $f$ can be extended to a map $T_r : A_I \to A$ such that
a monomial $i_1 \cdots i_n$ gets mapped to $(-1)^{K(σ)}f(i_{σ(1)})...f(i_{σ(n)})$ where $σ$ is the
permutation oredering the $i$'s by $r$ (preserving the order of terms which are equal under $r$),
and $K(σ)$ is the Koszul sign factor. (see e.g. PSE:24157)
<br><br>
There are two types $I$ we are intrested in.
"
/-- The indexing set of constructive and destructive position based operators. -/
def ConstDestAlgebra.index : Type := Fin 2 × S.𝓯 × SpaceTime
/-- The free algebra generated by constructive and destructive parts of fields position-based
fields. -/
abbrev ConstDestAlgebra := FreeAlgebra (ConstDestAlgebra.index S)
/-- The indexing set of position based field operators. -/
def FieldAlgebra.index : Type := S.𝓯 × SpaceTime
/-- The free algebra generated by fields. -/
abbrev FieldAlgebra := FreeAlgebra (FieldAlgebra.index S)
namespace ConstDestAlgebra
variable {S} (𝓞 : WickAlgebra S)
/-- The inclusion of constructive and destructive fields into the full operator algebra. -/
def toWickAlgebra : ConstDestAlgebra S →ₐ[] 𝓞.A :=
FreeAlgebra.lift (fun (i, f, x) =>
match i with
| 0 => 𝓞.ψc f x
| 1 => 𝓞.ψd f x)
@[simp]
lemma toWickAlgebra_ι_zero (x : S.𝓯 × SpaceTime) :
toWickAlgebra 𝓞 (FreeAlgebra.ι (0, x)) = 𝓞.ψc x.1 x.2 := by
simp [toWickAlgebra]
@[simp]
lemma toWickAlgebra_ι_one (x : S.𝓯 × SpaceTime) :
toWickAlgebra 𝓞 (FreeAlgebra.ι (1, x)) = 𝓞.ψd x.1 x.2 := by
simp [toWickAlgebra]
/-- The time ordering relation on constructive and destructive operators. -/
def timeOrderRel : index S → index S → Prop := fun x y => y.2.2 0 ≤ x.2.2 0
/-- The normal ordering relation on constructive and destructive operators. -/
def normalOrderRel : index S → index S → Prop := fun x y => x.1 ≤ y.1
/-- The normal ordering relation of constructive and destructive operators is decidable. -/
instance : DecidableRel (@normalOrderRel S) := fun a b => a.1.decLe b.1
noncomputable section
/-- The time ordering relation of constructive and destructive operators is decidable. -/
instance : DecidableRel (@timeOrderRel S) :=
fun a b => Real.decidableLE (b.2.2 0) (a.2.2 0)
/-- The time ordering of constructive and destructive operators. -/
def timeOrder (q : index S → Fin 2) : S.ConstDestAlgebra →ₗ[] S.ConstDestAlgebra :=
koszulOrder timeOrderRel q
/-- The normal ordering of constructive and destructive operators. -/
def normalOrder (q : index S → Fin 2) : S.ConstDestAlgebra →ₗ[] S.ConstDestAlgebra :=
koszulOrder normalOrderRel q
lemma constructive_mul_normalOrder (q : index S → Fin 2) (i : FieldAlgebra.index S)
(A : S.ConstDestAlgebra) :
(FreeAlgebra.ι ((0, i) : index S) : S.ConstDestAlgebra) * normalOrder q A =
normalOrder q ((FreeAlgebra.ι (0, i) : S.ConstDestAlgebra) * A) := by
apply mul_koszulOrder_le
intro j
simp [normalOrderRel]
/-- Contraction of constructive and destructive operators, defined as their time
ordering minus their normal ordering. -/
def contract (q : index S → Fin 2) : S.ConstDestAlgebra →ₗ[] S.ConstDestAlgebra :=
timeOrder q - normalOrder q
end
end ConstDestAlgebra
namespace FieldAlgebra
variable {S} (𝓞 : WickAlgebra S)
/-- The inclusion fo the field algebra into the operator algebra. -/
def toWickAlgebra : FieldAlgebra S →ₐ[] 𝓞.A :=
FreeAlgebra.lift (fun i => 𝓞.ψ i.1 i.2)
@[simp]
lemma toWickAlgebra_ι (i : index S) : toWickAlgebra 𝓞 (FreeAlgebra.ι i) = 𝓞.ψ i.1 i.2 := by
simp [toWickAlgebra]
/-- The time ordering relation in the field algebra. -/
def timeOrderRel : index S → index S → Prop := fun x y => y.2 0 ≤ x.2 0
/-- The time ordering relation in the field algebra is decidable. -/
noncomputable instance : DecidableRel (@timeOrderRel S) :=
fun a b => Real.decidableLE (b.2 0) (a.2 0)
/-- The time ordering in the field algebra. -/
noncomputable def timeOrder (q : index S → Fin 2) : S.FieldAlgebra →ₗ[] S.FieldAlgebra :=
koszulOrder timeOrderRel q
/-- Given a list of fields and a map `f` tell us which field is constructive and
which is destructive, a list of constructive and destructive fields. -/
def listToConstDestList : (l : List (index S)) →
(f : Fin l.length → Fin 2) → List (ConstDestAlgebra.index S)
| [], _ => []
| i :: l, f =>
(f ⟨0, Nat.zero_lt_succ l.length⟩, i.1, i.2) :: listToConstDestList l (f ∘ Fin.succ)
@[simp]
lemma listToConstDestList_length (l : List (index S)) (f : Fin l.length → Fin 2) :
(listToConstDestList l f).length = l.length := by
induction l with
| nil => rfl
| cons i l ih =>
simp only [listToConstDestList, List.length_cons, Fin.zero_eta, Prod.mk.eta, add_left_inj]
rw [ih]
lemma listToConstDestList_insertionSortEquiv (l : List (index S))
(f : Fin l.length → Fin 2) :
(HepLean.List.insertionSortEquiv ConstDestAlgebra.timeOrderRel (listToConstDestList l f))
= (Fin.castOrderIso (by simp)).toEquiv.trans
((HepLean.List.insertionSortEquiv timeOrderRel l).trans
(Fin.castOrderIso (by simp)).toEquiv) := by
induction l with
| nil =>
simp [listToConstDestList, HepLean.List.insertionSortEquiv]
| cons i l ih =>
simp only [listToConstDestList, List.length_cons, Fin.zero_eta, List.insertionSort]
conv_lhs => simp [HepLean.List.insertionSortEquiv]
have h1 (l' : List (ConstDestAlgebra.index S)) :
(HepLean.List.insertEquiv ConstDestAlgebra.timeOrderRel (f ⟨0, by simp⟩, i.1, i.2) l') =
(Fin.castOrderIso (by simp)).toEquiv.trans
((HepLean.List.insertEquiv timeOrderRel (i.1, i.2) (l'.unzip).2).trans
(Fin.castOrderIso (by simp [List.orderedInsert_length])).toEquiv) := by
induction l' with
| nil =>
simp only [List.length_cons, Nat.add_zero, Nat.zero_eq, Fin.zero_eta, List.length_singleton,
List.orderedInsert, HepLean.List.insertEquiv, Fin.castOrderIso_refl,
OrderIso.refl_toEquiv, Equiv.trans_refl]
rfl
| cons j l' ih' =>
by_cases hr : ConstDestAlgebra.timeOrderRel (f ⟨0, by simp⟩, i) j
· rw [HepLean.List.insertEquiv_cons_pos]
· erw [HepLean.List.insertEquiv_cons_pos]
· rfl
· exact hr
· exact hr
· rw [HepLean.List.insertEquiv_cons_neg]
· erw [HepLean.List.insertEquiv_cons_neg]
· simp only [List.length_cons, Nat.add_zero, Nat.zero_eq, Fin.zero_eta,
List.orderedInsert, Prod.mk.eta, Fin.mk_one]
erw [ih']
ext x
simp only [Prod.mk.eta, List.length_cons, Nat.add_zero, Nat.zero_eq, Fin.zero_eta,
HepLean.Fin.equivCons_trans, Nat.succ_eq_add_one,
HepLean.Fin.equivCons_castOrderIso, Equiv.trans_apply, RelIso.coe_fn_toEquiv,
Fin.castOrderIso_apply, Fin.cast_trans, Fin.coe_cast]
congr 2
match x with
| ⟨0, h⟩ => rfl
| ⟨1, h⟩ => rfl
| ⟨Nat.succ (Nat.succ x), h⟩ => rfl
· exact hr
· exact hr
erw [h1]
rw [ih]
simp only [HepLean.Fin.equivCons_trans, Nat.succ_eq_add_one,
HepLean.Fin.equivCons_castOrderIso, List.length_cons, Nat.add_zero, Nat.zero_eq,
Fin.zero_eta]
ext x
conv_rhs => simp [HepLean.List.insertionSortEquiv]
simp only [Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply, Fin.cast_trans,
Fin.coe_cast]
have h2' (i : ConstDestAlgebra.index S) (l' : List (ConstDestAlgebra.index S)) :
(List.orderedInsert ConstDestAlgebra.timeOrderRel i l').unzip.2 =
List.orderedInsert timeOrderRel i.2 l'.unzip.2 := by
induction l' with
| nil =>
simp [HepLean.List.insertEquiv]
| cons j l' ih' =>
by_cases hij : ConstDestAlgebra.timeOrderRel i j
· rw [List.orderedInsert_of_le]
· erw [List.orderedInsert_of_le]
· simp
· exact hij
· exact hij
· simp only [List.orderedInsert, hij, ↓reduceIte, List.unzip_snd, List.map_cons]
have hn : ¬ timeOrderRel i.2 j.2 := hij
simp only [hn, ↓reduceIte, List.cons.injEq, true_and]
simpa using ih'
have h2 (l' : List (ConstDestAlgebra.index S)) :
(List.insertionSort ConstDestAlgebra.timeOrderRel l').unzip.2 =
List.insertionSort timeOrderRel l'.unzip.2 := by
induction l' with
| nil =>
simp [HepLean.List.insertEquiv]
| cons i l' ih' =>
simp only [List.insertionSort, List.unzip_snd]
simp only [List.unzip_snd] at h2'
rw [h2']
congr
simpa using ih'
rw [HepLean.List.insertEquiv_congr _ _ _ (h2 _)]
simp only [List.length_cons, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_trans, Fin.coe_cast]
have h3 : (List.insertionSort timeOrderRel (listToConstDestList l (f ∘ Fin.succ)).unzip.2) =
List.insertionSort timeOrderRel l := by
congr
have h3' (l : List (index S)) (f : Fin l.length → Fin 2) :
(listToConstDestList l (f)).unzip.2 = l := by
induction l with
| nil => rfl
| cons i l ih' =>
simp only [listToConstDestList, List.length_cons, Fin.zero_eta, Prod.mk.eta,
List.unzip_snd, List.map_cons, List.cons.injEq, true_and]
simpa using ih' (f ∘ Fin.succ)
rw [h3']
rw [HepLean.List.insertEquiv_congr _ _ _ h3]
simp only [List.length_cons, Equiv.trans_apply, RelIso.coe_fn_toEquiv, Fin.castOrderIso_apply,
Fin.cast_trans, Fin.cast_eq_self, Fin.coe_cast]
rfl
lemma listToConstDestList_get (l : List (index S)) (f : Fin l.length → Fin 2) :
(listToConstDestList l f).get = (fun i => (f i, l.get i)) ∘ Fin.cast (by simp) := by
induction l with
| nil =>
funext i
exact Fin.elim0 i
| cons i l ih =>
simp only [listToConstDestList, List.length_cons, Fin.zero_eta, List.get_eq_getElem]
funext x
match x with
| ⟨0, h⟩ => rfl
| ⟨x + 1, h⟩ =>
simp only [List.length_cons, List.get_eq_getElem, Prod.mk.eta, List.getElem_cons_succ,
Function.comp_apply, Fin.cast_mk]
change (listToConstDestList l _).get _ = _
rw [ih]
simp
lemma listToConstDestList_timeOrder (l : List (index S)) (f : Fin l.length → Fin 2) :
List.insertionSort ConstDestAlgebra.timeOrderRel (listToConstDestList l f) =
listToConstDestList (List.insertionSort timeOrderRel l)
(f ∘ (HepLean.List.insertionSortEquiv (timeOrderRel) l).symm) := by
let l1 := List.insertionSort (ConstDestAlgebra.timeOrderRel) (listToConstDestList l f)
let l2 := listToConstDestList (List.insertionSort timeOrderRel l)
(f ∘ (HepLean.List.insertionSortEquiv (timeOrderRel) l).symm)
change l1 = l2
have hlen : l1.length = l2.length := by
simp [l1, l2]
have hget : l1.get = l2.get ∘ Fin.cast hlen := by
rw [← HepLean.List.insertionSortEquiv_get]
rw [listToConstDestList_get]
rw [listToConstDestList_get]
rw [← HepLean.List.insertionSortEquiv_get]
funext i
simp only [List.get_eq_getElem, Function.comp_apply, Fin.coe_cast, Fin.cast_trans]
congr 2
· rw [listToConstDestList_insertionSortEquiv]
simp
· rw [listToConstDestList_insertionSortEquiv]
simp
apply List.ext_get hlen
rw [hget]
simp
lemma listToConstDestList_koszulSignInsert (q : index S → Fin 2) (l : List (index S)) (i : index S)
(f : Fin l.length → Fin 2) (a : Fin 2) :
koszulSignInsert ConstDestAlgebra.timeOrderRel (fun i => q i.2) (a, i)
(listToConstDestList l f) = koszulSignInsert timeOrderRel q i l := by
induction l with
| nil =>
simp [listToConstDestList, koszulSignInsert]
| cons j s ih =>
simp only [koszulSignInsert, List.length_cons, Fin.zero_eta, Prod.mk.eta, Fin.isValue]
by_cases hr : ConstDestAlgebra.timeOrderRel (a, i) (f ⟨0, by simp⟩, j)
· rw [if_pos]
· rw [if_pos]
· exact hr
· exact hr
· rw [if_neg]
· nth_rewrite 2 [if_neg]
· rw [ih (f ∘ Fin.succ)]
· exact hr
· exact hr
lemma listToConstDestList_koszulSign (q : index S → Fin 2) (l : List (index S))
(f : Fin l.length → Fin 2) :
koszulSign ConstDestAlgebra.timeOrderRel (fun i => q i.2) (listToConstDestList l f) =
koszulSign timeOrderRel q l := by
induction l with
| nil => rfl
| cons i l ih =>
simp only [koszulSign, List.length_cons, Fin.zero_eta, Prod.mk.eta]
rw [ih]
simp only [mul_eq_mul_right_iff]
apply Or.inl
exact listToConstDestList_koszulSignInsert q l i _ _
/-- The map from the field algebra to the algebra of constructive and destructive fields. -/
def toConstDestAlgebra : FieldAlgebra S →ₐ[] ConstDestAlgebra S :=
FreeAlgebra.lift (fun i => FreeAlgebra.ι (0, i) + FreeAlgebra.ι (1, i))
@[simp]
lemma toConstDestAlgebra_ι (i : index S) : toConstDestAlgebra (FreeAlgebra.ι i) =
FreeAlgebra.ι (0, i) + FreeAlgebra.ι (1, i) := by
simp [toConstDestAlgebra]
lemma toConstDestAlgebra_single (x : ) : (l : FreeMonoid (index S)) →
toConstDestAlgebra (FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x))
= ∑ (f : Fin l.length → Fin 2), FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm
(MonoidAlgebra.single (listToConstDestList l f) x)
| [] => by
simp only [MonoidAlgebra.single, FreeMonoid.length, List.length_nil, Finset.univ_unique,
listToConstDestList, Finset.sum_const, Finset.card_singleton, one_smul]
trans x • 1
· trans toConstDestAlgebra (x • 1)
· congr
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
FreeMonoid.lift, FreeMonoid.prodAux, FreeMonoid.toList, Equiv.coe_fn_mk,
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, MonoidHom.coe_mk, OneHom.coe_mk]
rfl
· simp only [toConstDestAlgebra, Fin.isValue, map_smul, map_one]
· simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply, FreeMonoid.lift,
FreeMonoid.prodAux, FreeMonoid.toList, Equiv.coe_fn_mk, AlgEquiv.ofAlgHom_symm_apply,
MonoidAlgebra.lift_single, MonoidHom.coe_mk, OneHom.coe_mk]
rfl
| i :: l => by
simp only [MonoidAlgebra.single, FreeMonoid.length, List.length_cons, listToConstDestList,
Fin.zero_eta, Prod.mk.eta]
have h1 : FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (Finsupp.single (i :: l) x) =
(FreeAlgebra.ι i) *
(FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (Finsupp.single l x)) := by
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, Algebra.mul_smul_comm]
congr
simp only [FreeMonoid.lift, FreeMonoid.prodAux, FreeMonoid.toList, Equiv.coe_fn_mk,
MonoidHom.coe_mk, OneHom.coe_mk]
change List.foldl (fun x1 x2 => x1 * x2)
(FreeAlgebra.ι i) (List.map (FreeAlgebra.ι ) l) = _
match l with
| [] =>
simp only [List.map_nil, List.foldl_nil, ne_eq, FreeAlgebra.ι_ne_zero, not_false_eq_true,
left_eq_mul₀]
rfl
| x :: l =>
simp only [List.map_cons, List.foldl_cons]
change _ = FreeAlgebra.ι i * List.foldl (fun x1 x2 => x1 * x2) _ _
rw [List.foldl_assoc]
rw [h1]
rw [map_mul]
trans ∑ f : Fin (l.length + 1) → Fin 2, (FreeAlgebra.ι ((f 0, i)) : ConstDestAlgebra S) *
(FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm
(Finsupp.single (listToConstDestList l (f ∘ Fin.succ)) x) : ConstDestAlgebra S)
· rw [← (Fin.consEquiv (n := l.length) (fun _ => Fin 2)).sum_comp
(α := FreeAlgebra (ConstDestAlgebra.index S))]
erw [Finset.sum_product]
simp only [toConstDestAlgebra_ι, Fin.isValue, Fin.consEquiv_apply, Fin.cons_zero,
Fin.sum_univ_two]
rw [← Finset.mul_sum, ← Finset.mul_sum]
erw [← toConstDestAlgebra_single]
rw [add_mul]
· congr
funext f
simp only [FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, Algebra.mul_smul_comm]
congr
simp only [FreeMonoid.lift, FreeMonoid.prodAux, FreeMonoid.toList, Equiv.coe_fn_mk,
MonoidHom.coe_mk, OneHom.coe_mk]
match l with
| [] =>
simp only [listToConstDestList]
change FreeAlgebra.ι (f 0, i) * 1 = _
simp only [mul_one]
rfl
| x :: l =>
simp only [listToConstDestList, List.length_cons, Fin.zero_eta, Function.comp_apply,
Fin.succ_zero_eq_one, Prod.mk.eta]
change FreeAlgebra.ι (f 0, i) * List.foldl _ _ _ = List.foldl _ _ _
simp only [List.map_cons, List.foldl_cons]
haveI : Std.Associative fun
(x1 x2 : FreeAlgebra (ConstDestAlgebra.index S)) => x1 * x2 := by
exact Semigroup.to_isAssociative
refine Eq.symm List.foldl_assoc
lemma toWickAlgebra_factor_toConstDestAlgebra :
toWickAlgebra 𝓞 = (ConstDestAlgebra.toWickAlgebra 𝓞).comp toConstDestAlgebra := by
refine FreeAlgebra.hom_ext ?_
funext i
simp only [Function.comp_apply, toWickAlgebra_ι, ConstDestAlgebra.toWickAlgebra, AlgHom.coe_comp,
toConstDestAlgebra_ι, Fin.isValue, map_add, FreeAlgebra.lift_ι_apply]
split
rename_i x i_1 f x_1 heq
simp_all only [Fin.isValue, Prod.mk.injEq]
obtain ⟨left, right⟩ := heq
subst left right
exact Eq.symm (𝓞.ψc_ψd f x_1)
/-- Time ordering fields and then mapping to constructive and destructive fields is the same as
mapping to constructive and destructive fields and then time ordering. -/
lemma timeOrder_comm_toConstDestAlgebra (q : index S → Fin 2) :
(ConstDestAlgebra.timeOrder (fun i => q i.2)).comp toConstDestAlgebra.toLinearMap =
toConstDestAlgebra.toLinearMap.comp (timeOrder q) := by
let e : S.FieldAlgebra ≃ₗ[] MonoidAlgebra (FreeMonoid (index S)) :=
FreeAlgebra.equivMonoidAlgebraFreeMonoid.toLinearEquiv
apply (LinearEquiv.eq_comp_toLinearMap_iff (e₁₂ := e.symm) _ _).mp
apply MonoidAlgebra.lhom_ext'
intro l
apply LinearMap.ext
intro x
simp only [AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_toLinearMap, LinearMap.coe_comp,
Function.comp_apply, MonoidAlgebra.lsingle_apply, AlgEquiv.toLinearMap_apply,
AlgHom.toLinearMap_apply, toConstDestAlgebra_single, map_sum, timeOrder, koszulOrder_single, e]
simp only [FreeMonoid.length]
let ew := Equiv.piCongrLeft' (fun _ => Fin 2)
(HepLean.List.insertionSortEquiv (timeOrderRel) l)
rw [← ew.sum_comp (α := FreeAlgebra (ConstDestAlgebra.index S))]
congr
funext f
simp only [ConstDestAlgebra.timeOrder, koszulOrder_single, EmbeddingLike.apply_eq_iff_eq]
congr 1
· rw [listToConstDestList_timeOrder]
simp only [ew]
rfl
· simp only [mul_eq_mul_right_iff]
exact Or.inl (listToConstDestList_koszulSign q l f)
/-- The contraction of fields defined as the time order minus normal order once mapped down
to constructive and destructive fields. -/
noncomputable def contract (q : index S → Fin 2) : FieldAlgebra S →ₗ[] ConstDestAlgebra S :=
ConstDestAlgebra.contract (fun i => q i.2) ∘ₗ toConstDestAlgebra.toLinearMap
lemma contract_timeOrderRel (q : index S → Fin 2) {i j : index S} (h : timeOrderRel i j) :
contract q (FreeAlgebra.ι i * FreeAlgebra.ι j) =
superCommute (fun i => q i.2) (FreeAlgebra.ι (1, i)) (FreeAlgebra.ι (0, j)) := by
simp only [contract, LinearMap.coe_comp, Function.comp_apply, AlgHom.toLinearMap_apply, map_mul,
toConstDestAlgebra_ι, Fin.isValue]
trans (ConstDestAlgebra.contract fun i => q i.2)
(FreeAlgebra.ι (0, i) * FreeAlgebra.ι (0, j)
+ FreeAlgebra.ι (0, i) * FreeAlgebra.ι (1, j)
+ FreeAlgebra.ι (1, i) * FreeAlgebra.ι (0, j)
+ FreeAlgebra.ι (1, i) * FreeAlgebra.ι (1, j))
· congr
simp [mul_add, add_mul]
abel
simp only [ConstDestAlgebra.contract, ConstDestAlgebra.timeOrder, ConstDestAlgebra.normalOrder,
Fin.isValue, map_add, LinearMap.sub_apply, koszulOrder_ι_pair, ConstDestAlgebra.normalOrderRel,
le_refl, ↓reduceIte, Fin.zero_le, Fin.le_zero_iff, Nat.reduceAdd, one_ne_zero]
have h1 : ConstDestAlgebra.timeOrderRel (0, i) (0, j) := h
rw [if_pos h1]
repeat rw [if_pos]
abel_nf
simp [koszulSign, koszulSignInsert , ConstDestAlgebra.normalOrderRel]
· exact h1
· exact h1
· exact h1
lemma contract_not_timeOrderRel (q : index S → Fin 2) {i j : index S} (h : ¬ timeOrderRel i j) :
contract q (FreeAlgebra.ι i * FreeAlgebra.ι j) = -
(superCommute (fun i => q i.2) (FreeAlgebra.ι (0, i)) (FreeAlgebra.ι (0, j)) +
superCommute (fun i => q i.2) (FreeAlgebra.ι (1, i)) (FreeAlgebra.ι (1, j)) +
superCommute (fun i => q i.2) (FreeAlgebra.ι (0, i)) (FreeAlgebra.ι (1, j))) := by
simp [contract]
trans (ConstDestAlgebra.contract fun i => q i.2)
(FreeAlgebra.ι (0, i) * FreeAlgebra.ι (0, j)
+ FreeAlgebra.ι (0, i) * FreeAlgebra.ι (1, j)
+ FreeAlgebra.ι (1, i) * FreeAlgebra.ι (0, j)
+ FreeAlgebra.ι (1, i) * FreeAlgebra.ι (1, j))
· congr
simp [mul_add, add_mul]
abel
simp [ConstDestAlgebra.contract, ConstDestAlgebra.timeOrder,
ConstDestAlgebra.normalOrder, ConstDestAlgebra.normalOrderRel, ConstDestAlgebra.normalOrderRel]
rw [if_neg, if_neg, if_neg, if_neg]
simp [koszulSign, koszulSignInsert , ConstDestAlgebra.normalOrderRel]
rw [if_neg]
nth_rewrite 2 [if_neg]
nth_rewrite 3 [if_neg]
nth_rewrite 4 [if_neg]
by_cases hq : q i = 1 ∧ q j = 1
· simp [hq, superCommute]
abel
· simp [hq, superCommute]
abel
· exact h
· exact h
· exact h
· exact h
· exact h
· exact h
· exact h
· exact h
end FieldAlgebra
end Species
informal_definition asymptoicContract where
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 : S.𝓯 × SpaceTime`, the super-commutator [ψ(i), φc(j)]."
informal_definition asymptoicContractAsymptotic where
math :≈ "Given two `i j : S.𝓯 × SpaceTime`, the super-commutator
[φd(i), φc(j)]."
informal_lemma contraction_in_center where
math :≈ "The contraction of two fields is in the center of the algebra."
deps :≈ [``WickAlgebra]
informal_lemma contraction_non_dual_is_zero where
math :≈ "The contraction of two fields is zero if the fields are not dual to each other."
deps :≈ [``WickAlgebra]
informal_lemma timeOrder_single where
math :≈ "The time ordering of a single field is the normal ordering of that field."
proof :≈ "Follows from the definitions."
deps :≈ [``WickAlgebra]
informal_lemma timeOrder_pair where
math :≈ "The time ordering of two fields is the normal ordering of the fields plus the
contraction of the fields."
proof :≈ "Follows from the definition of contraction."
deps :≈ [``WickAlgebra]
informal_definition WickMap where
math :≈ "A linear map `vev` from the Wick algebra `A` to the underlying field such that
`vev(...ψd(t)) = 0` and `vev(ψc(t)...) = 0`."
physics :≈ "An abstraction of the notion of a vacuum expectation value, containing
the necessary properties for lots of theorems to hold."
deps :≈ [``WickAlgebra]
informal_lemma normalOrder_wickMap where
math :≈ "Any normal ordering maps to zero under a Wick map."
deps :≈ [``WickMap]
end Wick

View file

@ -1,726 +0,0 @@
/-
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 HepLean.PerturbationTheory.Wick.String
import Mathlib.Algebra.Order.Ring.Nat
import Mathlib.Data.Fintype.Sum
import Mathlib.Logic.Equiv.Fin
import HepLean.Meta.Notes.Basic
import HepLean.Mathematics.Fin
/-!
# Wick Contract
## Further reading
- https://www.imperial.ac.uk/media/imperial-college/research-centres-and-groups/theoretical-physics/msc/current/qft/handouts/qftwickstheorem.pdf
-/
namespace Wick
variable {S : Species}
note r"
<h2>Wick Contractions</h2>
"
/-- 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 PreContract {α : Type} : (l : List α) → Type where
| nil : PreContract []
| cons (φ : α) {l : List α} (i : Option (Fin l.length)) (p : PreContract l) : PreContract (φ :: l)
namespace PreContract
def example1 : PreContract [0, 1, 2, 3] :=
cons 0 none (cons 1 none (cons 2 (some ⟨0,Nat.zero_lt_succ [].length⟩) (cons 3 none nil)))
def nonContracted {α : Type} : {l : List α} → PreContract l → List α := fun
| nil => []
| cons φ none p => φ :: p.nonContracted
| cons _ (some _) p => p.nonContracted
def remove {α : Type} : {l : List α} → PreContract l → (i : Fin l.length) → PreContract (List.eraseIdx l i) := fun
| nil, i => Fin.elim0 i
| cons φ none p, ⟨0, h⟩ => p
| cons φ none p, ⟨i + 1, h⟩ => cons φ none (p.remove ⟨i,Nat.succ_lt_succ_iff.mp h⟩)
| cons φ (some i) p, ⟨0, h⟩ => p
| cons (l := a :: b :: l) φ (some i) p, ⟨j + 1, h⟩ =>
if ⟨j, Nat.succ_lt_succ_iff.mp h⟩ = i then
cons φ none (p.remove ⟨j, Nat.succ_lt_succ_iff.mp h⟩)
else
cons φ (some (Fin.cast (by
simp [List.length_eraseIdx, h]
rw [if_pos]
simpa using h) (HepLean.Fin.predAboveI ⟨j, Nat.succ_lt_succ_iff.mp h⟩ i ))) (p.remove ⟨j, Nat.succ_lt_succ_iff.mp h⟩)
| cons (l := b ::[ ]) φ (some i) p, ⟨j + 1, h⟩ => cons φ none (p.remove ⟨j, Nat.succ_lt_succ_iff.mp h⟩)
@[nolint unusedArguments]
def length {l : List α} (_ : PreContract l) : := l.length
def dual : {l : List α} → PreContract l → List (Option (Fin l.length)) := fun
| nil => []
| cons (l := l) _ none p => none :: List.map (Option.map Fin.succ) p.dual
| cons (l := l) _ (some i) p => some (Fin.succ i) ::
List.set (List.map (Option.map (Fin.succ)) p.dual) i (some ⟨0, Nat.zero_lt_succ l.length⟩)
lemma dual_length : {l : List α} → (p : PreContract l) → p.dual.length = l.length := fun
| nil => rfl
| cons _ none p => by
simp [dual, length, dual_length p]
| cons _ (some _) p => by simp [dual, dual_length p]
def HasUniqueContr : {l : List α} → PreContract l → Bool
| _, nil => True
| _, cons _ none p => HasUniqueContr p
| _, cons _ (some i) p => p.dual.get ⟨i, by rw [dual_length]; exact i.isLt⟩ = none ∧ HasUniqueContr p
#eval HasUniqueContr example1
@[nolint unusedArguments]
def consDual (φ : α) {l : List α} (p : PreContract l) (i : Option (Fin l.length))
(_ : (Option.map (dual p) i).isNone) : PreContract (φ :: l) := cons φ p i
end PreContract
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 → 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 = 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)
namespace WickContract
/-- The number of nodes of a Wick contraction. -/
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 → 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
simpa [size] using w.size_eq_k
/-- The map giving the vertices on the left-hand-side of a contraction. -/
@[nolint unusedArguments]
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 → 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 = 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) (r : Fin k) :
(contr i j h hilej hb1 hb2i hb2j w).boundFst r.castSucc = w.boundFst r := by
simp only [boundFst, Fin.snoc_castSucc]
@[simp]
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 = 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) :
(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 → 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
intro r s hrs
rcases Fin.eq_castSucc_or_eq_last r with hr | hr
· obtain ⟨r, hr⟩ := hr
subst hr
rcases Fin.eq_castSucc_or_eq_last s with hs | hs
· obtain ⟨s, hs⟩ := hs
subst hs
simp only [boundFst_contr_castSucc]
apply w.boundFst_strictMono _
simpa using hrs
· subst hs
simp only [boundFst_contr_castSucc, boundFst_contr_last]
exact hb1 r
· subst hr
rcases Fin.eq_castSucc_or_eq_last s with hs | hs
· obtain ⟨s, hs⟩ := hs
subst hs
rw [Fin.lt_def] at hrs
simp only [Fin.val_last, Fin.coe_castSucc] at hrs
omega
· subst hs
simp at hrs
/-- The map giving the vertices on the right-hand-side of a contraction. -/
@[nolint unusedArguments]
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 → 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 = 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) (r : Fin k) :
(contr i j h hilej hb1 hb2i hb2j w).boundSnd r.castSucc = w.boundSnd r := by
simp only [boundSnd, Fin.snoc_castSucc]
@[simp]
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 = 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) :
(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 → 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
intro i j _
exact Fin.elim0 i
| contr i j hij hilej hi h2i h2j w => by
intro r s hrs
rcases Fin.eq_castSucc_or_eq_last r with hr | hr
· obtain ⟨r, hr⟩ := hr
subst hr
rcases Fin.eq_castSucc_or_eq_last s with hs | hs
· obtain ⟨s, hs⟩ := hs
subst hs
simp only [boundSnd_contr_castSucc] at hrs
simpa using w.boundSnd_injective hrs
· subst hs
simp only [boundSnd_contr_castSucc, boundSnd_contr_last] at hrs
exact False.elim (h2j r hrs)
· subst hr
rcases Fin.eq_castSucc_or_eq_last s with hs | hs
· obtain ⟨s, hs⟩ := hs
subst hs
simp only [boundSnd_contr_last, boundSnd_contr_castSucc] at hrs
exact False.elim (h2j s hrs.symm)
· subst hs
rfl
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) = 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
· obtain ⟨r, hr⟩ := hr
subst hr
simpa using w.color_boundSnd_eq_dual_boundFst r
· subst hr
simpa using hij
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
| contr i j hij hilej hi _ _ w => fun r => by
rcases Fin.eq_castSucc_or_eq_last r with hr | hr
· obtain ⟨r, hr⟩ := hr
subst hr
simpa using w.boundFst_lt_boundSnd r
· subst hr
simp only [boundFst_contr_last, boundSnd_contr_last]
exact hilej
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
| contr i j _ hilej h1 h2i h2j w => fun r s => by
rcases Fin.eq_castSucc_or_eq_last r with hr | hr
<;> rcases Fin.eq_castSucc_or_eq_last s with hs | hs
· obtain ⟨r, hr⟩ := hr
obtain ⟨s, hs⟩ := hs
subst hr hs
simpa using w.boundFst_neq_boundSnd r s
· obtain ⟨r, hr⟩ := hr
subst hr hs
simp only [Fin.snoc_castSucc, Fin.snoc_last, ne_eq]
have hn := h1 r
omega
· obtain ⟨s, hs⟩ := hs
subst hr hs
simp only [Fin.snoc_last, Fin.snoc_castSucc, ne_eq]
exact (h2i s).symm
· subst hr hs
simp only [Fin.snoc_last, ne_eq]
omega
/-- 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 → 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) :
WickContract str b1' b2' :=
cast (by subst hk; rfl) (hb2 ▸ hb1 ▸ w)
@[simp]
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 → 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 = 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') →
∃ (w' : WickContract str b1 b2), w = castMaps hk' hb1' hb2'
(contr i j h hilej hb1 hb2i hb2j w') := fun
| string => fun hk' => by
simp at hk'
| contr i' j' h' hilej' hb1' hb2i' hb2j' w' => by
intro hk b1 b2 i j h hilej hb1 hb2i hb2j hb1' hb2'
rename_i k' k b1' b2' f
have hk2 : k' = k := Nat.succ_inj'.mp hk
subst hk2
simp_all
have hb2'' : b2 = b2' := by
funext k
trans (@Fin.snoc k' (fun _ => Fin n) b2 j) (Fin.castSucc k)
· simp
· rw [hb2']
simp
have hb1'' : b1 = b1' := by
funext k
trans (@Fin.snoc k' (fun _ => Fin n) b1 i) (Fin.castSucc k)
· simp
· rw [hb1']
simp
have hi : i = i' := by
trans (@Fin.snoc k' (fun _ => Fin n) b1 i) (Fin.last k')
· simp
· rw [hb1']
simp
have hj : j = j' := by
trans (@Fin.snoc k' (fun _ => Fin n) b2 j) (Fin.last k')
· simp
· rw [hb2']
simp
subst hb1'' hb2'' hi hj
simp
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 = 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 → 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
| zero =>
have hb1 : b1 = Fin.elim0 := Subsingleton.elim _ _
have hb2 : b2 = Fin.elim0 := Subsingleton.elim _ _
subst hb1 hb2
match w1, w2 with
| string, string => rfl
| succ k hI =>
match w1, w2 with
| contr i j h hilej hb1 hb2i hb2j w, w2 =>
let ⟨w', hw'⟩ := mem_snoc i j h hilej hb1 hb2i hb2j w2
rw [hw']
apply congrArg (contr i j _ _ _ _ _) (hI w w')
lemma eq_snoc_castSucc {k n : } (b1 : Fin k.succ → Fin n) :
b1 = Fin.snoc (b1 ∘ Fin.castSucc) (b1 (Fin.last k)) := by
funext i
rcases Fin.eq_castSucc_or_eq_last i with h1 | h1
· obtain ⟨i, rfl⟩ := h1
simp
· subst h1
simp
/-- 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 → 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) = S.ξ (c (b1 i)))
(hb1ltb2 : ∀ i, b1 i < b2 i)
(hb1 : StrictMono b1)
(hb1neb2 : ∀ r1 r2, b1 r1 ≠ b2 r2)
(hb2 : Function.Injective b2) :
WickContract str b1 b2 := by
match k with
| 0 =>
refine castMaps ?_ ?_ ?_ string
· rfl
· exact funext (fun i => Fin.elim0 i)
· exact funext (fun i => Fin.elim0 i)
| Nat.succ k =>
refine castMaps rfl (eq_snoc_castSucc b1).symm (eq_snoc_castSucc b2).symm
(contr (b1 (Fin.last k)) (b2 (Fin.last k))
(hi (Fin.last k))
(hb1ltb2 (Fin.last k))
(fun r => hb1 (Fin.castSucc_lt_last r))
(fun r a => hb1neb2 (Fin.last k) r.castSucc a.symm)
(fun r => hb2.eq_iff.mp.mt (Fin.ne_last_of_lt (Fin.castSucc_lt_last r)))
(fromMaps (b1 ∘ Fin.castSucc) (b2 ∘ Fin.castSucc) (fun i => hi (Fin.castSucc i))
(fun i => hb1ltb2 (Fin.castSucc i)) (StrictMono.comp hb1 Fin.strictMono_castSucc)
?_ ?_))
· exact fun r1 r2 => hb1neb2 r1.castSucc r2.castSucc
· exact Function.Injective.comp hb2 (Fin.castSucc_injective k)
/-- 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 → 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)
(fun i => color_boundSnd_eq_dual_boundFst w i.castSucc)
(fun i => boundFst_lt_boundSnd w i.castSucc)
(StrictMono.comp w.boundFst_strictMono Fin.strictMono_castSucc)
(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 → 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 → 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))
(w.color_boundSnd_eq_dual_boundFst (Fin.last k))
(w.boundFst_lt_boundSnd (Fin.last k))
(fun r => w.boundFst_strictMono (Fin.castSucc_lt_last r))
(fun r a => w.boundFst_neq_boundSnd (Fin.last k) r.castSucc a.symm)
(fun r => w.boundSnd_injective.eq_iff.mp.mt (Fin.ne_last_of_lt (Fin.castSucc_lt_last r)))
(dropLast w)) := by
rw [eq_from_maps w]
rfl
/-- Wick contractions of a given Wick string with `k` different contractions. -/
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 → 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
intro x
match x with
| ⟨b1, b2, w⟩ =>
have hb1 : b1 = Fin.elim0 := Subsingleton.elim _ _
have hb2 : b2 = Fin.elim0 := Subsingleton.elim _ _
subst hb1 hb2
simp only [Finset.mem_singleton]
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 → 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 = S.ξ (c i)
hilej : i < j
hb1 : ∀ r, b1 r < i
hb2i : ∀ r, b2 r ≠ i
hb2j : ∀ r, b2 r ≠ j
/-- 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 → 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 = 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 :=
match x with
| ⟨⟨i, j⟩, ⟨h, hilej, hb1, hb2i, hb2j⟩⟩ => ⟨i, j, h, hilej, hb1, hb2i, hb2j⟩
left_inv x := by rfl
right_inv x := by
simp_all only [ne_eq]
obtain ⟨val, property⟩ := x
obtain ⟨fst, snd⟩ := val
obtain ⟨left, right⟩ := property
obtain ⟨left_1, right⟩ := right
obtain ⟨left_2, right⟩ := right
obtain ⟨left_3, right⟩ := right
simp_all only [ne_eq]
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
subst h1 h2
simp only [heq_eq_eq]
exact is_subsingleton.allEq w w'
/-- 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 → 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
| ⟨b1, b2, w⟩ =>
⟨⟨b1 ∘ Fin.castSucc, b2 ∘ Fin.castSucc, dropLast w⟩,
⟨b1 (Fin.last k), b2 (Fin.last k),
w.color_boundSnd_eq_dual_boundFst (Fin.last k),
w.boundFst_lt_boundSnd (Fin.last k),
fun r => w.boundFst_strictMono (Fin.castSucc_lt_last r),
fun r a => w.boundFst_neq_boundSnd (Fin.last k) r.castSucc a.symm,
fun r => w.boundSnd_injective.eq_iff.mp.mt (Fin.ne_last_of_lt (Fin.castSucc_lt_last r))⟩⟩
invFun w :=
match w with
| ⟨⟨b1, b2, w⟩, cp⟩ => ⟨Fin.snoc b1 cp.i, Fin.snoc b2 cp.j,
contr cp.i cp.j cp.h cp.hilej cp.hb1 cp.hb2i cp.hb2j w⟩
left_inv w := by
match w with
| ⟨b1, b2, w⟩ =>
simp only [Nat.succ_eq_add_one, Function.comp_apply]
congr
· exact Eq.symm (eq_snoc_castSucc b1)
· funext b2
congr
exact Eq.symm (eq_snoc_castSucc b1)
· exact Eq.symm (eq_snoc_castSucc b2)
· rw [eq_dropLast_contr w]
simp only [castMaps, Nat.succ_eq_add_one, cast_eq, heq_eqRec_iff_heq, heq_eq_eq,
contr.injEq]
rfl
right_inv w := by
match w with
| ⟨⟨b1, b2, w⟩, cp⟩ =>
simp only [Nat.succ_eq_add_one, Fin.snoc_last, Sigma.mk.inj_iff]
apply And.intro
· congr
· exact Fin.snoc_comp_castSucc
· funext b2
congr
exact Fin.snoc_comp_castSucc
· exact Fin.snoc_comp_castSucc
· exact heq_eq _ _ Fin.snoc_comp_castSucc Fin.snoc_comp_castSucc
· congr
· exact Fin.snoc_comp_castSucc
· exact Fin.snoc_comp_castSucc
· exact heq_eq _ _ Fin.snoc_comp_castSucc Fin.snoc_comp_castSucc
· simp
· simp
· simp
/-- 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 → 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 → 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 → 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 → 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
match x, y with
| Sum.inl x, Sum.inl y =>
simp only [bound_inl] at h
simpa using (StrictMono.injective w.boundFst_strictMono).eq_iff.mp h
| Sum.inr x, Sum.inr y =>
simp only [bound_inr] at h
simpa using w.boundSnd_injective h
| Sum.inl x, Sum.inr y =>
simp only [bound_inl, bound_inr] at h
exact False.elim (w.boundFst_neq_boundSnd x y h)
| Sum.inr x, Sum.inl y =>
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 → 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),
?_⟩
apply Function.Injective.comp (Function.Injective.comp _ finSumFinEquiv.symm.injective)
· exact Fin.cast_injective (Nat.two_mul k)
· exact bound_injection w
/-- 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 → 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 → 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 → 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]
rw [← List.card_toFinset, unboundList]
rw [List.toFinset_filter, List.toFinset_finRange]
have hn := Finset.filter_card_add_filter_neg_card_eq_card (s := Finset.univ)
(fun (i : Fin n) => i ∈ Finset.image w.bound Finset.univ)
have hn' :(Finset.filter (fun i => i ∈ Finset.image w.bound Finset.univ) Finset.univ).card =
(Finset.image w.bound Finset.univ).card := by
refine Finset.card_equiv (Equiv.refl _) fun i => ?_
simp
rw [hn'] at hn
rw [Finset.card_image_of_injective] at hn
simp only [Finset.card_univ, Fintype.card_sum, Fintype.card_fin,
Finset.mem_univ, true_and, Sum.exists, bound_inl, bound_inr, not_or, not_exists] at hn
have hn'' : (Finset.filter (fun a => a ∉ Finset.image w.bound Finset.univ) Finset.univ).card =
n - 2 * k := by
omega
rw [← hn'']
congr
funext x
simp only [ne_eq, Sum.forall, bound_inl, bound_inr, Bool.decide_and, Bool.and_eq_true,
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 → 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 → 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
inj' := by
apply Function.Injective.comp
· rw [← List.nodup_iff_injective_get]
exact w.unboundList_nodup
· exact Fin.cast_injective _
map_rel_iff' := by
refine fun {a b} => StrictMono.le_iff_le ?_
rw [Function.Embedding.coeFn_mk]
apply StrictMono.comp
· exact List.Sorted.get_strictMono w.unboundList_sorted
· exact fun ⦃a b⦄ a => a
informal_lemma level_fintype where
math :≈ "Level is a finite type, as there are only finitely many ways to contract a Wick string."
deps :≈ [``Level]
informal_definition HasEqualTimeContractions where
math :≈ "The condition for a Wick contraction to have two fields contracted
which are of equal time, i.e. come from the same vertex."
deps :≈ [``WickContract]
informal_definition IsConnected where
math :≈ "The condition for a full Wick contraction that for any two vertices
(including external vertices) are connected by contractions."
deps :≈ [``WickContract]
informal_definition HasVacuumContributions where
math :≈ "The condition for a full Wick contraction to have a vacuum contribution."
deps :≈ [``WickContract]
informal_definition IsOneParticleIrreducible where
math :≈ "The condition for a full Wick contraction to be one-particle irreducible."
deps :≈ [``WickContract]
end WickContract
end Wick

View file

@ -194,4 +194,3 @@ lemma superCommute_ofList_ofListM_sum {I : Type} {f : I → Type} [∀ i, Finty
end
end Wick
#min_imports

View file

@ -1,35 +0,0 @@
/-
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 HepLean.PerturbationTheory.Wick.Contract
/-!
# Wick contraction in momentum space
Every complete Wick contraction leads to a function on momenta, following
the Feynman rules.
-/
namespace Wick
informal_definition toMomentumTensorTree where
math :≈ "A function which takes a Wick contraction,
and corresponding momenta, and outputs the corresponding
tensor tree associated with that contraction. The rules for how this is done
is given by the `Feynman rules`.
The appropriate ring to consider here is a ring permitting the abstract notion of a
Dirac delta function. "
ref :≈ "
Some references for Feynman rules are:
- QED Feynman rules: http://hitoshi.berkeley.edu/public_html/129A/point.pdf,
- Weyl Fermions: http://scipp.ucsc.edu/~haber/susybook/feyn115.pdf."
informal_definition toMomentumTensor where
math :≈ "The tensor associated to `toMomentumTensorTree` associated with a Wick contraction,
and corresponding internal momenta, and external momenta."
deps :≈ [``toMomentumTensorTree]
end Wick

View file

@ -1,24 +0,0 @@
/-
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 HepLean.PerturbationTheory.FeynmanDiagrams.Basic
import HepLean.Meta.Informal.Basic
/-!
# Wick contraction in position space
Every complete Wick contraction leads to a function on positions, following
the Feynman rules.
## Further reading
The following reference provides a good resource for Wick contractions of external fields.
- http://www.dylanjtemples.com:82/solutions/QFT_Solution_I-6.pdf
-/
namespace Wick
end Wick

View file

@ -1,81 +0,0 @@
/-
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.Logic.Function.Basic
import HepLean.Meta.Informal.Basic
import HepLean.Meta.Notes.Basic
import HepLean.Lorentz.RealVector.Basic
/-!
# Wick Species
Note: There is very likely a much better name for what we here call a Wick Species.
A Wick Species is a structure containing the basic information needed to write wick contractions
for a theory, and calculate their corresponding Feynman diagrams.
-/
/-! TODO: There should be some sort of notion of a group action on a Wick Species. -/
namespace Wick
note "
<h2>Wick Species</h2>
To do perturbation theory for a quantum field theory, we need a quantum field theory, or
at least enough data from a quantum field theory to write down necessary constructions.
The first bit of data we need is a type of fields `𝓯`. We also need to know what fields
are dual to what other fields, for example in a complex scalar theory `φ` is dual to `φ†`.
We can encode this information in an involution `ξ : 𝓯𝓯`.
<br><br>
The second bit of data we need is how the fields interact with each other. In other words,
a list of interaction vertices `𝓘`, and the type of fields associated to each vertex.
<br><br>
This necessary information to do perturbation theory is encoded in a `Wick Species`, which
we define as:
"
/-- 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. -/
@[note_attr]
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
to Feynman diagrams. -/
𝓯 : Type
/-- The map taking a field operator to its dual operator. -/
ξ : 𝓯𝓯
/-- The condition that `ξ` is an involution. -/
ξ_involutive : Function.Involutive ξ
/-- The color of interaction terms which appear in a theory.
One may wish to call these `vertices`, however we restrict this terminology
to Feynman diagrams. -/
𝓘 : Type
/-- The fields associated to each interaction term. -/
𝓘Fields : 𝓘 → Σ n, Fin n → 𝓯
/-- The map taking a field to `0` if it is a boson and `1` if it is a fermion.
Note that this definition suffers a similar problem to Boolean Blindness. -/
grade : 𝓯 → Fin 2
namespace Species
variable (S : Species)
/-- When commuting two fields `f` and `g`, in the super commuator which is sematically
`[f, g] = f g + c * g f`, this is `c`. -/
def commFactor (f g : S.𝓯) : := - (- 1) ^ (S.grade f * S.grade g : )
informal_definition 𝓕 where
math :≈ "The orbits of the involution `ξ`.
May have to define a multiplicative action of ℤ₂ on `𝓯`, and
take the orbits of this."
physics :≈ "The different types of fields present in a theory."
deps :≈ [``Species]
end Species
end Wick

View file

@ -1,133 +0,0 @@
/-
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 HepLean.Meta.Informal.Basic
import HepLean.PerturbationTheory.Wick.Species
import Mathlib.Data.Fin.Tuple.Basic
/-!
# 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.
-/
namespace Wick
variable {S : Species}
/-- A helper function for `WickString`. It is used to seperate incoming, vertex, and
outgoing nodes. -/
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.
The use of vertices in the Wick string
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
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.
-/
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 → 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 → 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 → 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 → 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
| 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
/-- The number of vertices in a Wick string. This does NOT include external vertices. -/
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
| endIncoming w => numIntVertex w
| vertex w v => numIntVertex w + 1
| endVertex w => numIntVertex w
| outgoing w e => numIntVertex w
| endOutgoing w => numIntVertex w
/-- The vertices present in a Wick string. This does NOT include external vertices. -/
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
| 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."
deps :≈ [``WickString]
informal_definition exponentialPrefactor where
math :≈ "The combinatorical prefactor from the expansion of the
exponential associated with a Wick string."
deps :≈ [``intVertex, ``WickString]
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."
deps :≈ [``intVertex, ``WickString]
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]
end WickString
end Wick

View file

@ -1,34 +0,0 @@
/-
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 HepLean.PerturbationTheory.Wick.Species
/-!
# Wick's theorem
Wick's theorem is related to a result in probability theory called Isserlis' theorem.
-/
namespace Wick
note r"
<h2>Wick's theorem</h2>
"
informal_lemma_note wicks_theorem where
math :≈ "Wick's theorem for fields which are not normally ordered."
informal_lemma wicks_theorem_normal_order where
math :≈ "Wick's theorem for which fields at the same space-time point are normally ordered."
ref :≈ "https://www.physics.purdue.edu/~clarkt/Courses/Physics662/ps/qftch32.pdf"
informal_lemma wicks_theorem_vev where
math :≈ "Wick's theorem in a vev leaving only full contractions of Wick strings left."
informal_lemma wicks_theorem_asymptotic_states where
math :≈ "Wick's theorem for a term in the Dyson series within asymptotic states
leaves only full contractions with the asymptotic states."
end Wick