docs: Docs for Wick contractions
This commit is contained in:
parent
ce9668a3cd
commit
c9607c459f
13 changed files with 123 additions and 64 deletions
|
@ -118,6 +118,9 @@ lemma crPart_mul_normalOrder (φ : 𝓕.FieldOp) (a : 𝓕.FieldOpAlgebra) :
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
|
|
||||||
|
/-- For a field specfication `𝓕`, and `a` and `b` in `𝓕.FieldOpAlgebra` the normal ordering
|
||||||
|
of the super commutator of `a` and `b` vanishes. I.e. `𝓝([a,b]ₛ) = 0`. -/
|
||||||
@[simp]
|
@[simp]
|
||||||
lemma normalOrder_superCommute_eq_zero (a b : 𝓕.FieldOpAlgebra) :
|
lemma normalOrder_superCommute_eq_zero (a b : 𝓕.FieldOpAlgebra) :
|
||||||
𝓝([a, b]ₛ) = 0 := by
|
𝓝([a, b]ₛ) = 0 := by
|
||||||
|
@ -226,7 +229,7 @@ The proof of this result ultimetly goes as follows
|
||||||
- The definition of `normalOrder` is used to rewrite `𝓝(φ₀…φₙ)` as a scalar multiple of
|
- The definition of `normalOrder` is used to rewrite `𝓝(φ₀…φₙ)` as a scalar multiple of
|
||||||
a `ofCrAnList φsn` where `φsn` is the normal ordering of `φ₀…φₙ`.
|
a `ofCrAnList φsn` where `φsn` is the normal ordering of `φ₀…φₙ`.
|
||||||
- `superCommuteF_ofCrAnListF_ofCrAnListF_eq_sum` is used to rewrite the super commutator of `φ`
|
- `superCommuteF_ofCrAnListF_ofCrAnListF_eq_sum` is used to rewrite the super commutator of `φ`
|
||||||
(considered as a list with one lement) with
|
(considered as a list with one element) with
|
||||||
`ofCrAnList φsn` as a sum of supercommutors, one for each element of `φsn`.
|
`ofCrAnList φsn` as a sum of supercommutors, one for each element of `φsn`.
|
||||||
- The fact that super-commutors are in the center of `𝓕.FieldOpAlgebra` is used to rearange terms.
|
- The fact that super-commutors are in the center of `𝓕.FieldOpAlgebra` is used to rearange terms.
|
||||||
- Properties of ordered lists, and `normalOrderSign_eraseIdx` is then used to complete the proof.
|
- Properties of ordered lists, and `normalOrderSign_eraseIdx` is then used to complete the proof.
|
||||||
|
|
|
@ -25,10 +25,14 @@ variable {𝓕 : FieldSpecification}
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
/--
|
/-- For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||||
Let `c` be a Wick Contraction for `φs := φ₀φ₁…φₙ`.
|
`𝓕.FieldOp`, and a `i ≤ φs.length` the following relation holds
|
||||||
We have (roughly) `𝓝ᶠ([φsΛ ↩Λ φ i none]ᵘᶜ) = s • 𝓝ᶠ(φ :: [φsΛ]ᵘᶜ)`
|
|
||||||
Where `s` is the exchange sign for `φ` and the uncontracted fields in `φ₀φ₁…φᵢ₋₁`.
|
`𝓝([φsΛ ↩Λ φ i none]ᵘᶜ) = s • 𝓝(φ :: [φsΛ]ᵘᶜ)`
|
||||||
|
|
||||||
|
where `s` is the exchange sign for `φ` and the uncontracted fields in `φ₀…φᵢ₋₁`.
|
||||||
|
|
||||||
|
The prove of this result ultimately a consequence of `normalOrder_superCommute_eq_zero`.
|
||||||
-/
|
-/
|
||||||
lemma normalOrder_uncontracted_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
lemma normalOrder_uncontracted_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
(i : Fin φs.length.succ) (φsΛ : WickContraction φs.length) :
|
(i : Fin φs.length.succ) (φsΛ : WickContraction φs.length) :
|
||||||
|
@ -95,10 +99,12 @@ lemma normalOrder_uncontracted_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp
|
||||||
rw [insertAndContract_uncontractedList_none_map]
|
rw [insertAndContract_uncontractedList_none_map]
|
||||||
|
|
||||||
/--
|
/--
|
||||||
Let `c` be a Wick Contraction for `φ₀φ₁…φₙ`.
|
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||||
We have (roughly) `N(c ↩Λ φ i k).uncontractedList`
|
`𝓕.FieldOp`, a `i ≤ φs.length` and a `k` in `φsΛ.uncontracted`, then
|
||||||
is equal to `N((c.uncontractedList).eraseIdx k')`
|
`𝓝([φsΛ ↩Λ φ i none]ᵘᶜ)` is equal to the normal ordering of `[φsΛ]ᵘᶜ` with the `FieldOp`
|
||||||
where `k'` is the position in `c.uncontractedList` corresponding to `k`.
|
corresponding to `k` removed.
|
||||||
|
|
||||||
|
The proof of this result ultimately a consequence of definitions.
|
||||||
-/
|
-/
|
||||||
lemma normalOrder_uncontracted_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
lemma normalOrder_uncontracted_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
(i : Fin φs.length.succ) (φsΛ : WickContraction φs.length) (k : φsΛ.uncontracted) :
|
(i : Fin φs.length.succ) (φsΛ : WickContraction φs.length) (k : φsΛ.uncontracted) :
|
||||||
|
|
|
@ -14,13 +14,18 @@ open FieldSpecification
|
||||||
variable {𝓕 : FieldSpecification}
|
variable {𝓕 : FieldSpecification}
|
||||||
|
|
||||||
/--
|
/--
|
||||||
Given a natural number `n` corresponding to the number of fields, a Wick contraction
|
Given a natural number `n`, which will correspond to the number of fields needing
|
||||||
is a finite set of pairs of `Fin n`, such that no element of `Fin n` occurs in more then one pair.
|
contracting, a Wick contraction
|
||||||
|
is a finite set of pairs of `Fin n` (numbers `0`, …, `n-1`), such that no
|
||||||
|
element of `Fin n` occurs in more then one pair. The pairs are the positions of fields we
|
||||||
|
'contract' together.
|
||||||
|
|
||||||
For example for `n = 3` there are `4` Wick contractions:
|
For example for `n = 3` there are `4` Wick contractions:
|
||||||
- `∅`, corresponding to the case where no fields are contracted.
|
- `∅`, corresponding to the case where no fields are contracted.
|
||||||
- `{{0, 1}}`, corresponding to the case where the field at position `0` and `1` are contracted.
|
- `{{0, 1}}`, corresponding to the case where the field at position `0` and `1` are contracted.
|
||||||
- `{{0, 2}}`, corresponding to the case where the field at position `0` and `2` are contracted.
|
- `{{0, 2}}`, corresponding to the case where the field at position `0` and `2` are contracted.
|
||||||
- `{{1, 2}}`, corresponding to the case where the field at position `1` and `2` are contracted.
|
- `{{1, 2}}`, corresponding to the case where the field at position `1` and `2` are contracted.
|
||||||
|
|
||||||
For `n=4` some possible Wick contractions are
|
For `n=4` some possible Wick contractions are
|
||||||
- `∅`, corresponding to the case where no fields are contracted.
|
- `∅`, corresponding to the case where no fields are contracted.
|
||||||
- `{{0, 1}, {2, 3}}`, corresponding to the case where the field at position `0` and `1` are
|
- `{{0, 1}, {2, 3}}`, corresponding to the case where the field at position `0` and `1` are
|
||||||
|
@ -37,6 +42,12 @@ namespace WickContraction
|
||||||
variable {n : ℕ} (c : WickContraction n)
|
variable {n : ℕ} (c : WickContraction n)
|
||||||
open HepLean.List
|
open HepLean.List
|
||||||
|
|
||||||
|
remark contraction_notation := "Given a field specification `𝓕`, and a list `φs`
|
||||||
|
of `𝓕.FieldOp`, a Wick contraction of `φs` will mean a Wick contraction in
|
||||||
|
`WickContraction φs.length`. The notation `φsΛ` will be used for such contractions.
|
||||||
|
The terminology that `φsΛ` contracts pairs within of `φs` will also be used, even though
|
||||||
|
`φsΛ` is really contains positions of `φs`."
|
||||||
|
|
||||||
/-- Wick contractions are decidable. -/
|
/-- Wick contractions are decidable. -/
|
||||||
instance : DecidableEq (WickContraction n) := Subtype.instDecidableEq
|
instance : DecidableEq (WickContraction n) := Subtype.instDecidableEq
|
||||||
|
|
||||||
|
@ -520,8 +531,11 @@ lemma prod_finset_eq_mul_fst_snd (c : WickContraction n) (a : c.1)
|
||||||
rw [← (c.contractEquivFinTwo a).symm.prod_comp]
|
rw [← (c.contractEquivFinTwo a).symm.prod_comp]
|
||||||
simp [contractEquivFinTwo]
|
simp [contractEquivFinTwo]
|
||||||
|
|
||||||
/-- A Wick contraction associated with a list of states is said to be `GradingCompliant` if in any
|
/-- For a field specification `𝓕`, `φs` a list of `𝓕.FieldOp` and a Wick contraction
|
||||||
contracted pair of states they are either both fermionic or both bosonic. -/
|
`φsΛ` of `φs`, the Wick contraction `φsΛ` is said to be `GradingCompliant` if
|
||||||
|
for every pair in `φsΛ` the contracted fields are either both `fermionic` or both `bosonic`.
|
||||||
|
I.e. in a `GradingCompliant` Wick contraction no contractions occur between `fermionic` and
|
||||||
|
`bosonic` fields. -/
|
||||||
def GradingCompliant (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) :=
|
def GradingCompliant (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) :=
|
||||||
∀ (a : φsΛ.1), (𝓕 |>ₛ φs[φsΛ.fstFieldOfContract a]) = (𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
∀ (a : φsΛ.1), (𝓕 |>ₛ φs[φsΛ.fstFieldOfContract a]) = (𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||||
|
|
||||||
|
|
|
@ -236,9 +236,9 @@ def cardFun : ℕ → ℕ
|
||||||
| 1 => 1
|
| 1 => 1
|
||||||
| Nat.succ (Nat.succ n) => cardFun (Nat.succ n) + (n + 1) * cardFun n
|
| Nat.succ (Nat.succ n) => cardFun (Nat.succ n) + (n + 1) * cardFun n
|
||||||
|
|
||||||
/-- The number of Wick contractions for `n : ℕ` fields, i.e. the cardinality of
|
/-- The number of Wick contractions in `WickContraction n` is equal to the terms in
|
||||||
`WickContraction n`, is equal to the terms in
|
Online Encyclopedia of Integer Sequences (OEIS) A000085. That is:
|
||||||
Online Encyclopedia of Integer Sequences (OEIS) A000085. -/
|
1, 1, 2, 4, 10, 26, 76, 232, 764, 2620, 9496, ... -/
|
||||||
theorem card_eq_cardFun : (n : ℕ) → Fintype.card (WickContraction n) = cardFun n
|
theorem card_eq_cardFun : (n : ℕ) → Fintype.card (WickContraction n) = cardFun n
|
||||||
| 0 => by decide
|
| 0 => by decide
|
||||||
| 1 => by decide
|
| 1 => by decide
|
||||||
|
|
|
@ -24,15 +24,19 @@ open HepLean.Fin
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
/-- Given a Wick contraction `φsΛ` associated to a list `φs`,
|
/-- Given a Wick contraction `φsΛ` for a list `φs` of `𝓕.FieldOp`,
|
||||||
a position `i : Fin φs.lengthsucc`, an element `φ`, and an optional uncontracted element
|
a `𝓕.FieldOp` `φ`, an `i ≤ φs.length` and a `j` which is either `none` or
|
||||||
`j : Option (φsΛ.uncontracted)` of `φsΛ`.
|
some element of `φsΛ.uncontracted`, the new Wick contraction
|
||||||
The Wick contraction `φsΛ.insertAndContract φ i j` is defined to be the Wick contraction
|
`φsΛ.insertAndContract φ i j` is defined by inserting `φ` into `φs` after
|
||||||
associated with `(φs.insertIdx i φ)` formed by 'inserting' `φ` into `φs` after the first `i`
|
the first `i`-elements and moving the values representing the contracted pairs in `φsΛ`
|
||||||
elements and contracting `φ` optionally with `j`.
|
accordingly.
|
||||||
|
If `j` is not `none`, but rather `some j`, to this contraction is added the contraction
|
||||||
|
of `φ` (at position `i`) with the new position of `j` after `φ` is added.
|
||||||
|
|
||||||
The notation `φsΛ ↩Λ φ i j` is used to denote `φsΛ.insertAndContract φ i j`. Thus,
|
In other words, `φsΛ.insertAndContract φ i j` is formed by adding `φ` to `φs` at position `i`,
|
||||||
`φsΛ ↩Λ φ i none` indicates the case when we insert `φ` into `φs` but do not contract it. -/
|
and contracting `φ` with the field orginally at position `j` if `j` is not none.
|
||||||
|
|
||||||
|
The notation `φsΛ ↩Λ φ i j` is used to denote `φsΛ.insertAndContract φ i j`. -/
|
||||||
def insertAndContract {φs : List 𝓕.FieldOp} (φ : 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
def insertAndContract {φs : List 𝓕.FieldOp} (φ : 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||||
(i : Fin φs.length.succ) (j : Option φsΛ.uncontracted) :
|
(i : Fin φs.length.succ) (j : Option φsΛ.uncontracted) :
|
||||||
WickContraction (φs.insertIdx i φ).length :=
|
WickContraction (φs.insertIdx i φ).length :=
|
||||||
|
|
|
@ -23,9 +23,9 @@ variable {n : ℕ} (c : WickContraction n)
|
||||||
open HepLean.List
|
open HepLean.List
|
||||||
open FieldOpAlgebra
|
open FieldOpAlgebra
|
||||||
|
|
||||||
/-- Given a Wick contraction `φsΛ` on `φs` and a Wick contraction `φsucΛ` on the uncontracted fields
|
/-- Given a list `φs` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs` and a Wick contraction
|
||||||
within `φsΛ`, a Wick contraction on `φs`consisting of the contractins in `φsΛ` and
|
`φsucΛ` of `[φsΛ]ᵘᶜ`, `join φsΛ φsucΛ` is defined as the Wick contraction of `φs` consisting of
|
||||||
the contractions in `φsucΛ`. -/
|
the contractions in `φsΛ` and those in `φsucΛ`. -/
|
||||||
def join {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
def join {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) : WickContraction φs.length :=
|
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) : WickContraction φs.length :=
|
||||||
⟨φsΛ.1 ∪ φsucΛ.1.map (Finset.mapEmbedding uncontractedListEmd).toEmbedding, by
|
⟨φsΛ.1 ∪ φsucΛ.1.map (Finset.mapEmbedding uncontractedListEmd).toEmbedding, by
|
||||||
|
|
|
@ -28,10 +28,11 @@ def signFinset (c : WickContraction n) (i1 i2 : Fin n) : Finset (Fin n) :=
|
||||||
Finset.univ.filter (fun i => i1 < i ∧ i < i2 ∧
|
Finset.univ.filter (fun i => i1 < i ∧ i < i2 ∧
|
||||||
(c.getDual? i = none ∨ ∀ (h : (c.getDual? i).isSome), i1 < (c.getDual? i).get h))
|
(c.getDual? i = none ∨ ∀ (h : (c.getDual? i).isSome), i1 < (c.getDual? i).get h))
|
||||||
|
|
||||||
/-- Given a Wick contraction `φsΛ` associated with a list of states `φs`
|
/-- For a list `φs` of `𝓕.FieldOp`, and a Wick contraction `φsΛ` of `φs`,
|
||||||
the sign associated with `φsΛ` is the sign corresponding to the number
|
the complex number `φsΛ.sign` is defined to be the sign (`1` or `-1`) corresponding
|
||||||
of fermionic-fermionic exchanges one must do to put elements in contracted pairs
|
to the number of `fermionic`-`fermionic` exchanges that must done to put
|
||||||
of `φsΛ` next to each other. -/
|
contracted pairs with `φsΛ` next to one another, starting from the contracted pair
|
||||||
|
whose first element occurs at the left-most position. -/
|
||||||
def sign (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) : ℂ :=
|
def sign (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) : ℂ :=
|
||||||
∏ (a : φsΛ.1), 𝓢(𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a],
|
∏ (a : φsΛ.1), 𝓢(𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a],
|
||||||
𝓕 |>ₛ ⟨φs.get, φsΛ.signFinset (φsΛ.fstFieldOfContract a) (φsΛ.sndFieldOfContract a)⟩)
|
𝓕 |>ₛ ⟨φs.get, φsΛ.signFinset (φsΛ.fstFieldOfContract a) (φsΛ.sndFieldOfContract a)⟩)
|
||||||
|
|
|
@ -238,10 +238,14 @@ lemma signInsertNone_eq_filterset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
simp [h]
|
simp [h]
|
||||||
· exact hG
|
· exact hG
|
||||||
|
|
||||||
/-- For `φsΛ` a grading compliant Wick contraction, and `i : Fin φs.length.succ` we have
|
/-- For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a graded compliant Wick contraction `φsΛ` of `φs`,
|
||||||
|
an `i ≤ φs.length` and a `φ` in `𝓕.FieldOp`, the following relation holds
|
||||||
`(φsΛ ↩Λ φ i none).sign = s * φsΛ.sign`
|
`(φsΛ ↩Λ φ i none).sign = s * φsΛ.sign`
|
||||||
where `s` is the sign got by moving `φ` through the elements of `φ₀…φᵢ₋₁` which
|
where `s` is the sign got by moving `φ` through the elements of `φ₀…φᵢ₋₁` which
|
||||||
are contracted. -/
|
are contracted with some element.
|
||||||
|
|
||||||
|
The proof of this result involves a careful consideration of the contributions of different
|
||||||
|
`FieldOp`s in `φs` to the sign of `φsΛ ↩Λ φ i none`. -/
|
||||||
lemma sign_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
lemma sign_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
|
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
|
||||||
(φsΛ ↩Λ φ i none).sign = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, Finset.univ.filter
|
(φsΛ ↩Λ φ i none).sign = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, Finset.univ.filter
|
||||||
|
@ -250,6 +254,11 @@ lemma sign_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
rw [signInsertNone_eq_filterset]
|
rw [signInsertNone_eq_filterset]
|
||||||
exact hG
|
exact hG
|
||||||
|
|
||||||
|
/-- For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a graded compliant Wick contraction `φsΛ` of `φs`,
|
||||||
|
and a `φ` in `𝓕.FieldOp`, the following relation holds
|
||||||
|
`(φsΛ ↩Λ φ 0 none).sign = φsΛ.sign`.
|
||||||
|
|
||||||
|
This is a direct corollary of `sign_insert_none`. -/
|
||||||
lemma sign_insert_none_zero (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
lemma sign_insert_none_zero (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
(φsΛ : WickContraction φs.length) : (φsΛ ↩Λ φ 0 none).sign = φsΛ.sign := by
|
(φsΛ : WickContraction φs.length) : (φsΛ ↩Λ φ 0 none).sign = φsΛ.sign := by
|
||||||
rw [sign_insert_none_eq_signInsertNone_mul_sign]
|
rw [sign_insert_none_eq_signInsertNone_mul_sign]
|
||||||
|
|
|
@ -855,13 +855,15 @@ lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.FieldOp) (φs :
|
||||||
exact fun h => lt_of_le_of_ne h (Fin.succAbove_ne i ((φsΛ.getDual? j).get hj))
|
exact fun h => lt_of_le_of_ne h (Fin.succAbove_ne i ((φsΛ.getDual? j).get hj))
|
||||||
|
|
||||||
/--
|
/--
|
||||||
For `k < i`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
|
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||||
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ`,
|
`𝓕.FieldOp`, a `i ≤ φs.length` and a `k` in `φsΛ.uncontracted` such that `k<i`,
|
||||||
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
|
the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
|
||||||
|
- the sign associated with moving `φ` through the `φsΛ`-uncontracted `FieldOp` in `φ₀…φₖ`,
|
||||||
|
- the sign associated with moving `φ` through all `FieldOp` in `φ₀…φᵢ₋₁`,
|
||||||
- the sign of `φsΛ`.
|
- the sign of `φsΛ`.
|
||||||
|
|
||||||
The proof of this result involves a careful consideration of the contributions of different
|
The proof of this result involves a careful consideration of the contributions of different
|
||||||
fields in `φs` to the sign of `φsΛ ↩Λ φ i (some k)`.
|
`FieldOp` in `φs` to the sign of `φsΛ ↩Λ φ i (some k)`.
|
||||||
-/
|
-/
|
||||||
lemma sign_insert_some_of_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
lemma sign_insert_some_of_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||||
|
@ -878,13 +880,15 @@ lemma sign_insert_some_of_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
simp
|
simp
|
||||||
|
|
||||||
/--
|
/--
|
||||||
For `i ≤ k`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
|
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||||
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ₋₁`,
|
`𝓕.FieldOp`, a `i ≤ φs.length` and a `k` in `φsΛ.uncontracted` such that `i ≤ k`,
|
||||||
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
|
the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
|
||||||
|
- the sign associated with moving `φ` through the `φsΛ`-uncontracted `FieldOp` in `φ₀…φₖ₋₁`,
|
||||||
|
- the sign associated with moving `φ` through all the `FieldOp` in `φ₀…φᵢ₋₁`,
|
||||||
- the sign of `φsΛ`.
|
- the sign of `φsΛ`.
|
||||||
|
|
||||||
The proof of this result involves a careful consideration of the contributions of different
|
The proof of this result involves a careful consideration of the contributions of different
|
||||||
fields in `φs` to the sign of `φsΛ ↩Λ φ i (some k)`.
|
`FieldOp` in `φs` to the sign of `φsΛ ↩Λ φ i (some k)`.
|
||||||
-/
|
-/
|
||||||
lemma sign_insert_some_of_not_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
lemma sign_insert_some_of_not_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||||
|
@ -900,6 +904,15 @@ lemma sign_insert_some_of_not_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
rw [mul_comm, ← mul_assoc]
|
rw [mul_comm, ← mul_assoc]
|
||||||
simp
|
simp
|
||||||
|
|
||||||
|
/--
|
||||||
|
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||||
|
`𝓕.FieldOp`, and a `k` in `φsΛ.uncontracted`,
|
||||||
|
the sign of `φsΛ ↩Λ φ 0 (some k)` is equal to the product of
|
||||||
|
- the sign associated with moving `φ` through the `φsΛ`-uncontracted `FieldOp` in `φ₀…φₖ₋₁`,
|
||||||
|
- the sign of `φsΛ`.
|
||||||
|
|
||||||
|
This is a direct corollary of `sign_insert_some_of_not_lt`.
|
||||||
|
-/
|
||||||
lemma sign_insert_some_zero (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
lemma sign_insert_some_zero (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||||
(φsΛ : WickContraction φs.length) (k : φsΛ.uncontracted)
|
(φsΛ : WickContraction φs.length) (k : φsΛ.uncontracted)
|
||||||
(hn : GradingCompliant φs φsΛ ∧ (𝓕|>ₛφ) = 𝓕|>ₛφs[k.1]) :
|
(hn : GradingCompliant φs φsΛ ∧ (𝓕|>ₛφ) = 𝓕|>ₛφs[k.1]) :
|
||||||
|
|
|
@ -417,11 +417,13 @@ lemma join_sign_induction {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs
|
||||||
apply sign_congr
|
apply sign_congr
|
||||||
exact join_uncontractedListGet (singleton hij) φsucΛ'
|
exact join_uncontractedListGet (singleton hij) φsucΛ'
|
||||||
|
|
||||||
/-- Let `φsΛ` be a grading compliant Wick contraction for `φs` and
|
/-- For a list `φs` of `𝓕.FieldOp`, a grading compliant Wick contraction `φsΛ` of `φs`,
|
||||||
`φsucΛ` a Wick contraction for `[φsΛ]ᵘᶜ`. Then `(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign`.
|
and a Wick contraction `φsucΛ` of `[φsΛ]ᵘᶜ`, the following relation holds
|
||||||
|
`(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign`.
|
||||||
|
|
||||||
This lemma manifests the fact that it does not matter which order contracted pairs are brought
|
In `φsΛ.sign` the sign is determined by starting with the contracted pair
|
||||||
together when defining the sign of a contraction. -/
|
whose first element occurs at the left-most position. This lemma manifests that
|
||||||
|
choice does not matter, and that contracted pairs can be brought together in any order. -/
|
||||||
lemma join_sign {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
lemma join_sign {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (hc : φsΛ.GradingCompliant) :
|
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (hc : φsΛ.GradingCompliant) :
|
||||||
(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign := by
|
(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign := by
|
||||||
|
|
|
@ -16,7 +16,8 @@ namespace WickContraction
|
||||||
variable {n : ℕ} (c : WickContraction n)
|
variable {n : ℕ} (c : WickContraction n)
|
||||||
open HepLean.List
|
open HepLean.List
|
||||||
|
|
||||||
/-- Given a Wick contraction, the finset of elements of `Fin n` which are not contracted. -/
|
/-- For a Wick contraction `c`, `c.uncontracted` is defined as the finset of elements of `Fin n`
|
||||||
|
which are not in any contracted pair. -/
|
||||||
def uncontracted : Finset (Fin n) := Finset.filter (fun i => c.getDual? i = none) (Finset.univ)
|
def uncontracted : Finset (Fin n) := Finset.filter (fun i => c.getDual? i = none) (Finset.univ)
|
||||||
|
|
||||||
lemma congr_uncontracted {n m : ℕ} (c : WickContraction n) (h : n = m) :
|
lemma congr_uncontracted {n m : ℕ} (c : WickContraction n) (h : n = m) :
|
||||||
|
|
|
@ -314,8 +314,11 @@ lemma take_uncontractedIndexEquiv_symm (k : c.uncontracted) :
|
||||||
|
|
||||||
-/
|
-/
|
||||||
|
|
||||||
/-- Given a Wick Contraction `φsΛ` for a list of states `φs`. The list of uncontracted
|
/-- Given a Wick Contraction `φsΛ` of a list `φs` of `𝓕.FieldOp`. The list
|
||||||
states in `φs`. -/
|
`φsΛ.uncontractedListGet` of `𝓕.FieldOp` is defined as the list `φs` with
|
||||||
|
all contracted positions removed, leaving the uncontracted `𝓕.FieldOp`.
|
||||||
|
|
||||||
|
The notation `[φsΛ]ᵘᶜ` is used for `φsΛ.uncontractedListGet`. -/
|
||||||
def uncontractedListGet {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
def uncontractedListGet {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||||
List 𝓕.FieldOp := φsΛ.uncontractedList.map φs.get
|
List 𝓕.FieldOp := φsΛ.uncontractedList.map φs.get
|
||||||
|
|
||||||
|
|
|
@ -192,30 +192,33 @@ def perturbationTheory : Note where
|
||||||
.name ``FieldSpecification.normalOrderSign_eraseIdx .complete,
|
.name ``FieldSpecification.normalOrderSign_eraseIdx .complete,
|
||||||
.name ``FieldSpecification.FieldOpFreeAlgebra.normalOrderF .complete,
|
.name ``FieldSpecification.FieldOpFreeAlgebra.normalOrderF .complete,
|
||||||
.name ``FieldSpecification.FieldOpAlgebra.normalOrder .complete,
|
.name ``FieldSpecification.FieldOpAlgebra.normalOrder .complete,
|
||||||
|
.name ``FieldSpecification.FieldOpAlgebra.normalOrder_superCommute_eq_zero .complete,
|
||||||
.name ``FieldSpecification.FieldOpAlgebra.ofCrAnOp_superCommute_normalOrder_ofCrAnList_sum .complete,
|
.name ``FieldSpecification.FieldOpAlgebra.ofCrAnOp_superCommute_normalOrder_ofCrAnList_sum .complete,
|
||||||
.name ``FieldSpecification.FieldOpAlgebra.ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum .complete,
|
.name ``FieldSpecification.FieldOpAlgebra.ofFieldOp_mul_normalOrder_ofFieldOpList_eq_sum .complete,
|
||||||
.h1 "Wick Contractions",
|
.h1 "Wick Contractions",
|
||||||
.h2 "Definition",
|
.h2 "Definition",
|
||||||
.name ``WickContraction .incomplete,
|
.name ``WickContraction .complete,
|
||||||
.name ``WickContraction.GradingCompliant .incomplete,
|
.name `WickContraction.contraction_notation .complete,
|
||||||
|
.name ``WickContraction.GradingCompliant .complete,
|
||||||
.h2 "Aside: Cardinality",
|
.h2 "Aside: Cardinality",
|
||||||
.name ``WickContraction.card_eq_cardFun .incomplete,
|
.name ``WickContraction.card_eq_cardFun .complete,
|
||||||
|
.h2 "Uncontracted elements",
|
||||||
|
.name ``WickContraction.uncontracted .complete,
|
||||||
|
.name ``WickContraction.uncontractedListGet .complete,
|
||||||
.h2 "Constructors",
|
.h2 "Constructors",
|
||||||
.p "There are a number of ways to construct a Wick contraction from
|
.name ``WickContraction.insertAndContract .complete,
|
||||||
other Wick contractions or single contractions.",
|
.name ``WickContraction.join .complete,
|
||||||
.name ``WickContraction.insertAndContract .incomplete,
|
|
||||||
.name ``WickContraction.join .incomplete,
|
|
||||||
.h2 "Sign",
|
.h2 "Sign",
|
||||||
.name ``WickContraction.sign .incomplete,
|
.name ``WickContraction.sign .complete,
|
||||||
.name ``WickContraction.join_sign .incomplete,
|
.name ``WickContraction.join_sign .complete,
|
||||||
.name ``WickContraction.sign_insert_none .incomplete,
|
.name ``WickContraction.sign_insert_none .complete,
|
||||||
.name ``WickContraction.sign_insert_none_zero .incomplete,
|
.name ``WickContraction.sign_insert_none_zero .complete,
|
||||||
.name ``WickContraction.sign_insert_some_of_not_lt .incomplete,
|
.name ``WickContraction.sign_insert_some_of_not_lt .complete,
|
||||||
.name ``WickContraction.sign_insert_some_of_lt .incomplete,
|
.name ``WickContraction.sign_insert_some_of_lt .complete,
|
||||||
.name ``WickContraction.sign_insert_some_zero .incomplete,
|
.name ``WickContraction.sign_insert_some_zero .complete,
|
||||||
.h2 "Normal order",
|
.h2 "Normal order",
|
||||||
.name ``FieldSpecification.FieldOpAlgebra.normalOrder_uncontracted_none .incomplete,
|
.name ``FieldSpecification.FieldOpAlgebra.normalOrder_uncontracted_none .complete,
|
||||||
.name ``FieldSpecification.FieldOpAlgebra.normalOrder_uncontracted_some .incomplete,
|
.name ``FieldSpecification.FieldOpAlgebra.normalOrder_uncontracted_some .complete,
|
||||||
.h1 "Static contractions",
|
.h1 "Static contractions",
|
||||||
.name ``WickContraction.staticContract .incomplete,
|
.name ``WickContraction.staticContract .incomplete,
|
||||||
.name ``WickContraction.staticContract_insert_some_of_lt .incomplete,
|
.name ``WickContraction.staticContract_insert_some_of_lt .incomplete,
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue