refactor: More notes on Wick's thoerem

This commit is contained in:
jstoobysmith 2025-02-04 09:58:30 +00:00
parent 20df8ece6c
commit 034f6c8c91
6 changed files with 115 additions and 77 deletions

View file

@ -32,8 +32,19 @@ def fieldOpIdealSet : Set (FieldOpFreeAlgebra 𝓕) :=
(∃ (φ φ' : 𝓕.CrAnFieldOp) (_ : ¬ (𝓕 |>ₛ φ) = (𝓕 |>ₛ φ')),
x = [ofCrAnOpF φ, ofCrAnOpF φ']ₛca)}
/-- The algebra spanned by cr and an parts of fields, with appropriate super-commutors
set to zero. -/
/-- For a field specification `𝓕`, the algebra `𝓕.FieldOpAlgebra` is defined as the quotient
of the free algebra `𝓕.FieldOpFreeAlgebra` by the ideal generated by elements of the form
- `[ofCrAnOpF φ1, [ofCrAnOpF φ2, ofCrAnOpF φ3]ₛca]ₛca`, which (with the other conditions)
gives us that super-commutors are in the center.
- `[ofCrAnOpF φc, ofCrAnOpF φc']ₛca` for `φc` and `φc'` creation operators. I.e two creation
operators always super-commute.
- `[ofCrAnOpF φa, ofCrAnOpF φa']ₛca` for `φa` and `φa'` annihilation operators. I.e two
annihilation operators always super-commute.
- `[ofCrAnOpF φ, ofCrAnOpF φ']ₛca` for `φ` and `φ'` field operators with different statistics.
I.e. Fermions super-commute with bosons.
The algebra `𝓕.FieldOpAlgebra` is the most general (in the correct sense) algebra
satisfying these properties.
-/
abbrev FieldOpAlgebra : Type := (TwoSidedIdeal.span 𝓕.fieldOpIdealSet).ringCon.Quotient
namespace FieldOpAlgebra

View file

@ -469,9 +469,10 @@ lemma ofCrAnFieldOp_superCommute_normalOrder_ofFieldOpList_sum (φ : 𝓕.CrAnFi
rw [← Finset.sum_mul, ← map_sum, ← map_sum, ← ofFieldOp_eq_sum, ← ofFieldOpList_eq_sum]
/--
Within a proto-operator algebra we have that
`[anPartF φ, 𝓝(φs)] = ∑ i, sᵢ • [anPartF φ, φᵢ]ₛ * 𝓝(φ₀…φᵢ₋₁φᵢ₊₁…φₙ)`
where `sᵢ` is the exchange sign for `φ` and `φ₀…φᵢ₋₁`.
The commutor of the annihilation part of a field operator with a normal ordered list of field
operators can be decomponsed into the sum of the commutators of the annihilation part with each
element of the list of field operators, i.e.
`[anPart φ, 𝓝(φ₀…φₙ)]ₛ= ∑ i, 𝓢(φ, φ₀…φᵢ₋₁) • [anPart φ, φᵢ]ₛ * 𝓝(φ₀…φᵢ₋₁φᵢ₊₁…φₙ)`.
-/
lemma anPart_superCommute_normalOrder_ofFieldOpList_sum (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) :
[anPart φ, 𝓝(ofFieldOpList φs)]ₛ = ∑ n : Fin φs.length, 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (φs.take n)) •

View file

@ -67,13 +67,13 @@ def statesIsPosition : 𝓕.FieldOp → Bool
| FieldOp.position _ => true
| _ => false
/-- The statistics associated to a state. -/
/-- The statistics associated to a field operator. -/
def statesStatistic : 𝓕.FieldOp → FieldStatistic := 𝓕.statistics ∘ 𝓕.fieldOpToField
/-- The field statistics associated with a state. -/
/-- The field statistics associated with a field operator. -/
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => statesStatistic 𝓕 φ
/-- The field statistics associated with a list states. -/
/-- The field statistics associated with a list field operators. -/
scoped[FieldSpecification] notation 𝓕 "|>ₛ" φ => FieldStatistic.ofList
(statesStatistic 𝓕) φ

View file

@ -48,6 +48,11 @@ def isFullInvolutionEquiv : {c : WickContraction n // IsFull c} ≃
left_inv c := by simp
right_inv f := by simp
remark card_in_general := "The cardinality of `WickContraction n` in general
follows OEIS:A000085. I.e.
1, 1, 2, 4, 10, 26, 76, 232, 764, 2620, 9496, 35696, 140152, 568504, 2390480,
10349536, 46206736, 211799312, 997313824,...."
/-- If `n` is even then the number of full contractions is `(n-1)!!`. -/
theorem card_of_isfull_even (he : Even n) :
Fintype.card {c : WickContraction n // IsFull c} = (n - 1)‼ := by

View file

@ -20,10 +20,9 @@ open HepLean.List
open FieldStatistic
lemma signFinset_insertAndContract_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (i1 i2 : Fin φs.length) :
(φsΛ ↩Λ φ i none).signFinset (finCongr (insertIdx_length_fin φ φs i).symm
(i.succAbove i1)) (finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove i2)) =
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (i1 i2 : Fin φs.length) :
(φsΛ ↩Λ φ i none).signFinset (finCongr (insertIdx_length_fin φ φs i).symm
(i.succAbove i1)) (finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove i2)) =
if i.succAbove i1 < i ∧ i < i.succAbove i2 then
Insert.insert (finCongr (insertIdx_length_fin φ φs i).symm i)
(insertAndContractLiftFinset φ i (φsΛ.signFinset i1 i2))
@ -214,8 +213,18 @@ lemma signInsertNone_eq_filter_map (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
exact List.nodup_finRange φs.length
· exact hG
/-- The change in sign when inserting a field `φ` at `i` into `φsΛ` is equal
to the sign got by moving `φ` through each field `φ₀…φᵢ₋₁` which has a dual. -/
/-- The following signs for a grading compliant Wick contraction are equal:
- The sign `φsΛ.signInsertNone φ φs i` which is given by the following: For each
contracted pair `{a1, a2}` in `φsΛ` if `a1 < a2`
such that `i` is within the range `a1 < i < a2` we pick up a sign equal to `𝓢(φ, φs[a2])`.
- The sign got by moving `φ` through `φ₀…φᵢ₋₁` and only picking up a sign when `φᵢ` has a dual.
These are equal since: Both ignore uncontracted fields, and for a contracted pair `{a1, a2}`
with `a1 < a2`
- if `i < a1 < a2` then we don't pick up a sign from either `φₐ₁` or `φₐ₂`.
- if `a1 < i < a2` then we pick up a sign from `φₐ₁` cases which is equal to `𝓢(φ, φs[a2])`
(since `φsΛ` is grading compliant).
- if `a1 < a2 < i` then we pick up a sign from both `φₐ₁` and `φₐ₂` which cancel each other out.
-/
lemma signInsertNone_eq_filterset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
φsΛ.signInsertNone φ φs i = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, Finset.univ.filter

View file

@ -198,11 +198,13 @@ lemma signFinset_insertAndContract_some (φ : 𝓕.FieldOp) (φs : List 𝓕.Fie
simp only [Fin.coe_cast, Option.get_map, Function.comp_apply, Fin.val_fin_lt]
rw [Fin.succAbove_lt_succAbove_iff]
/-- Given a Wick contraction `c` associated with a list of states `φs`
and an `i : Fin φs.length.succ`, the change in sign of the contraction associated with
inserting `φ` into `φs` at position `i` and contracting it with `j : c.uncontracted`
coming from contractions other then the `i` and `j` contraction but which
are effected by this new contraction. -/
/--
Given a Wick contraction `φsΛ` the sign defined in the followin way,
related to inserting a field `φ` at position `i` and contracting it with `j : φsΛ.uncontracted`.
- For each contracted pair `{a1, a2}` in `φsΛ` with `a1 < a2` the sign
`𝓢(φ, φₐ₂)` if `a₁ < i ≤ a₂` and `a₁ < j`.
- For each contracted pair `{a1, a2}` in `φsΛ` with `a1 < a2` the sign
`𝓢(φⱼ, φₐ₂)` if `a₁ < j < a₂` and `i < a₁`. -/
def signInsertSomeProd (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) : :=
∏ (a : φsΛ.1),
@ -216,10 +218,12 @@ def signInsertSomeProd (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : Wi
else
1
/-- Given a Wick contraction `c` associated with a list of states `φs`
and an `i : Fin φs.length.succ`, the change in sign of the contraction associated with
inserting `φ` into `φs` at position `i` and contracting it with `j : c.uncontracted`
coming from putting `i` next to `j`. -/
/-- Given a Wick contraction `φsΛ` the sign defined in the followin way,
related to inserting a field `φ` at position `i` and contracting it with `j : φsΛ.uncontracted`.
- If `j < i`, for each field `φₐ` in `φⱼ₊₁…φᵢ₋₁` without a dual at position `< j`
the sign `𝓢(φₐ, φᵢ)`.
- Else, for each field `φₐ` in `φᵢ…φⱼ₋₁` of `φs` without dual at position `< i` the sign
`𝓢(φₐ, φⱼ)`. -/
def signInsertSomeCoef (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) : :=
let a : (φsΛ ↩Λ φ i (some j)).1 := congrLift (insertIdx_length_fin φ φs i).symm
@ -239,8 +243,7 @@ def signInsertSome (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickCo
lemma sign_insert_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
(φsΛ ↩Λ φ i (some j)).sign = (φsΛ.signInsertSome φ φs i j) * φsΛ.sign := by
rw [sign]
rw [signInsertSome, signInsertSomeProd, sign, mul_assoc, ← Finset.prod_mul_distrib]
rw [sign, signInsertSome, signInsertSomeProd, sign, mul_assoc, ← Finset.prod_mul_distrib]
rw [insertAndContract_some_prod_contractions]
congr
funext a
@ -377,40 +380,6 @@ lemma signInsertSomeProd_eq_prod_fin (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldO
simp only [hφj, Fin.getElem_fin]
exact hg
lemma signInsertSomeProd_eq_list (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1]))
(hg : GradingCompliant φs φsΛ) :
φsΛ.signInsertSomeProd φ φs i j =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (List.filter (fun x => (φsΛ.getDual? x).isSome ∧
∀ (h : (φsΛ.getDual? x).isSome), x < j ∧ (i.succAbove x < i ∧
i < i.succAbove ((φsΛ.getDual? x).get h)
j < ((φsΛ.getDual? x).get h) ∧ ¬ i.succAbove x < i))
(List.finRange φs.length)).map φs.get) := by
rw [signInsertSomeProd_eq_prod_fin]
rw [FieldStatistic.ofList_map_eq_finset_prod]
rw [map_prod]
congr
funext x
split
· rename_i h
simp only [Nat.succ_eq_add_one, not_lt, instCommGroup.eq_1, Bool.decide_and,
Bool.decide_eq_true, List.mem_filter, List.mem_finRange, h, forall_true_left, Bool.decide_or,
Bool.true_and, Bool.and_eq_true, decide_eq_true_eq, Bool.or_eq_true, true_and,
Fin.getElem_fin]
split
· rename_i h1
simp [h1]
· rename_i h1
simp [h1]
· rename_i h
simp [h]
refine
List.Nodup.filter _ ?_
exact List.nodup_finRange φs.length
simp only [hφj, Fin.getElem_fin]
exact hg
lemma signInsertSomeProd_eq_finset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length)
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1]))
@ -630,10 +599,33 @@ lemma signInsertSomeCoef_eq_finset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
stat_signFinset_insert_some_self_fst]
simp [hφj]
/-- The change in sign when inserting a field `φ` at `i` into `φsΛ` and
contracting it with `k` (`k < i`) is equal
to the sign got by moving `φ` through each field `φ₀…φᵢ₋₁`
multiplied by the sign got moving `φ` through each uncontracted field `φ₀…φₖ`. -/
/--
The following two signs are equal for `i.succAbove k < i`. The sign `signInsertSome φ φs φsΛ i k` which is constructed
as follows:
1a. For each contracted pair `{a1, a2}` in `φsΛ` with `a1 < a2` the sign
`𝓢(φ, φₐ₂)` if `a₁ < i ≤ a₂` and `a₁ < k`.
1b. For each contracted pair `{a1, a2}` in `φsΛ` with `a1 < a2` the sign
`𝓢(φⱼ, φₐ₂)` if `a₁ < k < a₂` and `i < a₁`.
1c. For each field `φₐ` in `φₖ₊₁…φᵢ₋₁` without a dual at position `< k`
the sign `𝓢(φₐ, φᵢ)`.
and the sign constructed as follows:
2a. For each uncontracted field `φₐ` in `φ₀…φₖ` in `φsΛ` the sign `𝓢(φ, φₐ)`
2b. For each field in `φₐ` in `φ₀…φᵢ₋₁` the sign `𝓢(φ, φₐ)`.
The outline of why this is true can be got by considering contributions of fields.
- `φₐ`, `i ≤ a`. No contributions.
- `φₖ`, `k -> 2a`, `k -> 2b`
- `φₐ`, `a ≤ k` uncontracted `a -> 2a`, `a -> 2b`.
- `φₐ`, `k < a < i` uncontracted `a -> 1c`, `a -> 2b`.
For contracted fields `{a₁, a₂}` in `φsΛ` with `a₁ < a₂` we have the following cases:
- `φₐ₁` `φₐ₂` `a₁ < a₂ < k < i`, `a₁ -> 2b`, `a₂ -> 2b`,
- `φₐ₁` `φₐ₂` `a₁ < k < a₂ < i`, `a₁ -> 2b`, `a₂ -> 2b`,
- `φₐ₁` `φₐ₂` `a₁ < k < i ≤ a₂`, `a₁ -> 2b`, `a₂ -> 1a`
- `φₐ₁` `φₐ₂` `k < a₁ < a₂ < i`, `a₁ -> 2b`, `a₂ -> 2b`, `a₁ -> 1c`, `a₂ -> 1c`
- `φₐ₁` `φₐ₂` `k < a₁ < i ≤ a₂ `,`a₁ -> 2b`, `a₁ -> 1c`
- `φₐ₁` `φₐ₂` `k < i ≤ a₁ < a₂ `, No contributions.
-/
lemma signInsertSome_mul_filter_contracted_of_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
(hk : i.succAbove k < i) (hg : GradingCompliant φs φsΛ ∧ (𝓕 |>ₛ φ) = 𝓕 |>ₛ φs[k.1]) :
@ -669,14 +661,12 @@ lemma signInsertSome_mul_filter_contracted_of_lt (φ : 𝓕.FieldOp) (φs : List
simp only [not_lt] at hkj
have h2 := h.2 hkj
apply Fin.ne_succAbove i j
have hij : i.succAbove j ≤ i.succAbove k.1 :=
Fin.succAbove_le_succAbove_iff.mpr hkj
have hij : i.succAbove j ≤ i.succAbove k.1 := Fin.succAbove_le_succAbove_iff.mpr hkj
omega
· have h1' := h.1
rcases h1' with h1' | h1'
· have hl := h.2 h1'
have hij : i.succAbove j ≤ i.succAbove k.1 :=
Fin.succAbove_le_succAbove_iff.mpr h1'
have hij : i.succAbove j ≤ i.succAbove k.1 := Fin.succAbove_le_succAbove_iff.mpr h1'
by_contra hn
apply Fin.ne_succAbove i j
omega
@ -737,10 +727,35 @@ lemma signInsertSome_mul_filter_contracted_of_lt (φ : 𝓕.FieldOp) (φs : List
or_true, imp_self]
omega
/-- The change in sign when inserting a field `φ` at `i` into `φsΛ` and
contracting it with `k` (`i < k`) is equal
to the sign got by moving `φ` through each field `φ₀…φᵢ₋₁`
multiplied by the sign got moving `φ` through each uncontracted field `φ₀…φₖ₋₁`. -/
/--
The following two signs are equal for `i < i.succAbove k`.
The sign `signInsertSome φ φs φsΛ i k` which is constructed
as follows:
1a. For each contracted pair `{a1, a2}` in `φsΛ` with `a1 < a2` the sign
`𝓢(φ, φₐ₂)` if `a₁ < i ≤ a₂` and `a₁ < k`.
1b. For each contracted pair `{a1, a2}` in `φsΛ` with `a1 < a2` the sign
`𝓢(φⱼ, φₐ₂)` if `a₁ < k < a₂` and `i < a₁`.
1c. For each field `φₐ` in `φᵢ…φₖ₋₁` of `φs` without dual at position `< i` the sign
`𝓢(φₐ, φⱼ)`.
and the sign constructed as follows:
2a. For each uncontracted field `φₐ` in `φ₀…φₖ₋₁` in `φsΛ` the sign `𝓢(φ, φₐ)`
2b. For each field in `φₐ` in `φ₀…φᵢ₋₁` the sign `𝓢(φ, φₐ)`.
The outline of why this is true can be got by considering contributions of fields.
- `φₐ`, `k < a`. No contributions.
- `φₖ`, No Contributes
- `φₐ`, `a < i` uncontracted `a -> 2a`, `a -> 2b`.
- `φₐ`, `i ≤ a < k` uncontracted `a -> 1c`, `a -> 2a`.
For contracted fields `{a₁, a₂}` in `φsΛ` with `a₁ < a₂` we have the following cases:
- `φₐ₁` `φₐ₂` `a₁ < a₂ < i ≤ k`, `a₁ -> 2b`, `a₂ -> 2b`
- `φₐ₁` `φₐ₂` `a₁ < i ≤ a₂ < k`, `a₁ -> 2b`, `a₂ -> 1a`
- `φₐ₁` `φₐ₂` `a₁ < i ≤ k < a₂`, `a₁ -> 2b`, `a₂ -> 1a`
- `φₐ₁` `φₐ₂` `i ≤ a₁ < a₂ < k`, `a₂ -> 1c`, `a₁ -> 1c`
- `φₐ₁` `φₐ₂` `i ≤ a₁ < k < a₂ `, `a₁ -> 1c`, `a₁ -> 1b`
- `φₐ₁` `φₐ₂` `i ≤ k ≤ a₁ < a₂ `, No contributions
-/
lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
(hk : ¬ i.succAbove k < i) (hg : GradingCompliant φs φsΛ ∧ (𝓕 |>ₛ φ) = 𝓕 |>ₛ φs[k.1]) :
@ -804,9 +819,7 @@ lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.FieldOp) (φs :
imp_false, not_lt, true_and, implies_true, imp_self, and_true, forall_const, hik,
imp_forall_iff_forall]
· have hikn : j < k.1 := by omega
have hksucc : i.succAbove j < i.succAbove k.1 := by
rw [Fin.succAbove_lt_succAbove_iff]
omega
have hksucc : i.succAbove j < i.succAbove k.1 := Fin.succAbove_lt_succAbove_iff.mpr hikn
simp only [hikn, true_and, forall_const, hik, false_and, or_false, IsEmpty.forall_iff,
and_true]
by_cases hij: i < i.succAbove j
@ -822,9 +835,8 @@ lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.FieldOp) (φs :
· apply Or.inl
omega
· apply Or.inl
have hi : i.succAbove k.1 < i.succAbove ((φsΛ.getDual? j).get hj) := by
rw [Fin.succAbove_lt_succAbove_iff]
omega
have hi : i.succAbove k.1 < i.succAbove ((φsΛ.getDual? j).get hj) :=
Fin.succAbove_lt_succAbove_iff.mpr h1
apply And.intro
· apply Or.inr
apply And.intro