docs: Docs for Wick contractions
This commit is contained in:
parent
ce9668a3cd
commit
c9607c459f
13 changed files with 123 additions and 64 deletions
|
@ -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 ∧
|
||||
(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`
|
||||
the sign associated with `φsΛ` is the sign corresponding to the number
|
||||
of fermionic-fermionic exchanges one must do to put elements in contracted pairs
|
||||
of `φsΛ` next to each other. -/
|
||||
/-- For a list `φs` of `𝓕.FieldOp`, and a Wick contraction `φsΛ` of `φs`,
|
||||
the complex number `φsΛ.sign` is defined to be the sign (`1` or `-1`) corresponding
|
||||
to the number of `fermionic`-`fermionic` exchanges that must done to put
|
||||
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) : ℂ :=
|
||||
∏ (a : φsΛ.1), 𝓢(𝓕 |>ₛ φs[φ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]
|
||||
· 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`
|
||||
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)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
|
||||
(φ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]
|
||||
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)
|
||||
(φsΛ : WickContraction φs.length) : (φsΛ ↩Λ φ 0 none).sign = φsΛ.sign := by
|
||||
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))
|
||||
|
||||
/--
|
||||
For `k < i`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
|
||||
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ`,
|
||||
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
|
||||
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||
`𝓕.FieldOp`, a `i ≤ φs.length` and a `k` in `φsΛ.uncontracted` such that `k<i`,
|
||||
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 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)
|
||||
(φ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
|
||||
|
||||
/--
|
||||
For `i ≤ k`, the sign of `φsΛ ↩Λ φ i (some k)` is equal to the product of
|
||||
- the sign associated with moving `φ` through the `φsΛ`-uncontracted fields in `φ₀…φₖ₋₁`,
|
||||
- the sign associated with moving `φ` through the fields in `φ₀…φᵢ₋₁`,
|
||||
For a list `φs = φ₀…φₙ` of `𝓕.FieldOp`, a Wick contraction `φsΛ` of `φs`, an element `φ` of
|
||||
`𝓕.FieldOp`, a `i ≤ φs.length` and a `k` in `φsΛ.uncontracted` such that `i ≤ k`,
|
||||
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 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)
|
||||
(φ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]
|
||||
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)
|
||||
(φsΛ : WickContraction φs.length) (k : φsΛ.uncontracted)
|
||||
(hn : GradingCompliant φs φsΛ ∧ (𝓕|>ₛφ) = 𝓕|>ₛφs[k.1]) :
|
||||
|
|
|
@ -417,11 +417,13 @@ lemma join_sign_induction {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs
|
|||
apply sign_congr
|
||||
exact join_uncontractedListGet (singleton hij) φsucΛ'
|
||||
|
||||
/-- Let `φsΛ` be a grading compliant Wick contraction for `φs` and
|
||||
`φsucΛ` a Wick contraction for `[φsΛ]ᵘᶜ`. Then `(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign`.
|
||||
/-- For a list `φs` of `𝓕.FieldOp`, a grading compliant Wick contraction `φsΛ` of `φs`,
|
||||
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
|
||||
together when defining the sign of a contraction. -/
|
||||
In `φsΛ.sign` the sign is determined by starting with the contracted pair
|
||||
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)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (hc : φsΛ.GradingCompliant) :
|
||||
(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign := by
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue