feat: Time dependent Wick theorem. (#274)

feat: Proof of the time-dependent Wick's theorem
This commit is contained in:
Joseph Tooby-Smith 2025-01-20 15:17:48 +00:00 committed by GitHub
parent 4d43698b3c
commit 17f84b7153
No known key found for this signature in database
GPG key ID: B5690EEEBB952194
53 changed files with 8563 additions and 3329 deletions

View file

@ -0,0 +1,525 @@
/-
Copyright (c) 2025 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.Algebras.OperatorAlgebra.NormalOrder
import HepLean.Mathematics.List.InsertIdx
/-!
# Contractions
-/
open FieldStruct
variable {𝓕 : FieldStruct}
/--
Given a natural number `n` corresponding to the number of fields, a Wick contraction
is a finite set of pairs of `Fin n`, such that no element of `Fin n` occurs in more then one pair.
For example for `n = 3` there are `4` Wick contractions:
- `∅`, 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, 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.
For `n=4` some possible Wick contractions are
- `∅`, 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
contracted and the field at position `2` and `3` are contracted.
- `{{0, 2}, {1, 3}}`, corresponding to the case where the field at position `0` and `2` are
contracted and the field at position `1` and `3` are contracted.
etc.
-/
def WickContraction (n : ) : Type :=
{f : Finset ((Finset (Fin n))) // (∀ a ∈ f, a.card = 2) ∧
(∀ a ∈ f, ∀ b ∈ f, a = b Disjoint a b)}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
/-- The contraction consisting of no contracted pairs. -/
def empty : WickContraction n := ⟨∅, by simp, by simp⟩
/-- The equivalence between `WickContraction n` and `WickContraction m`
derived from a propositional equality of `n` and `m`. -/
def congr : {n m : } → (h : n = m) → WickContraction n ≃ WickContraction m
| n, .(n), rfl => Equiv.refl _
@[simp]
lemma congr_refl : c.congr rfl = c := by
cases c
rfl
lemma congr_contractions {n m : } (h : n = m) (c : WickContraction n) :
((congr h) c).1 = Finset.map (Finset.mapEmbedding (finCongr h)).toEmbedding c.1 := by
subst h
simp only [congr_refl, Finset.le_eq_subset, finCongr_refl, Equiv.refl_toEmbedding]
ext a
apply Iff.intro
· intro ha
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
use a
simp only [ha, true_and]
rw [Finset.mapEmbedding_apply]
simp
· intro ha
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding] at ha
obtain ⟨b, hb, hab⟩ := ha
rw [Finset.mapEmbedding_apply] at hab
simp only [Finset.map_refl] at hab
subst hab
exact hb
@[simp]
lemma congr_trans {n m o : } (h1 : n = m) (h2 : m = o) :
(congr h1).trans (congr h2) = congr (h1.trans h2) := by
subst h1 h2
simp [congr]
@[simp]
lemma congr_trans_apply {n m o : } (h1 : n = m) (h2 : m = o) (c : WickContraction n) :
(congr h2) ((congr h1) c) = congr (h1.trans h2) c := by
subst h1 h2
simp
/-- Given a contracted pair in `c : WickContraction n` the contracted pair
in `congr h c`. -/
def congrLift {n m : } (h : n = m) {c : WickContraction n} (a : c.1) : (congr h c).1 :=
⟨a.1.map (finCongr h).toEmbedding, by
subst h
simp⟩
@[simp]
lemma congrLift_rfl {n : } {c : WickContraction n} :
c.congrLift rfl = id := by
funext a
simp [congrLift]
lemma congrLift_injective {n m : } {c : WickContraction n} (h : n = m) :
Function.Injective (c.congrLift h) := by
subst h
simp only [congrLift_rfl]
exact fun ⦃a₁ a₂⦄ a => a
lemma congrLift_surjective {n m : } {c : WickContraction n} (h : n = m) :
Function.Surjective (c.congrLift h) := by
subst h
simp only [congrLift_rfl]
exact Function.surjective_id
lemma congrLift_bijective {n m : } {c : WickContraction n} (h : n = m) :
Function.Bijective (c.congrLift h) := by
exact ⟨c.congrLift_injective h, c.congrLift_surjective h⟩
lemma eq_filter_mem_self : c.1 = Finset.filter (fun x => x ∈ c.1) Finset.univ := by
exact Eq.symm (Finset.filter_univ_mem c.1)
/-- For a contraction `c : WickContraction n` and `i : Fin n` the `j` such that
`{i, j}` is a contracted pair in `c`. If such an `j` does not exist, this returns `none`. -/
def getDual? (i : Fin n) : Option (Fin n) := Fin.find (fun j => {i, j} ∈ c.1)
lemma getDual?_congr {n m : } (h : n = m) (c : WickContraction n) (i : Fin m) :
(congr h c).getDual? i = Option.map (finCongr h) (c.getDual? (finCongr h.symm i)) := by
subst h
simp
lemma getDual?_congr_get {n m : } (h : n = m) (c : WickContraction n) (i : Fin m)
(hg : ((congr h c).getDual? i).isSome) :
((congr h c).getDual? i).get hg =
(finCongr h ((c.getDual? (finCongr h.symm i)).get (by simpa [getDual?_congr] using hg))) := by
simp only [getDual?_congr, finCongr_apply]
exact Option.get_map
lemma getDual?_eq_some_iff_mem (i j : Fin n) :
c.getDual? i = some j ↔ {i, j} ∈ c.1 := by
simp only [getDual?]
rw [Fin.find_eq_some_iff]
apply Iff.intro
· intro h
exact h.1
· intro h
simp only [h, true_and]
intro k hk
have hc := c.2.2 _ h _ hk
simp only [Finset.disjoint_insert_right, Finset.mem_insert, Finset.mem_singleton, true_or,
not_true_eq_false, Finset.disjoint_singleton_right, not_or, false_and, or_false] at hc
have hj : k ∈ ({i, j} : Finset (Fin n)) := by
simp [hc]
simp only [Finset.mem_insert, Finset.mem_singleton] at hj
rcases hj with hj | hj
· subst hj
simp only [Finset.mem_singleton, Finset.insert_eq_of_mem] at hk
have hc := c.2.1 _ hk
simp at hc
· subst hj
simp
@[simp]
lemma getDual?_one_eq_none (c : WickContraction 1) (i : Fin 1) : c.getDual? i = none := by
by_contra h
have hn : (c.getDual? i).isSome := by
rw [← Option.not_isSome_iff_eq_none] at h
simpa [- Option.not_isSome, -Option.isNone_iff_eq_none] using h
rw [@Option.isSome_iff_exists] at hn
obtain ⟨a, hn⟩ := hn
rw [getDual?_eq_some_iff_mem] at hn
have hc := c.2.1 {i, a} hn
fin_cases i
fin_cases a
simp at hc
@[simp]
lemma getDual?_get_self_mem (i : Fin n) (h : (c.getDual? i).isSome) :
{(c.getDual? i).get h, i} ∈ c.1 := by
rw [@Finset.pair_comm]
rw [← getDual?_eq_some_iff_mem]
simp
@[simp]
lemma self_getDual?_get_mem (i : Fin n) (h : (c.getDual? i).isSome) :
{i, (c.getDual? i).get h} ∈ c.1 := by
rw [← getDual?_eq_some_iff_mem]
simp
lemma getDual?_eq_some_neq (i j : Fin n) (h : c.getDual? i = some j) :
¬ i = j := by
rw [getDual?_eq_some_iff_mem] at h
by_contra hn
subst hn
have hc := c.2.1 _ h
simp at hc
@[simp]
lemma self_neq_getDual?_get (i : Fin n) (h : (c.getDual? i).isSome) :
¬ i = (c.getDual? i).get h := by
by_contra hn
have hx : {i, (c.getDual? i).get h} ∈ c.1 := by simp
have hc := c.2.1 _ hx
nth_rewrite 1 [hn] at hc
simp at hc
@[simp]
lemma getDual?_get_self_neq (i : Fin n) (h : (c.getDual? i).isSome) :
¬ (c.getDual? i).get h = i := by
by_contra hn
have hx : {i, (c.getDual? i).get h} ∈ c.1 := by simp
have hc := c.2.1 _ hx
nth_rewrite 1 [hn] at hc
simp at hc
lemma getDual?_isSome_iff (i : Fin n) : (c.getDual? i).isSome ↔ ∃ (a : c.1), i ∈ a.1 := by
apply Iff.intro
· intro h
simp only [getDual?] at h
rw [Fin.isSome_find_iff] at h
obtain ⟨a, ha⟩ := h
use ⟨{i, a}, ha⟩
simp
· intro h
obtain ⟨a, ha⟩ := h
have ha := c.2.1 a a.2
rw [@Finset.card_eq_two] at ha
obtain ⟨x, y, hx, hy⟩ := ha
rw [hy] at ha
simp only [Finset.mem_insert, Finset.mem_singleton] at ha
match ha with
| Or.inl ha =>
subst ha
simp only [getDual?]
rw [Fin.isSome_find_iff]
use y
rw [← hy]
exact a.2
| Or.inr ha =>
subst ha
simp only [getDual?]
rw [Fin.isSome_find_iff]
use x
rw [Finset.pair_comm]
rw [← hy]
exact a.2
lemma getDual?_isSome_of_mem (a : c.1) (i : a.1) : (c.getDual? i).isSome := by
rw [getDual?_isSome_iff]
use ⟨a.1, a.2⟩
simp
@[simp]
lemma getDual?_getDual?_get_get (i : Fin n) (h : (c.getDual? i).isSome) :
c.getDual? ((c.getDual? i).get h) = some i := by
rw [getDual?_eq_some_iff_mem]
simp
lemma getDual?_getDual?_get_isSome (i : Fin n) (h : (c.getDual? i).isSome) :
(c.getDual? ((c.getDual? i).get h)).isSome := by
simp
lemma getDual?_getDual?_get_not_none (i : Fin n) (h : (c.getDual? i).isSome) :
¬ (c.getDual? ((c.getDual? i).get h)) = none := by
simp
/-!
## Extracting parts from a contraction.
-/
/-- The smallest of the two positions in a contracted pair given a Wick contraction. -/
def fstFieldOfContract (c : WickContraction n) (a : c.1) : Fin n :=
(a.1.sort (· ≤ ·)).head (by
have hx : (Finset.sort (fun x1 x2 => x1 ≤ x2) a.1).length = a.1.card := by
exact Finset.length_sort fun x1 x2 => x1 ≤ x2
by_contra hn
rw [hn, c.2.1 a.1 a.2] at hx
simp at hx)
@[simp]
lemma fstFieldOfContract_congr {n m : } (h : n = m) (c : WickContraction n) (a : c.1) :
(congr h c).fstFieldOfContract (c.congrLift h a) = (finCongr h) (c.fstFieldOfContract a) := by
subst h
simp [congr]
/-- The largest of the two positions in a contracted pair given a Wick contraction. -/
def sndFieldOfContract (c : WickContraction n) (a : c.1) : Fin n :=
(a.1.sort (· ≤ ·)).tail.head (by
have hx : (Finset.sort (fun x1 x2 => x1 ≤ x2) a.1).length = a.1.card := by
exact Finset.length_sort fun x1 x2 => x1 ≤ x2
by_contra hn
have hn := congrArg List.length hn
simp only [List.length_tail, Finset.length_sort, List.length_nil] at hn
rw [c.2.1 a.1 a.2] at hn
omega)
@[simp]
lemma sndFieldOfContract_congr {n m : } (h : n = m) (c : WickContraction n) (a : c.1) :
(congr h c).sndFieldOfContract (c.congrLift h a) = (finCongr h) (c.sndFieldOfContract a) := by
subst h
simp [congr]
lemma finset_eq_fstFieldOfContract_sndFieldOfContract (c : WickContraction n) (a : c.1) :
a.1 = {c.fstFieldOfContract a, c.sndFieldOfContract a} := by
have h1 := c.2.1 a.1 a.2
rw [Finset.card_eq_two] at h1
obtain ⟨x, y, hxy, ha⟩ := h1
rw [ha]
by_cases hxyle : x ≤ y
· have ha : a.1.sort (· ≤ ·) = [x, y] := by
rw [ha]
trans Finset.sort (· ≤ ·) (Finset.cons x {y} (by simp [hxy]))
· congr
simp
rw [Finset.sort_cons]
simp only [Finset.sort_singleton]
intro b hb
simp only [Finset.mem_singleton] at hb
subst hb
omega
simp [fstFieldOfContract, ha, sndFieldOfContract]
· have ha : a.1.sort (· ≤ ·) = [y, x] := by
rw [ha]
trans Finset.sort (· ≤ ·) (Finset.cons y {x} (by simp only [Finset.mem_singleton]; omega))
· congr
simp only [Finset.cons_eq_insert]
rw [@Finset.pair_comm]
rw [Finset.sort_cons]
simp only [Finset.sort_singleton]
intro b hb
simp only [Finset.mem_singleton] at hb
subst hb
omega
simp only [fstFieldOfContract, ha, List.head_cons, sndFieldOfContract, List.tail_cons]
rw [Finset.pair_comm]
lemma fstFieldOfContract_neq_sndFieldOfContract (c : WickContraction n) (a : c.1) :
c.fstFieldOfContract a ≠ c.sndFieldOfContract a := by
have h1 := c.2.1 a.1 a.2
have h2 := c.finset_eq_fstFieldOfContract_sndFieldOfContract a
by_contra hn
rw [h2, hn] at h1
simp at h1
lemma fstFieldOfContract_le_sndFieldOfContract (c : WickContraction n) (a : c.1) :
c.fstFieldOfContract a ≤ c.sndFieldOfContract a := by
simp only [fstFieldOfContract, sndFieldOfContract, List.head_tail]
have h1 (n : ) (l : List (Fin n)) (h : l ≠ []) (hl : l.Sorted (· ≤ ·)) :
∀ a ∈ l, l.head h ≤ a := by
induction l with
| nil => simp at h
| cons i l ih =>
simp only [List.sorted_cons] at hl
simpa using hl.1
apply h1
· exact Finset.sort_sorted (fun x1 x2 => x1 ≤ x2) _
· exact List.getElem_mem _
lemma fstFieldOfContract_lt_sndFieldOfContract (c : WickContraction n) (a : c.1) :
c.fstFieldOfContract a < c.sndFieldOfContract a :=
lt_of_le_of_ne (c.fstFieldOfContract_le_sndFieldOfContract a)
(c.fstFieldOfContract_neq_sndFieldOfContract a)
@[simp]
lemma fstFieldOfContract_mem (c : WickContraction n) (a : c.1) :
c.fstFieldOfContract a ∈ a.1 := by
rw [finset_eq_fstFieldOfContract_sndFieldOfContract]
simp
lemma fstFieldOfContract_getDual?_isSome (c : WickContraction n) (a : c.1) :
(c.getDual? (c.fstFieldOfContract a)).isSome := by
rw [getDual?_isSome_iff]
use a
simp
@[simp]
lemma fstFieldOfContract_getDual? (c : WickContraction n) (a : c.1) :
c.getDual? (c.fstFieldOfContract a) = some (c.sndFieldOfContract a) := by
rw [getDual?_eq_some_iff_mem]
erw [← finset_eq_fstFieldOfContract_sndFieldOfContract]
exact a.2
@[simp]
lemma sndFieldOfContract_mem (c : WickContraction n) (a : c.1) :
c.sndFieldOfContract a ∈ a.1 := by
rw [finset_eq_fstFieldOfContract_sndFieldOfContract]
simp
lemma sndFieldOfContract_getDual?_isSome (c : WickContraction n) (a : c.1) :
(c.getDual? (c.sndFieldOfContract a)).isSome := by
rw [getDual?_isSome_iff]
use a
simp
@[simp]
lemma sndFieldOfContract_getDual? (c : WickContraction n) (a : c.1) :
c.getDual? (c.sndFieldOfContract a) = some (c.fstFieldOfContract a) := by
rw [getDual?_eq_some_iff_mem]
rw [Finset.pair_comm]
erw [← finset_eq_fstFieldOfContract_sndFieldOfContract]
exact a.2
lemma eq_fstFieldOfContract_of_mem (c : WickContraction n) (a : c.1) (i j : Fin n)
(hi : i ∈ a.1) (hj : j ∈ a.1) (hij : i < j) :
c.fstFieldOfContract a = i := by
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hi hj
simp_all only [Finset.mem_insert, Finset.mem_singleton]
match hi, hj with
| Or.inl hi, Or.inl hj =>
subst hi hj
simp at hij
| Or.inl hi, Or.inr hj =>
subst hi
rfl
| Or.inr hi, Or.inl hj =>
subst hi hj
have hn := fstFieldOfContract_lt_sndFieldOfContract c a
omega
| Or.inr hi, Or.inr hj =>
subst hi hj
simp at hij
lemma eq_sndFieldOfContract_of_mem (c : WickContraction n) (a : c.1) (i j : Fin n)
(hi : i ∈ a.1) (hj : j ∈ a.1) (hij : i < j) :
c.sndFieldOfContract a = j := by
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hi hj
simp_all only [Finset.mem_insert, Finset.mem_singleton]
match hi, hj with
| Or.inl hi, Or.inl hj =>
subst hi hj
simp at hij
| Or.inl hi, Or.inr hj =>
subst hi hj
omega
| Or.inr hi, Or.inl hj =>
subst hi hj
have hn := fstFieldOfContract_lt_sndFieldOfContract c a
omega
| Or.inr hi, Or.inr hj =>
subst hi hj
simp at hij
/-- As a type, any pair of contractions is equivalent to `Fin 2`
with `0` being associated with `c.fstFieldOfContract a` and `1` being associated with
`c.sndFieldOfContract`. -/
def contractEquivFinTwo (c : WickContraction n) (a : c.1) :
a ≃ Fin 2 where
toFun i := if i = c.fstFieldOfContract a then 0 else 1
invFun i :=
match i with
| 0 => ⟨c.fstFieldOfContract a, fstFieldOfContract_mem c a⟩
| 1 => ⟨c.sndFieldOfContract a, sndFieldOfContract_mem c a⟩
left_inv i := by
simp only [Fin.isValue]
have hi := i.2
have ha := c.finset_eq_fstFieldOfContract_sndFieldOfContract a
simp only [ha, Finset.mem_insert, Finset.mem_singleton] at hi
rcases hi with hi | hi
· rw [hi]
simp only [↓reduceIte, Fin.isValue]
exact Subtype.eq (id (Eq.symm hi))
· rw [hi]
rw [if_neg]
simp only
exact Subtype.eq (id (Eq.symm hi))
have h := fstFieldOfContract_neq_sndFieldOfContract c a
exact (Ne.symm h)
right_inv i := by
fin_cases i
· simp
· simp only [Fin.isValue, Fin.mk_one, ite_eq_right_iff, zero_ne_one, imp_false]
have h := fstFieldOfContract_neq_sndFieldOfContract c a
exact (Ne.symm h)
lemma prod_finset_eq_mul_fst_snd (c : WickContraction n) (a : c.1)
(f : a.1 → M) [CommMonoid M] :
∏ (x : a), f x = f (⟨c.fstFieldOfContract a, fstFieldOfContract_mem c a⟩)
* f (⟨c.sndFieldOfContract a, sndFieldOfContract_mem c a⟩) := by
rw [← (c.contractEquivFinTwo a).symm.prod_comp]
simp [contractEquivFinTwo]
/-- A Wick contraction associated with a list of states is said to be `GradingCompliant` if in any
contracted pair of states they are either both fermionic or both bosonic. -/
def GradingCompliant (φs : List 𝓕.States) (c : WickContraction φs.length) :=
∀ (a : c.1), (𝓕 |>ₛ φs[c.fstFieldOfContract a]) = (𝓕 |>ₛ φs[c.sndFieldOfContract a])
/-- An equivalence from the sigma type `(a : c.1) × a` to the subtype of `Fin n` consisting of
those positions which are contracted. -/
def sigmaContractedEquiv : (a : c.1) × a ≃ {x : Fin n // (c.getDual? x).isSome} where
toFun := fun x => ⟨x.2, getDual?_isSome_of_mem c x.fst x.snd⟩
invFun := fun x => ⟨
⟨{x.1, (c.getDual? x.1).get x.2}, self_getDual?_get_mem c (↑x) x.prop⟩,
⟨x.1, by simp⟩⟩
left_inv x := by
have hxa (x1 x2 : (a : c.1) × a) (h1 : x1.1 = x2.1)
(h2 : x1.2.val = x2.2.val) : x1 = x2 := by
cases x1
cases x2
simp_all only [Sigma.mk.inj_iff, true_and]
subst h1
rename_i fst snd snd_1
simp_all only [heq_eq_eq]
obtain ⟨val, property⟩ := fst
obtain ⟨val_2, property_2⟩ := snd
subst h2
simp_all only
match x with
| ⟨a, i⟩ =>
apply hxa
· have hc := c.2.2 a.1 a.2 {i.1, (c.getDual? ↑i).get (getDual?_isSome_of_mem c a i)}
(self_getDual?_get_mem c (↑i) (getDual?_isSome_of_mem c a i))
have hn : ¬ Disjoint a.1 {i.1, (c.getDual? ↑i).get (getDual?_isSome_of_mem c a i)} := by
rw [Finset.disjoint_iff_inter_eq_empty]
rw [@Finset.eq_empty_iff_forall_not_mem]
simp only [Finset.coe_mem, Finset.inter_insert_of_mem, Finset.mem_insert, Finset.mem_inter,
Finset.mem_singleton, not_or, not_and, not_forall, Classical.not_imp, Decidable.not_not]
use i
simp
simp_all only [or_false, disjoint_self, Finset.bot_eq_empty, Finset.insert_ne_empty,
not_false_eq_true]
exact Subtype.eq (id (Eq.symm hc))
· simp
right_inv := by
intro x
cases x
rfl
end WickContraction

View file

@ -0,0 +1,160 @@
/-
Copyright (c) 2025 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.WickContraction.Uncontracted
/-!
# Erasing an element from a contraction
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open HepLean.Fin
/-- Given a Wick contraction `WickContraction n.succ` and a `i : Fin n.succ` the
Wick contraction associated with `n` obtained by removing `i`.
If `i` is contracted with `j` in the new Wick contraction `j` will be uncontracted. -/
def erase (c : WickContraction n.succ) (i : Fin n.succ) : WickContraction n := by
refine ⟨Finset.filter (fun x => Finset.map i.succAboveEmb x ∈ c.1) Finset.univ, ?_, ?_⟩
· intro a ha
simpa using c.2.1 (Finset.map i.succAboveEmb a) (by simpa using ha)
· intro a ha b hb
simp only [Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and] at ha hb
rw [← Finset.disjoint_map i.succAboveEmb, ← (Finset.map_injective i.succAboveEmb).eq_iff]
exact c.2.2 _ ha _ hb
lemma mem_erase_uncontracted_iff (c : WickContraction n.succ) (i : Fin n.succ) (j : Fin n) :
j ∈ (c.erase i).uncontracted ↔
i.succAbove j ∈ c.uncontracted c.getDual? (i.succAbove j) = some i := by
rw [getDual?_eq_some_iff_mem]
simp only [uncontracted, getDual?, erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ,
Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, true_and]
rw [Fin.find_eq_none_iff, Fin.find_eq_none_iff]
apply Iff.intro
· intro h
by_cases hi : {i.succAbove j, i} ∈ c.1
· simp [hi]
· apply Or.inl
intro k
by_cases hi' : k = i
· subst hi'
exact hi
· simp only [← Fin.exists_succAbove_eq_iff] at hi'
obtain ⟨z, hz⟩ := hi'
subst hz
exact h z
· intro h
intro k
rcases h with h | h
· exact h (i.succAbove k)
· by_contra hn
have hc := c.2.2 _ h _ hn
simp only [Nat.succ_eq_add_one, Finset.disjoint_insert_right, Finset.mem_insert,
Finset.mem_singleton, true_or, not_true_eq_false, Finset.disjoint_singleton_right, not_or,
false_and, or_false] at hc
have hi : i ∈ ({i.succAbove j, i.succAbove k} : Finset (Fin n.succ)) := by
simp [← hc]
simp only [Nat.succ_eq_add_one, Finset.mem_insert, Finset.mem_singleton] at hi
rcases hi with hi | hi
· exact False.elim (Fin.succAbove_ne _ _ hi.symm)
· exact False.elim (Fin.succAbove_ne _ _ hi.symm)
lemma mem_not_eq_erase_of_isSome (c : WickContraction n.succ) (i : Fin n.succ)
(h : (c.getDual? i).isSome) (ha : a ∈ c.1) (ha2 : a ≠ {i, (c.getDual? i).get h}) :
∃ a', a' ∈ (c.erase i).1 ∧ a = Finset.map i.succAboveEmb a' := by
have h2a := c.2.1 a ha
rw [@Finset.card_eq_two] at h2a
obtain ⟨x, y, hx,hy⟩ := h2a
subst hy
have hxn : ¬ x = i := by
by_contra hx
subst hx
rw [← @getDual?_eq_some_iff_mem] at ha
rw [(Option.get_of_mem h ha)] at ha2
simp at ha2
have hyn : ¬ y = i := by
by_contra hy
subst hy
rw [@Finset.pair_comm] at ha
rw [← @getDual?_eq_some_iff_mem] at ha
rw [(Option.get_of_mem h ha)] at ha2
simp [Finset.pair_comm] at ha2
simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hxn hyn
obtain ⟨x', hx'⟩ := hxn
obtain ⟨y', hy'⟩ := hyn
use {x', y'}
subst hx' hy'
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, Finset.map_insert,
Fin.succAboveEmb_apply, Finset.map_singleton, true_and, and_true]
exact ha
lemma mem_not_eq_erase_of_isNone (c : WickContraction n.succ) (i : Fin n.succ)
(h : (c.getDual? i).isNone) (ha : a ∈ c.1) :
∃ a', a' ∈ (c.erase i).1 ∧ a = Finset.map i.succAboveEmb a' := by
have h2a := c.2.1 a ha
rw [@Finset.card_eq_two] at h2a
obtain ⟨x, y, hx,hy⟩ := h2a
subst hy
have hi : i ∈ c.uncontracted := by
simp only [Nat.succ_eq_add_one, uncontracted, Finset.mem_filter, Finset.mem_univ, true_and]
simp_all only [Nat.succ_eq_add_one, Option.isNone_iff_eq_none, ne_eq]
rw [@mem_uncontracted_iff_not_contracted] at hi
have hxn : ¬ x = i := by
by_contra hx
subst hx
exact hi {x, y} ha (by simp)
have hyn : ¬ y = i := by
by_contra hy
subst hy
exact hi {x, y} ha (by simp)
simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hxn hyn
obtain ⟨x', hx'⟩ := hxn
obtain ⟨y', hy'⟩ := hyn
use {x', y'}
subst hx' hy'
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, Finset.map_insert,
Fin.succAboveEmb_apply, Finset.map_singleton, true_and, and_true]
exact ha
/-- Given a Wick contraction `c : WickContraction n.succ` and a `i : Fin n.succ` the (optional)
element of `(erase c i).uncontracted` which comes from the element in `c` contracted
with `i`. -/
def getDualErase {n : } (c : WickContraction n.succ) (i : Fin n.succ) :
Option ((erase c i).uncontracted) := by
match n with
| 0 => exact none
| Nat.succ n =>
refine if hj : (c.getDual? i).isSome then some ⟨(predAboveI i ((c.getDual? i).get hj)), ?_⟩
else none
rw [mem_erase_uncontracted_iff]
apply Or.inr
rw [succsAbove_predAboveI, getDual?_eq_some_iff_mem]
· simp
· apply c.getDual?_eq_some_neq _ _ _
simp
@[simp]
lemma getDualErase_isSome_iff_getDual?_isSome (c : WickContraction n.succ) (i : Fin n.succ) :
(c.getDualErase i).isSome ↔ (c.getDual? i).isSome := by
match n with
| 0 =>
fin_cases i
simp [getDualErase]
| Nat.succ n =>
simp [getDualErase]
@[simp]
lemma getDualErase_one (c : WickContraction 1) (i : Fin 1) :
c.getDualErase i = none := by
fin_cases i
simp [getDualErase]
end WickContraction

View file

@ -0,0 +1,110 @@
/-
Copyright (c) 2025 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.WickContraction.Insert
/-!
# Equivalence extracting element from contraction
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open HepLean.Fin
lemma extractEquiv_equiv {c1 c2 : (c : WickContraction n) × Option c.uncontracted}
(h : c1.1 = c2.1) (ho : c1.2 = uncontractedCongr (by rw [h]) c2.2) : c1 = c2 := by
cases c1
cases c2
simp_all only [Sigma.mk.inj_iff]
simp only at h
subst h
simp only [uncontractedCongr, Equiv.optionCongr_apply, heq_eq_eq, true_and]
rename_i a
match a with
| none => simp
| some a =>
simp only [Option.map_some', Option.some.injEq]
ext
simp
/-- The equivalence between `WickContraction n.succ` and the sigma type
`(c : WickContraction n) × Option c.uncontracted` formed by inserting
and erasing elements from a contraction. -/
def extractEquiv (i : Fin n.succ) : WickContraction n.succ ≃
(c : WickContraction n) × Option c.uncontracted where
toFun := fun c => ⟨erase c i, getDualErase c i⟩
invFun := fun ⟨c, j⟩ => insert c i j
left_inv f := by
simp
right_inv f := by
refine extractEquiv_equiv ?_ ?_
simp only [insert_erase]
simp only [Nat.succ_eq_add_one]
have h1 := insert_getDualErase f.fst i f.snd
exact insert_getDualErase _ i _
lemma extractEquiv_symm_none_uncontracted (i : Fin n.succ) (c : WickContraction n) :
((extractEquiv i).symm ⟨c, none⟩).uncontracted =
(Insert.insert i (c.uncontracted.map i.succAboveEmb)) := by
exact insert_none_uncontracted c i
@[simp]
lemma extractEquiv_apply_congr_symm_apply {n m : } (k : )
(hnk : k < n.succ) (hkm : k < m.succ) (hnm : n = m) (c : WickContraction n)
(i : c.uncontracted) : congr (by rw [hnm]) ((extractEquiv ⟨k, hkm⟩
(congr (by rw [hnm]) ((extractEquiv ⟨k, hnk⟩).symm ⟨c, i⟩)))).1 = c := by
subst hnm
simp
/-- The fintype instance of `WickContraction 0` defined through its single
element `empty`. -/
instance fintype_zero : Fintype (WickContraction 0) where
elems := {empty}
complete := by
intro c
simp only [Finset.mem_singleton]
apply Subtype.eq
simp only [empty]
ext a
apply Iff.intro
· intro h
have hc := c.2.1 a h
rw [Finset.card_eq_two] at hc
obtain ⟨x, y, hxy, ha⟩ := hc
exact Fin.elim0 x
· simp
lemma sum_WickContraction_nil (f : WickContraction 0 → M) [AddCommMonoid M] :
∑ c, f c = f empty := by
rw [Finset.sum_eq_single_of_mem]
simp only [Finset.mem_univ]
intro b hb
fin_cases b
simp
/-- The fintype instance of `WickContraction n`, for `n.succ` this is defined
through the equivalence `extractEquiv`. -/
instance fintype_succ : (n : ) → Fintype (WickContraction n)
| 0 => fintype_zero
| Nat.succ n => by
letI := fintype_succ n
exact Fintype.ofEquiv _ (extractEquiv 0).symm
lemma sum_extractEquiv_congr [AddCommMonoid M] {n m : } (i : Fin n) (f : WickContraction n → M)
(h : n = m.succ) :
∑ c, f c = ∑ (c : WickContraction m), ∑ (k : Option c.uncontracted),
f (congr h.symm ((extractEquiv (finCongr h i)).symm ⟨c, k⟩)) := by
subst h
simp only [finCongr_refl, Equiv.refl_apply, congr_refl]
rw [← (extractEquiv i).symm.sum_comp]
rw [Finset.sum_sigma']
rfl
end WickContraction

View file

@ -0,0 +1,718 @@
/-
Copyright (c) 2025 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.WickContraction.Erase
/-!
# Inserting an element into a contraction
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open HepLean.Fin
/-!
## Inserting an element into a contraction
-/
/-- Given a Wick contraction `c` for `n`, a position `i : Fin n.succ` and
an optional uncontracted element `j : Option (c.uncontracted)` of `c`.
The Wick contraction for `n.succ` formed by 'inserting' `i` into `Fin n`
and contracting it optionally with `j`. -/
def insert (c : WickContraction n) (i : Fin n.succ) (j : Option (c.uncontracted)) :
WickContraction n.succ := by
let f := Finset.map (Finset.mapEmbedding i.succAboveEmb).toEmbedding c.1
let f' := match j with | none => f | some j => Insert.insert {i, i.succAbove j} f
refine ⟨f', ?_, ?_⟩
· simp only [Nat.succ_eq_add_one, f']
match j with
| none =>
simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding,
forall_exists_index, and_imp, forall_apply_eq_imp_iff₂, f]
intro a ha
rw [Finset.mapEmbedding_apply]
simp only [Finset.card_map]
exact c.2.1 a ha
| some j =>
simp only [Finset.mem_insert, forall_eq_or_imp]
apply And.intro
· rw [@Finset.card_eq_two]
use i
use (i.succAbove j)
simp only [ne_eq, and_true]
exact Fin.ne_succAbove i j
· simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding,
forall_exists_index, and_imp, forall_apply_eq_imp_iff₂, f]
intro a ha
rw [Finset.mapEmbedding_apply]
simp only [Finset.card_map]
exact c.2.1 a ha
· intro a ha b hb
simp only [Nat.succ_eq_add_one, f'] at ha hb
match j with
| none =>
simp_all only [f, Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding,
Nat.succ_eq_add_one]
obtain ⟨a', ha', ha''⟩ := ha
obtain ⟨b', hb', hb''⟩ := hb
subst ha''
subst hb''
simp only [EmbeddingLike.apply_eq_iff_eq]
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply, Finset.disjoint_map]
exact c.2.2 a' ha' b' hb'
| some j =>
simp_all only [Finset.mem_insert, Nat.succ_eq_add_one]
match ha, hb with
| Or.inl ha, Or.inl hb =>
rw [ha, hb]
simp
| Or.inl ha, Or.inr hb =>
apply Or.inr
subst ha
simp only [Finset.disjoint_insert_left, Finset.disjoint_singleton_left]
simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding, f] at hb
obtain ⟨a', hb', hb''⟩ := hb
subst hb''
rw [Finset.mapEmbedding_apply]
apply And.intro
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
exact fun x _ => Fin.succAbove_ne i x
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
have hj := j.2
rw [mem_uncontracted_iff_not_contracted] at hj
intro a ha hja
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hja
subst hja
exact False.elim (hj a' hb' ha)
| Or.inr ha, Or.inl hb =>
apply Or.inr
subst hb
simp only [Finset.disjoint_insert_right, Nat.succ_eq_add_one,
Finset.disjoint_singleton_right]
simp only [Finset.le_eq_subset, Finset.mem_map, RelEmbedding.coe_toEmbedding, f] at ha
obtain ⟨a', ha', ha''⟩ := ha
subst ha''
rw [Finset.mapEmbedding_apply]
apply And.intro
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
exact fun x _ => Fin.succAbove_ne i x
· simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
have hj := j.2
rw [mem_uncontracted_iff_not_contracted] at hj
intro a ha hja
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hja
subst hja
exact False.elim (hj a' ha' ha)
| Or.inr ha, Or.inr hb =>
simp_all only [f, Finset.le_eq_subset,
or_true, Finset.mem_map, RelEmbedding.coe_toEmbedding]
obtain ⟨a', ha', ha''⟩ := ha
obtain ⟨b', hb', hb''⟩ := hb
subst ha''
subst hb''
simp only [EmbeddingLike.apply_eq_iff_eq]
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply, Finset.disjoint_map]
exact c.2.2 a' ha' b' hb'
lemma insert_of_isSome (c : WickContraction n) (i : Fin n.succ) (j : Option c.uncontracted)
(hj : j.isSome) : (insert c i j).1 = Insert.insert {i, i.succAbove (j.get hj)}
(Finset.map (Finset.mapEmbedding i.succAboveEmb).toEmbedding c.1) := by
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset]
rw [@Option.isSome_iff_exists] at hj
obtain ⟨j, hj⟩ := hj
subst hj
simp
@[simp]
lemma self_mem_uncontracted_of_insert_none (c : WickContraction n) (i : Fin n.succ) :
i ∈ (insert c i none).uncontracted := by
rw [mem_uncontracted_iff_not_contracted]
intro p hp
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
RelEmbedding.coe_toEmbedding] at hp
obtain ⟨a, ha, ha'⟩ := hp
have hc := c.2.1 a ha
rw [@Finset.card_eq_two] at hc
obtain ⟨x, y, hxy, ha⟩ := hc
subst ha
subst ha'
rw [Finset.mapEmbedding_apply]
simp only [Nat.succ_eq_add_one, Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton,
Finset.mem_insert, Finset.mem_singleton, not_or]
apply And.intro
· exact Fin.ne_succAbove i x
· exact Fin.ne_succAbove i y
@[simp]
lemma self_not_mem_uncontracted_of_insert_some (c : WickContraction n) (i : Fin n.succ)
(j : c.uncontracted) : i ∉ (insert c i (some j)).uncontracted := by
rw [mem_uncontracted_iff_not_contracted]
simp [insert]
lemma insert_succAbove_mem_uncontracted_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
(i.succAbove j) ∈ (insert c i none).uncontracted ↔ j ∈ c.uncontracted := by
rw [mem_uncontracted_iff_not_contracted, mem_uncontracted_iff_not_contracted]
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
RelEmbedding.coe_toEmbedding, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂]
apply Iff.intro
· intro h p hp
have hp' := h p hp
have hc := c.2.1 p hp
rw [Finset.card_eq_two] at hc
obtain ⟨x, y, hxy, hp⟩ := hc
subst hp
rw [Finset.mapEmbedding_apply] at hp'
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, Finset.mem_insert,
Finset.mem_singleton, not_or] at hp'
simp only [Finset.mem_insert, Finset.mem_singleton, not_or]
exact And.intro (fun a => hp'.1 (congrArg i.succAbove a))
(fun a => hp'.2 (congrArg i.succAbove a))
· intro h p hp
have hc := c.2.1 p hp
rw [Finset.card_eq_two] at hc
obtain ⟨x, y, hxy, hp⟩ := hc
subst hp
rw [Finset.mapEmbedding_apply]
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton, Finset.mem_insert,
Finset.mem_singleton, not_or]
have hp' := h {x, y} hp
simp only [Finset.mem_insert, Finset.mem_singleton, not_or] at hp'
apply And.intro
(fun a => hp'.1 (i.succAbove_right_injective a))
(fun a => hp'.2 (i.succAbove_right_injective a))
@[simp]
lemma mem_uncontracted_insert_none_iff (c : WickContraction n) (i : Fin n.succ) (k : Fin n.succ) :
k ∈ (insert c i none).uncontracted ↔
k = i ∃ j, k = i.succAbove j ∧ j ∈ c.uncontracted := by
by_cases hi : k = i
· subst hi
simp
· simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hi
obtain ⟨z, hk⟩ := hi
subst hk
have hn : ¬ i.succAbove z = i := by exact Fin.succAbove_ne i z
simp only [Nat.succ_eq_add_one, insert_succAbove_mem_uncontracted_iff, hn, false_or]
apply Iff.intro
· intro h
exact ⟨z, rfl, h⟩
· intro h
obtain ⟨j, hk⟩ := h
have hjk : z = j := Fin.succAbove_right_inj.mp hk.1
subst hjk
exact hk.2
lemma insert_none_uncontracted (c : WickContraction n) (i : Fin n.succ) :
(insert c i none).uncontracted = Insert.insert i (c.uncontracted.map i.succAboveEmb) := by
ext a
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_none_iff, Finset.mem_insert,
Finset.mem_map, Fin.succAboveEmb_apply]
apply Iff.intro
· intro a_1
cases a_1 with
| inl h =>
subst h
simp_all only [true_or]
| inr h_1 =>
obtain ⟨w, h⟩ := h_1
obtain ⟨left, right⟩ := h
subst left
apply Or.inr
apply Exists.intro
· apply And.intro
on_goal 2 => {rfl
}
· simp_all only
· intro a_1
cases a_1 with
| inl h =>
subst h
simp_all only [true_or]
| inr h_1 =>
obtain ⟨w, h⟩ := h_1
obtain ⟨left, right⟩ := h
subst right
apply Or.inr
apply Exists.intro
· apply And.intro
on_goal 2 => {exact left
}
· simp_all only
@[simp]
lemma mem_uncontracted_insert_some_iff (c : WickContraction n) (i : Fin n.succ)
(k : Fin n.succ) (j : c.uncontracted) :
k ∈ (insert c i (some j)).uncontracted ↔
∃ z, k = i.succAbove z ∧ z ∈ c.uncontracted ∧ z ≠ j := by
by_cases hki : k = i
· subst hki
simp only [Nat.succ_eq_add_one, self_not_mem_uncontracted_of_insert_some, ne_eq, false_iff,
not_exists, not_and, Decidable.not_not]
exact fun x hx => False.elim (Fin.ne_succAbove k x hx)
· simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hki
obtain ⟨z, hk⟩ := hki
subst hk
by_cases hjz : j = z
· subst hjz
rw [mem_uncontracted_iff_not_contracted]
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert,
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
or_true, not_true_eq_false, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂,
false_and, ne_eq, false_iff, not_exists, not_and, Decidable.not_not]
intro x
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)]
exact fun a _a => a.symm
· apply Iff.intro
· intro h
use z
simp only [Nat.succ_eq_add_one, ne_eq, true_and]
refine And.intro ?_ (fun a => hjz a.symm)
rw [mem_uncontracted_iff_not_contracted]
intro p hp
rw [mem_uncontracted_iff_not_contracted] at h
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert,
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] at h
have hc := h.2 p hp
rw [Finset.mapEmbedding_apply] at hc
exact (Finset.mem_map' (i.succAboveEmb)).mpr.mt hc
· intro h
obtain ⟨z', hz'1, hz'⟩ := h
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hz'1
subst hz'1
rw [mem_uncontracted_iff_not_contracted]
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert,
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂]
apply And.intro
· rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)]
exact And.intro (Fin.succAbove_ne i z) (fun a => hjz a.symm)
· rw [mem_uncontracted_iff_not_contracted] at hz'
exact fun a ha hc => hz'.1 a ha ((Finset.mem_map' (i.succAboveEmb)).mp hc)
lemma insert_some_uncontracted (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted) :
(insert c i (some j)).uncontracted = (c.uncontracted.erase j).map i.succAboveEmb := by
ext a
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_some_iff, ne_eq, Finset.map_erase,
Fin.succAboveEmb_apply, Finset.mem_erase, Finset.mem_map]
apply Iff.intro
· intro h
obtain ⟨z, h1, h2, h3⟩ := h
subst h1
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)]
simp only [h3, not_false_eq_true, true_and]
use z
· intro h
obtain ⟨z, h1, h2⟩ := h.2
use z
subst h2
simp only [true_and]
obtain ⟨a, ha1, ha2⟩ := h.2
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at ha2
subst ha2
simp_all only [true_and]
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at h
exact h.1
/-!
## Insert and getDual?
-/
lemma insert_none_getDual?_isNone (c : WickContraction n) (i : Fin n.succ) :
((insert c i none).getDual? i).isNone := by
have hi : i ∈ (insert c i none).uncontracted := by
simp
simp only [Nat.succ_eq_add_one, uncontracted, Finset.mem_filter, Finset.mem_univ, true_and] at hi
rw [hi]
simp
@[simp]
lemma insert_succAbove_getDual?_eq_none_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
(insert c i none).getDual? (i.succAbove j) = none ↔ c.getDual? j = none := by
have h1 := insert_succAbove_mem_uncontracted_iff c i j
simpa [uncontracted] using h1
@[simp]
lemma insert_succAbove_getDual?_isSome_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
((insert c i none).getDual? (i.succAbove j)).isSome ↔ (c.getDual? j).isSome := by
rw [← not_iff_not]
simp
@[simp]
lemma insert_succAbove_getDual?_get (c : WickContraction n) (i : Fin n.succ) (j : Fin n)
(h : ((insert c i none).getDual? (i.succAbove j)).isSome) :
((insert c i none).getDual? (i.succAbove j)).get h =
i.succAbove ((c.getDual? j).get (by simpa using h)) := by
have h1 : (insert c i none).getDual? (i.succAbove j) = some
(i.succAbove ((c.getDual? j).get (by simpa using h))) := by
rw [getDual?_eq_some_iff_mem]
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
RelEmbedding.coe_toEmbedding]
use {j, ((c.getDual? j).get (by simpa using h))}
simp only [self_getDual?_get_mem, true_and]
rw [Finset.mapEmbedding_apply]
simp
exact Option.get_of_mem h h1
@[simp]
lemma insert_some_getDual?_eq (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted) :
(insert c i (some j)).getDual? i = some (i.succAbove j) := by
rw [getDual?_eq_some_iff_mem]
simp [insert]
@[simp]
lemma insert_some_getDual?_neq_none (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
(k : Fin n) (hkj : k ≠ j.1) :
(insert c i (some j)).getDual? (i.succAbove k) = none ↔ c.getDual? k = none := by
apply Iff.intro
· intro h
have hk : (i.succAbove k) ∈ (insert c i (some j)).uncontracted := by
simp only [Nat.succ_eq_add_one, uncontracted, Finset.mem_filter, Finset.mem_univ, true_and]
exact h
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_some_iff, ne_eq] at hk
obtain ⟨z, hz1, hz2, hz3⟩ := hk
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hz1
subst hz1
simpa [uncontracted] using hz2
· intro h
have hk : (i.succAbove k) ∈ (insert c i (some j)).uncontracted := by
simp only [Nat.succ_eq_add_one, mem_uncontracted_insert_some_iff, ne_eq]
use k
simp only [hkj, not_false_eq_true, and_true, true_and]
simpa [uncontracted] using h
simpa [uncontracted, -mem_uncontracted_insert_some_iff, ne_eq] using hk
@[simp]
lemma insert_some_getDual?_neq_isSome (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
(k : Fin n) (hkj : k ≠ j.1) :
((insert c i (some j)).getDual? (i.succAbove k)).isSome ↔ (c.getDual? k).isSome := by
rw [← not_iff_not]
simp [hkj]
@[simp]
lemma insert_some_getDual?_neq_isSome_get (c : WickContraction n) (i : Fin n.succ)
(j : c.uncontracted) (k : Fin n) (hkj : k ≠ j.1)
(h : ((insert c i (some j)).getDual? (i.succAbove k)).isSome) :
((insert c i (some j)).getDual? (i.succAbove k)).get h =
i.succAbove ((c.getDual? k).get (by simpa [hkj] using h)) := by
have h1 : ((insert c i (some j)).getDual? (i.succAbove k))
= some (i.succAbove ((c.getDual? k).get (by simpa [hkj] using h))) := by
rw [getDual?_eq_some_iff_mem]
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert, Finset.mem_map,
RelEmbedding.coe_toEmbedding]
apply Or.inr
use { k, ((c.getDual? k).get (by simpa [hkj] using h))}
simp only [self_getDual?_get_mem, true_and]
rw [Finset.mapEmbedding_apply]
simp
exact Option.get_of_mem h h1
@[simp]
lemma insert_some_getDual?_of_neq (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
(k : Fin n) (hkj : k ≠ j.1) :
(insert c i (some j)).getDual? (i.succAbove k) =
Option.map i.succAbove (c.getDual? k) := by
by_cases h : (c.getDual? k).isSome
· have h1 : (c.insert i (some j)).getDual? (i.succAbove k) =
some (i.succAbove ((c.getDual? k).get h)) := by
rw [← insert_some_getDual?_neq_isSome_get c i j k hkj]
refine Eq.symm (Option.some_get ?_)
simpa [hkj] using h
rw [h1]
have h2 :(c.getDual? k) = some ((c.getDual? k).get h) := by simp
conv_rhs => rw [h2]
rw [@Option.map_coe']
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at h
simp only [Nat.succ_eq_add_one, h, Option.map_none']
simp only [ne_eq, hkj, not_false_eq_true, insert_some_getDual?_neq_none]
exact h
/-!
## Interaction with erase.
-/
@[simp]
lemma insert_erase (c : WickContraction n) (i : Fin n.succ) (j : Option (c.uncontracted)) :
erase (insert c i j) i = c := by
refine Subtype.eq ?_
simp only [erase, Nat.succ_eq_add_one, insert, Finset.le_eq_subset]
conv_rhs => rw [c.eq_filter_mem_self]
refine Finset.filter_inj'.mpr ?_
intro a _
match j with
| none =>
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
apply Iff.intro
· intro ha
obtain ⟨a', ha', ha''⟩ := ha
rw [Finset.mapEmbedding_apply] at ha''
simp only [Finset.map_inj] at ha''
subst ha''
exact ha'
· intro ha
use a
simp only [ha, true_and]
rw [Finset.mapEmbedding_apply]
| some j =>
simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
apply Iff.intro
· intro ha
rcases ha with ha | ha
· have hin : ¬ i ∈ Finset.map i.succAboveEmb a := by
simp only [Nat.succ_eq_add_one, Finset.mem_map, Fin.succAboveEmb_apply, not_exists,
not_and]
intro x
exact fun a => Fin.succAbove_ne i x
refine False.elim (hin ?_)
rw [ha]
simp
· obtain ⟨a', ha', ha''⟩ := ha
rw [Finset.mapEmbedding_apply] at ha''
simp only [Finset.map_inj] at ha''
subst ha''
exact ha'
· intro ha
apply Or.inr
use a
simp only [ha, true_and]
rw [Finset.mapEmbedding_apply]
lemma insert_getDualErase (c : WickContraction n) (i : Fin n.succ) (j : Option c.uncontracted) :
(insert c i j).getDualErase i =
uncontractedCongr (c := c) (c' := (c.insert i j).erase i) (by simp) j := by
match n with
| 0 =>
simp only [insert, Nat.succ_eq_add_one, Nat.reduceAdd, Finset.le_eq_subset, getDualErase]
fin_cases j
simp
| Nat.succ n =>
match j with
| none =>
have hi := insert_none_getDual?_isNone c i
simp [getDualErase, hi]
| some j =>
simp only [Nat.succ_eq_add_one, getDualErase, insert_some_getDual?_eq, Option.isSome_some,
↓reduceDIte, Option.get_some, predAboveI_succAbove, uncontractedCongr_some, Option.some.injEq]
rfl
@[simp]
lemma erase_insert (c : WickContraction n.succ) (i : Fin n.succ) :
insert (erase c i) i (getDualErase c i) = c := by
match n with
| 0 =>
apply Subtype.eq
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, insert, getDualErase, Finset.le_eq_subset]
ext a
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
apply Iff.intro
· intro h
simp only [erase, Nat.reduceAdd, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ,
true_and] at h
obtain ⟨a', ha', ha''⟩ := h
subst ha''
exact ha'
· intro ha
obtain ⟨a, ha⟩ := c.mem_not_eq_erase_of_isNone (a := a) i (by simp) ha
simp_all only [Nat.succ_eq_add_one]
obtain ⟨left, right⟩ := ha
subst right
apply Exists.intro
· apply And.intro
on_goal 2 => {rfl
}
· simp_all only
| Nat.succ n =>
apply Subtype.eq
by_cases hi : (c.getDual? i).isSome
· rw [insert_of_isSome]
simp only [Nat.succ_eq_add_one, getDualErase, hi, ↓reduceDIte, Option.get_some,
Finset.le_eq_subset]
rw [succsAbove_predAboveI]
ext a
apply Iff.intro
· simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
intro ha
rcases ha with ha | ha
· subst ha
simp
· obtain ⟨a', ha', ha''⟩ := ha
subst ha''
simpa [erase] using ha'
· intro ha
simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
by_cases hia : a = {i, (c.getDual? i).get hi}
· subst hia
simp
· apply Or.inr
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and]
obtain ⟨a', ha'⟩ := c.mem_not_eq_erase_of_isSome (a := a) i hi ha hia
use a'
simp_all only [Nat.succ_eq_add_one, true_and]
obtain ⟨left, right⟩ := ha'
subst right
rfl
simp only [Nat.succ_eq_add_one, ne_eq, self_neq_getDual?_get, not_false_eq_true]
exact (getDualErase_isSome_iff_getDual?_isSome c i).mpr hi
· simp only [Nat.succ_eq_add_one, insert, getDualErase, hi, Bool.false_eq_true, ↓reduceDIte,
Finset.le_eq_subset]
ext a
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
apply Iff.intro
· intro h
simp only [erase, Nat.succ_eq_add_one, Finset.mem_filter, Finset.mem_univ, true_and] at h
obtain ⟨a', ha', ha''⟩ := h
subst ha''
exact ha'
· intro ha
obtain ⟨a, ha⟩ := c.mem_not_eq_erase_of_isNone (a := a) i (by simpa using hi) ha
simp_all only [Nat.succ_eq_add_one, Bool.not_eq_true, Option.not_isSome,
Option.isNone_iff_eq_none]
obtain ⟨left, right⟩ := ha
subst right
apply Exists.intro
· apply And.intro
on_goal 2 => {rfl
}
· simp_all only
/-- Lifts a contraction in `c` to a contraction in `(c.insert i j)`. -/
def insertLift {c : WickContraction n} (i : Fin n.succ) (j : Option (c.uncontracted))
(a : c.1) : (c.insert i j).1 := ⟨a.1.map (Fin.succAboveEmb i), by
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset]
match j with
| none =>
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
use a
simp only [a.2, true_and]
rfl
| some j =>
simp only [Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding]
apply Or.inr
use a
simp only [a.2, true_and]
rfl⟩
lemma insertLift_injective {c : WickContraction n} (i : Fin n.succ) (j : Option (c.uncontracted)) :
Function.Injective (insertLift i j) := by
intro a b hab
simp only [Nat.succ_eq_add_one, insertLift, Subtype.mk.injEq, Finset.map_inj] at hab
exact Subtype.eq hab
lemma insertLift_none_surjective {c : WickContraction n} (i : Fin n.succ) :
Function.Surjective (c.insertLift i none) := by
intro a
have ha := a.2
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_map,
RelEmbedding.coe_toEmbedding] at ha
obtain ⟨a', ha', ha''⟩ := ha
use ⟨a', ha'⟩
exact Subtype.eq ha''
lemma insertLift_none_bijective {c : WickContraction n} (i : Fin n.succ) :
Function.Bijective (c.insertLift i none) := by
exact ⟨insertLift_injective i none, insertLift_none_surjective i⟩
@[simp]
lemma insert_fstFieldOfContract (c : WickContraction n) (i : Fin n.succ)
(j : Option (c.uncontracted)) (a : c.1) : (c.insert i j).fstFieldOfContract (insertLift i j a) =
i.succAbove (c.fstFieldOfContract a) := by
refine (c.insert i j).eq_fstFieldOfContract_of_mem (a := (insertLift i j a))
(i := i.succAbove (c.fstFieldOfContract a)) (j := i.succAbove (c.sndFieldOfContract a)) ?_ ?_ ?_
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
use (c.fstFieldOfContract a)
simp
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
use (c.sndFieldOfContract a)
simp
· refine Fin.succAbove_lt_succAbove_iff.mpr ?_
exact fstFieldOfContract_lt_sndFieldOfContract c a
@[simp]
lemma insert_sndFieldOfContract (c : WickContraction n) (i : Fin n.succ)
(j : Option (c.uncontracted)) (a : c.1) : (c.insert i j).sndFieldOfContract (insertLift i j a) =
i.succAbove (c.sndFieldOfContract a) := by
refine (c.insert i j).eq_sndFieldOfContract_of_mem (a := (insertLift i j a))
(i := i.succAbove (c.fstFieldOfContract a)) (j := i.succAbove (c.sndFieldOfContract a)) ?_ ?_ ?_
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
use (c.fstFieldOfContract a)
simp
· simp only [Nat.succ_eq_add_one, insertLift, Finset.mem_map, Fin.succAboveEmb_apply]
use (c.sndFieldOfContract a)
simp
· refine Fin.succAbove_lt_succAbove_iff.mpr ?_
exact fstFieldOfContract_lt_sndFieldOfContract c a
/-- Given a contracted pair for a Wick contraction `WickContraction n`, the
corresponding contracted pair of a wick contraction `(c.insert i (some j))` formed
by inserting an element `i` into the contraction. -/
def insertLiftSome {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted)
(a : Unit ⊕ c.1) : (c.insert i (some j)).1 :=
match a with
| Sum.inl () => ⟨{i, i.succAbove j}, by
simp [insert]⟩
| Sum.inr a => c.insertLift i j a
lemma insertLiftSome_injective {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted) :
Function.Injective (insertLiftSome i j) := by
intro a b hab
match a, b with
| Sum.inl (), Sum.inl () => rfl
| Sum.inl (), Sum.inr a =>
simp only [Nat.succ_eq_add_one, insertLiftSome, insertLift, Subtype.mk.injEq] at hab
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hab
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton] at hab
have hi : i ∈ ({i.succAbove (c.fstFieldOfContract a),
i.succAbove (c.sndFieldOfContract a)} : Finset (Fin (n + 1))) := by
rw [← hab]
simp
simp only [Nat.succ_eq_add_one, Finset.mem_insert, Finset.mem_singleton] at hi
rcases hi with hi | hi
· exact False.elim (Fin.ne_succAbove _ _ hi)
· exact False.elim (Fin.ne_succAbove _ _ hi)
| Sum.inr a, Sum.inl () =>
simp only [Nat.succ_eq_add_one, insertLiftSome, insertLift, Subtype.mk.injEq] at hab
rw [finset_eq_fstFieldOfContract_sndFieldOfContract] at hab
simp only [Finset.map_insert, Fin.succAboveEmb_apply, Finset.map_singleton] at hab
have hi : i ∈ ({i.succAbove (c.fstFieldOfContract a),
i.succAbove (c.sndFieldOfContract a)} : Finset (Fin (n + 1))) := by
rw [hab]
simp
simp only [Nat.succ_eq_add_one, Finset.mem_insert, Finset.mem_singleton] at hi
rcases hi with hi | hi
· exact False.elim (Fin.ne_succAbove _ _ hi)
· exact False.elim (Fin.ne_succAbove _ _ hi)
| Sum.inr a, Sum.inr b =>
simp only [Nat.succ_eq_add_one, insertLiftSome] at hab
simpa using insertLift_injective i (some j) hab
lemma insertLiftSome_surjective {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted) :
Function.Surjective (insertLiftSome i j) := by
intro a
have ha := a.2
simp only [Nat.succ_eq_add_one, insert, Finset.le_eq_subset, Finset.mem_insert, Finset.mem_map,
RelEmbedding.coe_toEmbedding] at ha
rcases ha with ha | ha
· use Sum.inl ()
exact Subtype.eq ha.symm
· obtain ⟨a', ha', ha''⟩ := ha
use Sum.inr ⟨a', ha'⟩
simp only [Nat.succ_eq_add_one, insertLiftSome, insertLift]
exact Subtype.eq ha''
lemma insertLiftSome_bijective {c : WickContraction n} (i : Fin n.succ) (j : c.uncontracted) :
Function.Bijective (insertLiftSome i j) :=
⟨insertLiftSome_injective i j, insertLiftSome_surjective i j⟩
end WickContraction

View file

@ -0,0 +1,293 @@
/-
Copyright (c) 2025 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.WickContraction.UncontractedList
/-!
# Inserting an element into a contraction
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open HepLean.Fin
/-!
## Inserting an element into a list
-/
/-- Given a Wick contraction `c` associated to a list `φs`,
a position `i : Fin n.succ`, an element `φ`, and an optional uncontracted element
`j : Option (c.uncontracted)` of `c`.
The Wick contraction associated with `(φs.insertIdx i φ).length` formed by 'inserting' `φ`
into `φs` after the first `i` elements and contracting it optionally with j. -/
def insertList (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option (c.uncontracted)) :
WickContraction (φs.insertIdx i φ).length :=
congr (by simp) (c.insert i j)
@[simp]
lemma insertList_fstFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option (c.uncontracted))
(a : c.1) : (insertList φ φs c i j).fstFieldOfContract
(congrLift (insertIdx_length_fin φ φs i).symm (insertLift i j a)) =
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove (c.fstFieldOfContract a)) := by
simp [insertList]
@[simp]
lemma insertList_sndFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option (c.uncontracted))
(a : c.1) : (insertList φ φs c i j).sndFieldOfContract
(congrLift (insertIdx_length_fin φ φs i).symm (insertLift i j a)) =
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove (c.sndFieldOfContract a)) := by
simp [insertList]
@[simp]
lemma insertList_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
(insertList φ φs c i (some j)).fstFieldOfContract
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩) =
if i < i.succAbove j.1 then
finCongr (insertIdx_length_fin φ φs i).symm i else
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) := by
split
· rename_i h
refine (insertList φ φs c i (some j)).eq_fstFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
· simp [congrLift]
· simp [congrLift]
· rw [Fin.lt_def] at h ⊢
simp_all
· rename_i h
refine (insertList φ φs c i (some j)).eq_fstFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
· simp [congrLift]
· simp [congrLift]
· rw [Fin.lt_def] at h ⊢
simp_all only [Nat.succ_eq_add_one, Fin.val_fin_lt, not_lt, finCongr_apply, Fin.coe_cast]
have hi : i.succAbove j ≠ i := by exact Fin.succAbove_ne i j
omega
/-!
## insertList and getDual?
-/
@[simp]
lemma insertList_none_getDual?_self (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) :
(insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm i) = none := by
simp only [Nat.succ_eq_add_one, insertList, getDual?_congr, finCongr_apply, Fin.cast_trans,
Fin.cast_eq_self, Option.map_eq_none']
have h1 := c.insert_none_getDual?_isNone i
simpa using h1
lemma insertList_isSome_getDual?_self (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
((insertList φ φs c i (some j)).getDual?
(Fin.cast (insertIdx_length_fin φ φs i).symm i)).isSome := by
simp [insertList, getDual?_congr]
lemma insertList_some_getDual?_self_not_none (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
¬ ((insertList φ φs c i (some j)).getDual?
(Fin.cast (insertIdx_length_fin φ φs i).symm i)) = none := by
simp [insertList, getDual?_congr]
@[simp]
lemma insertList_some_getDual?_self_eq (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
((insertList φ φs c i (some j)).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm i))
= some (Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j)) := by
simp [insertList, getDual?_congr]
@[simp]
lemma insertList_some_getDual?_some_eq (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
((insertList φ φs c i (some j)).getDual?
(Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j)))
= some (Fin.cast (insertIdx_length_fin φ φs i).symm i) := by
rw [getDual?_eq_some_iff_mem]
rw [@Finset.pair_comm]
rw [← getDual?_eq_some_iff_mem]
simp
@[simp]
lemma insertList_none_succAbove_getDual?_eq_none_iff (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length) :
(insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
(i.succAbove j)) = none ↔ c.getDual? j = none := by
simp [insertList, getDual?_congr]
@[simp]
lemma insertList_some_succAbove_getDual?_eq_option (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length)
(k : c.uncontracted) (hkj : j ≠ k.1) :
(insertList φ φs c i (some k)).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
(i.succAbove j)) = Option.map (Fin.cast (insertIdx_length_fin φ φs i).symm ∘ i.succAbove)
(c.getDual? j) := by
simp only [Nat.succ_eq_add_one, insertList, getDual?_congr, finCongr_apply, Fin.cast_trans,
Fin.cast_eq_self, ne_eq, hkj, not_false_eq_true, insert_some_getDual?_of_neq, Option.map_map]
rfl
@[simp]
lemma insertList_none_succAbove_getDual?_isSome_iff (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length) :
((insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
(i.succAbove j))).isSome ↔ (c.getDual? j).isSome := by
rw [← not_iff_not]
simp
@[simp]
lemma insertList_none_getDual?_get_eq (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length)
(h : ((insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
(i.succAbove j))).isSome) :
((insertList φ φs c i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
(i.succAbove j))).get h = Fin.cast (insertIdx_length_fin φ φs i).symm
(i.succAbove ((c.getDual? j).get (by simpa using h))) := by
simp [insertList, getDual?_congr_get]
/-........................................... -/
@[simp]
lemma insertList_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
(insertList φ φs c i (some j)).sndFieldOfContract
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩) =
if i < i.succAbove j.1 then
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) else
finCongr (insertIdx_length_fin φ φs i).symm i := by
split
· rename_i h
refine (insertList φ φs c i (some j)).eq_sndFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
· simp [congrLift]
· simp [congrLift]
· rw [Fin.lt_def] at h ⊢
simp_all
· rename_i h
refine (insertList φ φs c i (some j)).eq_sndFieldOfContract_of_mem
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insert]⟩)
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
· simp [congrLift]
· simp [congrLift]
· rw [Fin.lt_def] at h ⊢
simp_all only [Nat.succ_eq_add_one, Fin.val_fin_lt, not_lt, finCongr_apply, Fin.coe_cast]
have hi : i.succAbove j ≠ i := by exact Fin.succAbove_ne i j
omega
lemma insertList_none_prod_contractions (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ)
(f : (c.insertList φ φs i none).1 → M) [CommMonoid M] :
∏ a, f a = ∏ (a : c.1), f (congrLift (insertIdx_length_fin φ φs i).symm
(insertLift i none a)) := by
let e1 := Equiv.ofBijective (c.insertLift i none) (insertLift_none_bijective i)
let e2 := Equiv.ofBijective (congrLift (insertIdx_length_fin φ φs i).symm)
((c.insert i none).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
erw [← e2.prod_comp]
erw [← e1.prod_comp]
rfl
lemma insertList_some_prod_contractions (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted)
(f : (c.insertList φ φs i (some j)).1 → M) [CommMonoid M] :
∏ a, f a = f (congrLift (insertIdx_length_fin φ φs i).symm
⟨{i, i.succAbove j}, by simp [insert]⟩) *
∏ (a : c.1), f (congrLift (insertIdx_length_fin φ φs i).symm (insertLift i (some j) a)) := by
let e2 := Equiv.ofBijective (congrLift (insertIdx_length_fin φ φs i).symm)
((c.insert i (some j)).congrLift_bijective ((insertIdx_length_fin φ φs i).symm))
erw [← e2.prod_comp]
let e1 := Equiv.ofBijective (c.insertLiftSome i j) (insertLiftSome_bijective i j)
erw [← e1.prod_comp]
rw [Fintype.prod_sum_type]
simp only [Finset.univ_unique, PUnit.default_eq_unit, Nat.succ_eq_add_one, Finset.prod_singleton,
Finset.univ_eq_attach]
rfl
/-- Given a finite set of `Fin φs.length` the finite set of `(φs.insertIdx i φ).length`
formed by mapping elements using `i.succAboveEmb` and `finCongr`. -/
def insertListLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) :
Finset (Fin (φs.insertIdx i φ).length) :=
(a.map i.succAboveEmb).map (finCongr (insertIdx_length_fin φ φs i).symm).toEmbedding
@[simp]
lemma self_not_mem_insertListLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) :
Fin.cast (insertIdx_length_fin φ φs i).symm i ∉ insertListLiftFinset φ i a := by
simp only [Nat.succ_eq_add_one, insertListLiftFinset, Finset.mem_map_equiv, finCongr_symm,
finCongr_apply, Fin.cast_trans, Fin.cast_eq_self]
simp only [Finset.mem_map, Fin.succAboveEmb_apply, not_exists, not_and]
intro x
exact fun a => Fin.succAbove_ne i x
lemma succAbove_mem_insertListLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) (j : Fin φs.length) :
Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j) ∈ insertListLiftFinset φ i a ↔
j ∈ a := by
simp only [insertListLiftFinset, Finset.mem_map_equiv, finCongr_symm, finCongr_apply,
Fin.cast_trans, Fin.cast_eq_self]
simp only [Finset.mem_map, Fin.succAboveEmb_apply]
apply Iff.intro
· intro h
obtain ⟨x, hx1, hx2⟩ := h
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hx2
simp_all
· intro h
use j
lemma insert_fin_eq_self (φ : 𝓕.States) {φs : List 𝓕.States}
(i : Fin φs.length.succ) (j : Fin (List.insertIdx i φ φs).length) :
j = Fin.cast (insertIdx_length_fin φ φs i).symm i
∃ k, j = Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove k) := by
obtain ⟨k, hk⟩ := (finCongr (insertIdx_length_fin φ φs i).symm).surjective j
subst hk
by_cases hi : k = i
· simp [hi]
· simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hi
obtain ⟨z, hk⟩ := hi
subst hk
right
use z
rfl
lemma insertList_uncontractedList_none_map (φ : 𝓕.States) {φs : List 𝓕.States}
(c : WickContraction φs.length) (i : Fin φs.length.succ) :
List.map (List.insertIdx (↑i) φ φs).get (insertList φ φs c i none).uncontractedList =
List.insertIdx (c.uncontractedListOrderPos i) φ (List.map φs.get c.uncontractedList) := by
simp only [Nat.succ_eq_add_one, insertList]
rw [congr_uncontractedList]
erw [uncontractedList_extractEquiv_symm_none]
rw [orderedInsert_succAboveEmb_uncontractedList_eq_insertIdx]
rw [insertIdx_map, insertIdx_map]
congr 1
· simp
rw [List.map_map, List.map_map]
congr
conv_rhs => rw [get_eq_insertIdx_succAbove φ φs i]
rfl
lemma insertLift_sum (φ : 𝓕.States) {φs : List 𝓕.States}
(i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) :
∑ c, f c = ∑ (c : WickContraction φs.length), ∑ (k : Option (c.uncontracted)),
f (insertList φ φs c i k) := by
rw [sum_extractEquiv_congr (finCongr (insertIdx_length_fin φ φs i).symm i) f
(insertIdx_length_fin φ φs i)]
rfl
end WickContraction

View file

@ -0,0 +1,162 @@
/-
Copyright (c) 2025 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.WickContraction.Uncontracted
import HepLean.PerturbationTheory.Algebras.StateAlgebra.TimeOrder
import HepLean.PerturbationTheory.Algebras.OperatorAlgebra.TimeContraction
import HepLean.PerturbationTheory.WickContraction.InsertList
/-!
# Involution associated with a contraction
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open FieldStatistic
/-- The involution of `Fin n` associated with a Wick contraction `c : WickContraction n` as follows.
If `i : Fin n` is contracted in `c` then it is taken to its dual, otherwise
it is taken to itself. -/
def toInvolution : {f : Fin n → Fin n // Function.Involutive f} :=
⟨fun i => if h : (c.getDual? i).isSome then (c.getDual? i).get h else i, by
intro i
by_cases h : (c.getDual? i).isSome
· simp [h]
· simp [h]⟩
/-- The Wick contraction formed by an involution `f` of `Fin n` by taking as the contracted
sets of the contraction the orbits of `f` of cardinality `2`. -/
def fromInvolution (f : {f : Fin n → Fin n // Function.Involutive f}) : WickContraction n :=
⟨Finset.univ.filter (fun a => a.card = 2 ∧ ∃ i, {i, f.1 i} = a), by
intro a
simp only [Finset.mem_filter, Finset.mem_univ, true_and, and_imp, forall_exists_index]
intro h1 _ _
exact h1, by
intro a ha b hb
simp only [Finset.mem_filter, Finset.mem_univ, true_and] at ha hb
obtain ⟨i, hai⟩ := ha.2
subst hai
obtain ⟨j, hbj⟩ := hb.2
subst hbj
by_cases h : i = j
· subst h
simp
· by_cases hi : i = f.1 j
· subst hi
simp only [Finset.disjoint_insert_right, Finset.mem_insert, Finset.mem_singleton, not_or,
Finset.disjoint_singleton_right, true_or, not_true_eq_false, and_false, or_false]
rw [f.2]
rw [@Finset.pair_comm]
· apply Or.inr
simp only [Finset.disjoint_insert_right, Finset.mem_insert, Finset.mem_singleton, not_or,
Finset.disjoint_singleton_right]
apply And.intro
· apply And.intro
· exact fun a => h a.symm
· by_contra hn
subst hn
rw [f.2 i] at hi
simp at hi
· apply And.intro
· exact fun a => hi a.symm
· rw [Function.Injective.eq_iff]
exact fun a => h (id (Eq.symm a))
exact Function.Involutive.injective f.2⟩
@[simp]
lemma fromInvolution_getDual?_isSome (f : {f : Fin n → Fin n // Function.Involutive f})
(i : Fin n) : ((fromInvolution f).getDual? i).isSome ↔ f.1 i ≠ i := by
rw [getDual?_isSome_iff]
apply Iff.intro
· intro h
obtain ⟨a, ha⟩ := h
have ha2 := a.2
simp only [fromInvolution, Finset.mem_filter, Finset.mem_univ, true_and] at ha2
obtain ⟨j, h⟩ := ha2.2
rw [← h] at ha
have hj : f.1 j ≠ j := by
by_contra hn
rw [hn] at h
rw [← h] at ha2
simp at ha2
simp only [Finset.mem_insert, Finset.mem_singleton] at ha
rcases ha with ha | ha
· subst ha
exact hj
· subst ha
rw [f.2]
exact id (Ne.symm hj)
· intro hi
use ⟨{i, f.1 i}, by
simp only [fromInvolution, Finset.mem_filter, Finset.mem_univ, exists_apply_eq_apply,
and_true, true_and]
rw [Finset.card_pair (id (Ne.symm hi))]⟩
simp
lemma fromInvolution_getDual?_eq_some (f : {f : Fin n → Fin n // Function.Involutive f}) (i : Fin n)
(h : ((fromInvolution f).getDual? i).isSome) :
((fromInvolution f).getDual? i) = some (f.1 i) := by
rw [@getDual?_eq_some_iff_mem]
simp only [fromInvolution, Finset.mem_filter, Finset.mem_univ, exists_apply_eq_apply, and_true,
true_and]
apply Finset.card_pair
simp only [fromInvolution_getDual?_isSome, ne_eq] at h
exact fun a => h (id (Eq.symm a))
@[simp]
lemma fromInvolution_getDual?_get (f : {f : Fin n → Fin n // Function.Involutive f}) (i : Fin n)
(h : ((fromInvolution f).getDual? i).isSome) :
((fromInvolution f).getDual? i).get h = (f.1 i) := by
have h1 := fromInvolution_getDual?_eq_some f i h
exact Option.get_of_mem h h1
lemma toInvolution_fromInvolution : fromInvolution c.toInvolution = c := by
apply Subtype.eq
simp only [fromInvolution, toInvolution]
ext a
simp only [Finset.mem_filter, Finset.mem_univ, true_and]
apply Iff.intro
· intro h
obtain ⟨i, hi⟩ := h.2
split at hi
· subst hi
simp
· simp only [Finset.mem_singleton, Finset.insert_eq_of_mem] at hi
subst hi
simp at h
· intro ha
apply And.intro (c.2.1 a ha)
use c.fstFieldOfContract ⟨a, ha⟩
simp only [fstFieldOfContract_getDual?, Option.isSome_some, ↓reduceDIte, Option.get_some]
change _ = (⟨a, ha⟩ : c.1).1
conv_rhs => rw [finset_eq_fstFieldOfContract_sndFieldOfContract]
lemma fromInvolution_toInvolution (f : {f : Fin n → Fin n // Function.Involutive f}) :
(fromInvolution f).toInvolution = f := by
apply Subtype.eq
funext i
simp only [toInvolution, ne_eq, dite_not]
split
· rename_i h
simp
· rename_i h
simp only [fromInvolution_getDual?_isSome, ne_eq, Decidable.not_not] at h
exact id (Eq.symm h)
/-- The equivalence btween Wick contractions for `n` and involutions of `Fin n`.
The involution of `Fin n` associated with a Wick contraction `c : WickContraction n` as follows.
If `i : Fin n` is contracted in `c` then it is taken to its dual, otherwise
it is taken to itself. -/
def equivInvolution : WickContraction n ≃ {f : Fin n → Fin n // Function.Involutive f} where
toFun := toInvolution
invFun := fromInvolution
left_inv := toInvolution_fromInvolution
right_inv := fromInvolution_toInvolution
end WickContraction

View file

@ -0,0 +1,64 @@
/-
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.Mathematics.Fin.Involutions
import HepLean.PerturbationTheory.WickContraction.Involutions
/-!
# Full contraction
We say that a contraction is full if it has no uncontracted fields.
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open FieldStatistic
/-- A contraction is full if there are no uncontracted fields, i.e. the finite set
of uncontracted fields is empty. -/
def IsFull : Prop := c.uncontracted = ∅
/-- The condition on whether or not a contraction is full is decidable. -/
instance : Decidable (IsFull c) := decEq c.uncontracted ∅
lemma isFull_iff_equivInvolution_no_fixed_point :
IsFull c ↔ ∀ (i : Fin n), (equivInvolution c).1 i ≠ i := by
simp only [IsFull, ne_eq]
rw [Finset.eq_empty_iff_forall_not_mem]
simp [equivInvolution, toInvolution, uncontracted]
/-- The equivalence between full contractions and fixed-point free involutions. -/
def isFullInvolutionEquiv : {c : WickContraction n // IsFull c} ≃
{f : Fin n → Fin n // Function.Involutive f ∧ (∀ i, f i ≠ i)} where
toFun c := ⟨equivInvolution c.1, by
apply And.intro (equivInvolution c.1).2
rw [← isFull_iff_equivInvolution_no_fixed_point]
exact c.2⟩
invFun f := ⟨equivInvolution.symm ⟨f.1, f.2.1⟩, by
rw [isFull_iff_equivInvolution_no_fixed_point]
simpa using f.2.2⟩
left_inv c := by simp
right_inv f := by simp
open Nat
/-- 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
rw [Fintype.card_congr (isFullInvolutionEquiv)]
exact HepLean.Fin.involutionNoFixed_card_even n he
/-- If `n` is odd then there are no full contractions. This is because
there will always be at least one element unpaired. -/
theorem card_of_isfull_odd (ho : Odd n) :
Fintype.card {c : WickContraction n // IsFull c} = 0 := by
rw [Fintype.card_congr (isFullInvolutionEquiv)]
exact HepLean.Fin.involutionNoFixed_card_odd n ho
end WickContraction

File diff suppressed because it is too large Load diff

View file

@ -0,0 +1,165 @@
/-
Copyright (c) 2025 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.WickContraction.Sign
import HepLean.PerturbationTheory.Algebras.OperatorAlgebra.TimeContraction
/-!
# Time contractions
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
/-!
## Time contract.
-/
/-- Given a Wick contraction `c` associated with a list `φs`, the
product of all time-contractions of pairs of contracted elements in `φs`,
as a member of the center of `𝓞.A`. -/
noncomputable def timeContract (𝓞 : 𝓕.OperatorAlgebra) {φs : List 𝓕.States}
(c : WickContraction φs.length) :
Subalgebra.center 𝓞.A :=
∏ (a : c.1), ⟨𝓞.timeContract (φs.get (c.fstFieldOfContract a)) (φs.get (c.sndFieldOfContract a)),
𝓞.timeContract_mem_center _ _⟩
@[simp]
lemma timeContract_insertList_none (𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) :
(c.insertList φ φs i none).timeContract 𝓞 = c.timeContract 𝓞 := by
rw [timeContract, insertList_none_prod_contractions]
congr
ext a
simp
lemma timeConract_insertList_some (𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (j : c.uncontracted) :
(c.insertList φ φs i (some j)).timeContract 𝓞 =
(if i < i.succAbove j then
⟨𝓞.timeContract φ φs[j.1], 𝓞.timeContract_mem_center _ _⟩
else ⟨𝓞.timeContract φs[j.1] φ, 𝓞.timeContract_mem_center _ _⟩) * c.timeContract 𝓞 := by
rw [timeContract, insertList_some_prod_contractions]
congr 1
· simp only [Nat.succ_eq_add_one, insertList_fstFieldOfContract_some_incl, finCongr_apply,
List.get_eq_getElem, insertList_sndFieldOfContract_some_incl, Fin.getElem_fin]
split
· simp
· simp
· congr
ext a
simp
open FieldStatistic
lemma timeConract_insertList_some_eq_mul_contractStateAtIndex_lt
(𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (k : c.uncontracted)
(ht : 𝓕.timeOrderRel φ φs[k.1]) (hik : i < i.succAbove k) :
(c.insertList φ φs i (some k)).timeContract 𝓞 =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (c.uncontracted.filter (fun x => x < k))⟩)
• (𝓞.contractStateAtIndex φ (List.map φs.get c.uncontractedList)
((uncontractedStatesEquiv φs c) (some k)) * c.timeContract 𝓞) := by
rw [timeConract_insertList_some]
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
OperatorAlgebra.contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
List.getElem_map, uncontractedList_getElem_uncontractedFinEquiv_symm, List.get_eq_getElem,
Algebra.smul_mul_assoc]
· simp only [hik, ↓reduceIte, MulMemClass.coe_mul]
rw [𝓞.timeContract_of_timeOrderRel]
trans (1 : ) • (𝓞.crAnF ((CrAnAlgebra.superCommute
(CrAnAlgebra.anPart (StateAlgebra.ofState φ))) (CrAnAlgebra.ofState φs[k.1])) *
↑(timeContract 𝓞 c))
· simp
simp only [smul_smul]
congr
have h1 : ofList 𝓕.statesStatistic (List.take (↑(c.uncontractedFinEquiv.symm k))
(List.map φs.get c.uncontractedList))
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) c.uncontracted)⟩) := by
simp only [ofFinset]
congr
rw [← List.map_take]
congr
rw [take_uncontractedFinEquiv_symm]
rw [filter_uncontractedList]
rw [h1]
simp only [exchangeSign_mul_self]
· exact ht
lemma timeConract_insertList_some_eq_mul_contractStateAtIndex_not_lt
(𝓞 : 𝓕.OperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
(c : WickContraction φs.length) (i : Fin φs.length.succ) (k : c.uncontracted)
(ht : ¬ 𝓕.timeOrderRel φs[k.1] φ) (hik : ¬ i < i.succAbove k) :
(c.insertList φ φs i (some k)).timeContract 𝓞 =
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (c.uncontracted.filter (fun x => x ≤ k))⟩)
• (𝓞.contractStateAtIndex φ (List.map φs.get c.uncontractedList)
((uncontractedStatesEquiv φs c) (some k)) * c.timeContract 𝓞) := by
rw [timeConract_insertList_some]
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
OperatorAlgebra.contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
List.getElem_map, uncontractedList_getElem_uncontractedFinEquiv_symm, List.get_eq_getElem,
Algebra.smul_mul_assoc]
simp only [hik, ↓reduceIte, MulMemClass.coe_mul]
rw [𝓞.timeContract_of_not_timeOrderRel, 𝓞.timeContract_of_timeOrderRel]
simp only [instCommGroup.eq_1, Algebra.smul_mul_assoc, smul_smul]
congr
have h1 : ofList 𝓕.statesStatistic (List.take (↑(c.uncontractedFinEquiv.symm k))
(List.map φs.get c.uncontractedList))
= (𝓕 |>ₛ ⟨φs.get, (Finset.filter (fun x => x < k) c.uncontracted)⟩) := by
simp only [ofFinset]
congr
rw [← List.map_take]
congr
rw [take_uncontractedFinEquiv_symm, filter_uncontractedList]
rw [h1]
trans 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, {k.1}⟩)
· rw [exchangeSign_symm, ofFinset_singleton]
simp
rw [← map_mul]
congr
rw [ofFinset_union]
congr
ext a
simp only [Finset.mem_singleton, Finset.mem_sdiff, Finset.mem_union, Finset.mem_filter,
Finset.mem_inter, not_and, not_lt, and_imp]
apply Iff.intro
· intro h
subst h
simp
· intro h
have h1 := h.1
rcases h1 with h1 | h1
· have h2' := h.2 h1.1 h1.2 h1.1
omega
· have h2' := h.2 h1.1 (by omega) h1.1
omega
have ht := IsTotal.total (r := timeOrderRel) φs[k.1] φ
simp_all only [Fin.getElem_fin, Nat.succ_eq_add_one, not_lt, false_or]
exact ht
lemma timeContract_of_not_gradingCompliant (𝓞 : 𝓕.OperatorAlgebra) (φs : List 𝓕.States)
(c : WickContraction φs.length) (h : ¬ GradingCompliant φs c) :
c.timeContract 𝓞 = 0 := by
rw [timeContract]
simp only [GradingCompliant, Fin.getElem_fin, Subtype.forall, not_forall] at h
obtain ⟨a, ha⟩ := h
obtain ⟨ha, ha2⟩ := ha
apply Finset.prod_eq_zero (i := ⟨a, ha⟩)
simp only [Finset.univ_eq_attach, Finset.mem_attach]
apply Subtype.eq
simp only [List.get_eq_getElem, ZeroMemClass.coe_zero]
rw [OperatorAlgebra.timeContract_zero_of_diff_grade]
simp [ha2]
end WickContraction

View file

@ -0,0 +1,67 @@
/-
Copyright (c) 2025 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.WickContraction.Basic
/-!
# Uncontracted elements
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
/-- Given a Wick contraction, the finset of elements of `Fin n` which are not contracted. -/
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) :
(c.congr h).uncontracted = Finset.map (finCongr h).toEmbedding c.uncontracted := by
subst h
simp
/-- The equivalence of `Option c.uncontracted` for two propositionally equal Wick contractions. -/
def uncontractedCongr {c c': WickContraction n} (h : c = c') :
Option c.uncontracted ≃ Option c'.uncontracted :=
Equiv.optionCongr (Equiv.subtypeEquivRight (by rw [h]; simp))
@[simp]
lemma uncontractedCongr_none {c c': WickContraction n} (h : c = c') :
(uncontractedCongr h) none = none := by
simp [uncontractedCongr]
@[simp]
lemma uncontractedCongr_some {c c': WickContraction n} (h : c = c') (i : c.uncontracted) :
(uncontractedCongr h) (some i) = some (Equiv.subtypeEquivRight (by rw [h]; simp) i) := by
simp [uncontractedCongr]
lemma mem_uncontracted_iff_not_contracted (i : Fin n) :
i ∈ c.uncontracted ↔ ∀ p ∈ c.1, i ∉ p := by
simp only [uncontracted, getDual?, Finset.mem_filter, Finset.mem_univ, true_and]
apply Iff.intro
· intro h p hp
have hp := c.2.1 p hp
rw [Finset.card_eq_two] at hp
obtain ⟨a, b, ha, hb, hab⟩ := hp
rw [Fin.find_eq_none_iff] at h
by_contra hn
simp only [Finset.mem_insert, Finset.mem_singleton] at hn
rcases hn with hn | hn
· subst hn
exact h b hp
· subst hn
rw [Finset.pair_comm] at hp
exact h a hp
· intro h
rw [Fin.find_eq_none_iff]
by_contra hn
simp only [not_forall, Decidable.not_not] at hn
obtain ⟨j, hj⟩ := hn
apply h {i, j} hj
simp
end WickContraction

View file

@ -0,0 +1,490 @@
/-
Copyright (c) 2025 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.WickContraction.ExtractEquiv
/-!
# List of uncontracted elements
-/
open FieldStruct
variable {𝓕 : FieldStruct}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open HepLean.Fin
/-!
## Some properties of lists of fin
-/
lemma fin_list_sorted_monotone_sorted {n m : } (l: List (Fin n)) (hl : l.Sorted (· ≤ ·))
(f : Fin n → Fin m) (hf : StrictMono f) : ((List.map f l)).Sorted (· ≤ ·) := by
induction l with
| nil => simp
| cons a l ih =>
simp only [List.map_cons, List.sorted_cons, List.mem_map, forall_exists_index, and_imp,
forall_apply_eq_imp_iff₂]
apply And.intro
· simp only [List.sorted_cons] at hl
intro b hb
have hl1 := hl.1 b hb
exact (StrictMono.le_iff_le hf).mpr hl1
· simp only [List.sorted_cons] at hl
exact ih hl.2
lemma fin_list_sorted_succAboveEmb_sorted (l: List (Fin n)) (hl : l.Sorted (· ≤ ·))
(i : Fin n.succ) : ((List.map i.succAboveEmb l)).Sorted (· ≤ ·) := by
apply fin_list_sorted_monotone_sorted
exact hl
simp only [Fin.coe_succAboveEmb]
exact Fin.strictMono_succAbove i
lemma fin_list_sorted_split :
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : ) →
l = l.filter (fun x => x.1 < i) ++ l.filter (fun x => i ≤ x.1)
| [], _, _ => by simp
| a :: l, hl, i => by
simp only [List.sorted_cons] at hl
by_cases ha : a < i
· conv_lhs => rw [fin_list_sorted_split l hl.2 i]
rw [← List.cons_append]
rw [List.filter_cons_of_pos, List.filter_cons_of_neg]
simp only [decide_eq_true_eq, not_le, ha]
simp [ha]
· have hx : List.filter (fun x => decide (x.1 < i)) (a :: l) = [] := by
simp only [ha, decide_false, Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg,
List.filter_eq_nil_iff, decide_eq_true_eq, not_lt]
intro b hb
have hb' := hl.1 b hb
omega
simp only [hx, List.nil_append]
rw [List.filter_cons_of_pos]
simp only [List.cons.injEq, true_and]
have hl' := fin_list_sorted_split l hl.2 i
have hx : List.filter (fun x => decide (x.1 < i)) (l) = [] := by
simp only [List.filter_eq_nil_iff, decide_eq_true_eq, not_lt]
intro b hb
have hb' := hl.1 b hb
omega
simp only [hx, List.nil_append] at hl'
conv_lhs => rw [hl']
simp only [decide_eq_true_eq]
omega
lemma fin_list_sorted_indexOf_filter_le_mem :
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : Fin n) →
(hl : i ∈ l) →
List.indexOf i (List.filter (fun x => decide (↑i ≤ ↑x)) l) = 0
| [], _, _, _ => by simp
| a :: l, hl, i, hi => by
simp only [List.sorted_cons] at hl
by_cases ha : i ≤ a
· simp only [ha, decide_true, List.filter_cons_of_pos]
have ha : a = i := by
simp only [List.mem_cons] at hi
rcases hi with hi | hi
· subst hi
rfl
· have hl' := hl.1 i hi
exact Fin.le_antisymm hl' ha
subst ha
simp
· simp only [not_le] at ha
rw [List.filter_cons_of_neg (by simpa using ha)]
rw [fin_list_sorted_indexOf_filter_le_mem l hl.2]
simp only [List.mem_cons] at hi
rcases hi with hi | hi
· omega
· exact hi
lemma fin_list_sorted_indexOf_mem :
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : Fin n) →
(hi : i ∈ l) →
l.indexOf i = (l.filter (fun x => x.1 < i.1)).length := by
intro l hl i hi
conv_lhs => rw [fin_list_sorted_split l hl i]
rw [List.indexOf_append_of_not_mem]
erw [fin_list_sorted_indexOf_filter_le_mem l hl i hi]
· simp
· simp
lemma orderedInsert_of_fin_list_sorted :
(l : List (Fin n)) → (hl : l.Sorted (· ≤ ·)) → (i : Fin n) →
List.orderedInsert (· ≤ ·) i l = l.filter (fun x => x.1 < i.1) ++
i :: l.filter (fun x => i.1 ≤ x.1)
| [], _, _ => by simp
| a :: l, hl, i => by
simp only [List.sorted_cons] at hl
by_cases ha : i ≤ a
· simp only [List.orderedInsert, ha, ↓reduceIte, Fin.val_fin_lt, decide_eq_true_eq, not_lt,
List.filter_cons_of_neg, Fin.val_fin_le, decide_true, List.filter_cons_of_pos]
have h1 : List.filter (fun x => decide (↑x < ↑i)) l = [] := by
simp only [List.filter_eq_nil_iff, decide_eq_true_eq, not_lt]
intro a ha
have ha' := hl.1 a ha
omega
have hl : l = List.filter (fun x => decide (i ≤ x)) l := by
conv_lhs => rw [fin_list_sorted_split l hl.2 i]
simp [h1]
simp [← hl, h1]
· simp only [List.orderedInsert, ha, ↓reduceIte, Fin.val_fin_lt, Fin.val_fin_le, decide_false,
Bool.false_eq_true, not_false_eq_true, List.filter_cons_of_neg]
rw [List.filter_cons_of_pos]
rw [orderedInsert_of_fin_list_sorted l hl.2 i]
simp only [Fin.val_fin_lt, Fin.val_fin_le, List.cons_append]
simp only [decide_eq_true_eq]
omega
lemma orderedInsert_eq_insertIdx_of_fin_list_sorted (l : List (Fin n)) (hl : l.Sorted (· ≤ ·))
(i : Fin n) :
List.orderedInsert (· ≤ ·) i l = l.insertIdx (l.filter (fun x => x.1 < i.1)).length i := by
let n : Fin l.length.succ := ⟨(List.filter (fun x => decide (x < i)) l).length, by
have h1 := l.length_filter_le (fun x => x.1 < i.1)
simp only [Fin.val_fin_lt] at h1
omega⟩
simp only [Fin.val_fin_lt]
conv_rhs => rw [insertIdx_eq_take_drop _ _ n]
rw [orderedInsert_of_fin_list_sorted]
congr
· conv_rhs =>
rhs
rw [fin_list_sorted_split l hl i]
simp [n]
· conv_rhs =>
rhs
rw [fin_list_sorted_split l hl i]
simp [n]
exact hl
/-!
## Uncontracted List
-/
/-- Given a Wick contraction `c`, the ordered list of elements of `Fin n` which are not contracted,
i.e. do not appera anywhere in `c.1`. -/
def uncontractedList : List (Fin n) := List.filter (fun x => x ∈ c.uncontracted) (List.finRange n)
lemma uncontractedList_mem_iff (i : Fin n) :
i ∈ c.uncontractedList ↔ i ∈ c.uncontracted := by
simp [uncontractedList]
@[simp]
lemma nil_zero_uncontractedList : (empty (n := 0)).uncontractedList = [] := by
simp [empty, uncontractedList]
lemma congr_uncontractedList {n m : } (h : n = m) (c : WickContraction n) :
((congr h) c).uncontractedList = List.map (finCongr h) c.uncontractedList := by
subst h
simp [congr]
lemma uncontractedList_get_mem_uncontracted (i : Fin c.uncontractedList.length) :
c.uncontractedList.get i ∈ c.uncontracted := by
rw [← uncontractedList_mem_iff]
simp
lemma uncontractedList_sorted : List.Sorted (· ≤ ·) c.uncontractedList := by
rw [uncontractedList]
apply List.Sorted.filter
rw [← List.ofFn_id]
exact Monotone.ofFn_sorted fun ⦃a b⦄ a => a
lemma uncontractedList_nodup : c.uncontractedList.Nodup := by
rw [uncontractedList]
refine List.Nodup.filter (fun x => decide (x ∈ c.uncontracted)) ?_
exact List.nodup_finRange n
lemma uncontractedList_toFinset (c : WickContraction n) :
c.uncontractedList.toFinset = c.uncontracted := by
simp [uncontractedList]
lemma uncontractedList_eq_sort (c : WickContraction n) :
c.uncontractedList = c.uncontracted.sort (· ≤ ·) := by
symm
rw [← uncontractedList_toFinset]
refine (List.toFinset_sort (α := Fin n) (· ≤ ·) ?_).mpr ?_
· exact uncontractedList_nodup c
· exact uncontractedList_sorted c
lemma uncontractedList_length_eq_card (c : WickContraction n) :
c.uncontractedList.length = c.uncontracted.card := by
rw [uncontractedList_eq_sort]
exact Finset.length_sort fun x1 x2 => x1 ≤ x2
lemma uncontractedList_succAboveEmb_sorted (c : WickContraction n) (i : Fin n.succ) :
((List.map i.succAboveEmb c.uncontractedList)).Sorted (· ≤ ·) := by
apply fin_list_sorted_succAboveEmb_sorted
exact uncontractedList_sorted c
lemma uncontractedList_succAboveEmb_nodup (c : WickContraction n) (i : Fin n.succ) :
((List.map i.succAboveEmb c.uncontractedList)).Nodup := by
refine List.Nodup.map ?_ ?_
· exact Function.Embedding.injective i.succAboveEmb
· exact uncontractedList_nodup c
lemma uncontractedList_succAboveEmb_toFinset (c : WickContraction n) (i : Fin n.succ) :
(List.map i.succAboveEmb c.uncontractedList).toFinset =
(Finset.map i.succAboveEmb c.uncontracted) := by
ext a
simp only [Fin.coe_succAboveEmb, List.mem_toFinset, List.mem_map, Finset.mem_map,
Fin.succAboveEmb_apply]
rw [← c.uncontractedList_toFinset]
simp
lemma uncontractedList_succAboveEmb_eq_sort(c : WickContraction n) (i : Fin n.succ) :
(List.map i.succAboveEmb c.uncontractedList) =
(c.uncontracted.map i.succAboveEmb).sort (· ≤ ·) := by
rw [← uncontractedList_succAboveEmb_toFinset]
symm
refine (List.toFinset_sort (α := Fin n.succ) (· ≤ ·) ?_).mpr ?_
· exact uncontractedList_succAboveEmb_nodup c i
· exact uncontractedList_succAboveEmb_sorted c i
lemma uncontractedList_succAboveEmb_eraseIdx_sorted (c : WickContraction n) (i : Fin n.succ)
(k: ) : ((List.map i.succAboveEmb c.uncontractedList).eraseIdx k).Sorted (· ≤ ·) := by
apply HepLean.List.eraseIdx_sorted
exact uncontractedList_succAboveEmb_sorted c i
lemma uncontractedList_succAboveEmb_eraseIdx_nodup (c : WickContraction n) (i : Fin n.succ) (k: ) :
((List.map i.succAboveEmb c.uncontractedList).eraseIdx k).Nodup := by
refine List.Nodup.eraseIdx k ?_
exact uncontractedList_succAboveEmb_nodup c i
lemma uncontractedList_succAboveEmb_eraseIdx_toFinset (c : WickContraction n) (i : Fin n.succ)
(k : ) (hk : k < c.uncontractedList.length) :
((List.map i.succAboveEmb c.uncontractedList).eraseIdx k).toFinset =
(c.uncontracted.map i.succAboveEmb).erase (i.succAboveEmb c.uncontractedList[k]) := by
ext a
simp only [Fin.coe_succAboveEmb, List.mem_toFinset, Fin.succAboveEmb_apply, Finset.mem_erase,
ne_eq, Finset.mem_map]
rw [mem_eraseIdx_nodup _ _ _ (by simpa using hk)]
simp_all only [List.mem_map, List.getElem_map, ne_eq]
apply Iff.intro
· intro a_1
simp_all only [not_false_eq_true, true_and]
obtain ⟨left, right⟩ := a_1
obtain ⟨w, h⟩ := left
obtain ⟨left, right_1⟩ := h
subst right_1
use w
simp_all [uncontractedList]
· intro a_1
simp_all only [not_false_eq_true, and_true]
obtain ⟨left, right⟩ := a_1
obtain ⟨w, h⟩ := right
obtain ⟨left_1, right⟩ := h
subst right
use w
simp_all [uncontractedList]
exact uncontractedList_succAboveEmb_nodup c i
lemma uncontractedList_succAboveEmb_eraseIdx_eq_sort (c : WickContraction n) (i : Fin n.succ)
(k : ) (hk : k < c.uncontractedList.length) :
((List.map i.succAboveEmb c.uncontractedList).eraseIdx k) =
((c.uncontracted.map i.succAboveEmb).erase
(i.succAboveEmb c.uncontractedList[k])).sort (· ≤ ·) := by
rw [← uncontractedList_succAboveEmb_eraseIdx_toFinset]
symm
refine (List.toFinset_sort (α := Fin n.succ) (· ≤ ·) ?_).mpr ?_
· exact uncontractedList_succAboveEmb_eraseIdx_nodup c i k
· exact uncontractedList_succAboveEmb_eraseIdx_sorted c i k
lemma uncontractedList_succAbove_orderedInsert_sorted (c : WickContraction n) (i : Fin n.succ) :
(List.orderedInsert (· ≤ ·) i
(List.map i.succAboveEmb c.uncontractedList)).Sorted (· ≤ ·) := by
refine List.Sorted.orderedInsert i (List.map (⇑i.succAboveEmb) c.uncontractedList) ?_
exact uncontractedList_succAboveEmb_sorted c i
lemma uncontractedList_succAbove_orderedInsert_nodup (c : WickContraction n) (i : Fin n.succ) :
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)).Nodup := by
have h1 : (List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)).Perm
(i :: List.map i.succAboveEmb c.uncontractedList) := by
exact List.perm_orderedInsert (fun x1 x2 => x1 ≤ x2) i _
apply List.Perm.nodup h1.symm
simp only [Nat.succ_eq_add_one, List.nodup_cons, List.mem_map, not_exists,
not_and]
apply And.intro
· intro x _
exact Fin.succAbove_ne i x
· exact uncontractedList_succAboveEmb_nodup c i
lemma uncontractedList_succAbove_orderedInsert_toFinset (c : WickContraction n) (i : Fin n.succ) :
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)).toFinset =
(Insert.insert i (Finset.map i.succAboveEmb c.uncontracted)) := by
ext a
simp only [Nat.succ_eq_add_one, Fin.coe_succAboveEmb, List.mem_toFinset, List.mem_orderedInsert,
List.mem_map, Finset.mem_insert, Finset.mem_map, Fin.succAboveEmb_apply]
rw [← uncontractedList_toFinset]
simp
lemma uncontractedList_succAbove_orderedInsert_eq_sort (c : WickContraction n) (i : Fin n.succ) :
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)) =
(Insert.insert i (Finset.map i.succAboveEmb c.uncontracted)).sort (· ≤ ·) := by
rw [← uncontractedList_succAbove_orderedInsert_toFinset]
symm
refine (List.toFinset_sort (α := Fin n.succ) (· ≤ ·) ?_).mpr ?_
· exact uncontractedList_succAbove_orderedInsert_nodup c i
· exact uncontractedList_succAbove_orderedInsert_sorted c i
lemma uncontractedList_extractEquiv_symm_none (c : WickContraction n) (i : Fin n.succ) :
((extractEquiv i).symm ⟨c, none⟩).uncontractedList =
List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList) := by
rw [uncontractedList_eq_sort, extractEquiv_symm_none_uncontracted]
rw [uncontractedList_succAbove_orderedInsert_eq_sort]
/-- Given a Wick contraction `c : WickContraction n` and a `Fin n.succ`, the number of elements
of `c.uncontractedList` which are less then `i`.
Suppose we want to insert into `c` at position `i`, then this is the position we would
need to insert into `c.uncontractedList`. -/
def uncontractedListOrderPos (c : WickContraction n) (i : Fin n.succ) : :=
(List.filter (fun x => x.1 < i.1) c.uncontractedList).length
@[simp]
lemma uncontractedListOrderPos_lt_length_add_one (c : WickContraction n) (i : Fin n.succ) :
c.uncontractedListOrderPos i < c.uncontractedList.length + 1 := by
simp only [uncontractedListOrderPos, Nat.succ_eq_add_one]
have h1 := c.uncontractedList.length_filter_le (fun x => x.1 < i.1)
omega
lemma take_uncontractedListOrderPos_eq_filter (c : WickContraction n) (i : Fin n.succ) :
(c.uncontractedList.take (c.uncontractedListOrderPos i)) =
c.uncontractedList.filter (fun x => x.1 < i.1) := by
nth_rewrite 1 [fin_list_sorted_split c.uncontractedList _ i]
simp only [uncontractedListOrderPos, Nat.succ_eq_add_one, List.take_left']
exact uncontractedList_sorted c
lemma take_uncontractedListOrderPos_eq_filter_sort (c : WickContraction n) (i : Fin n.succ) :
(c.uncontractedList.take (c.uncontractedListOrderPos i)) =
(c.uncontracted.filter (fun x => x.1 < i.1)).sort (· ≤ ·) := by
rw [take_uncontractedListOrderPos_eq_filter]
have h1 : (c.uncontractedList.filter (fun x => x.1 < i.1)).Sorted (· ≤ ·) := by
apply List.Sorted.filter
exact uncontractedList_sorted c
have h2 : (c.uncontractedList.filter (fun x => x.1 < i.1)).Nodup := by
refine List.Nodup.filter _ ?_
exact uncontractedList_nodup c
have h3 : (c.uncontractedList.filter (fun x => x.1 < i.1)).toFinset =
(c.uncontracted.filter (fun x => x.1 < i.1)) := by
rw [uncontractedList_eq_sort]
simp
rw [← h3]
exact ((List.toFinset_sort (α := Fin n) (· ≤ ·) h2).mpr h1).symm
lemma orderedInsert_succAboveEmb_uncontractedList_eq_insertIdx (c : WickContraction n)
(i : Fin n.succ) :
(List.orderedInsert (· ≤ ·) i (List.map i.succAboveEmb c.uncontractedList)) =
(List.map i.succAboveEmb c.uncontractedList).insertIdx (uncontractedListOrderPos c i) i := by
rw [orderedInsert_eq_insertIdx_of_fin_list_sorted]
swap
exact uncontractedList_succAboveEmb_sorted c i
congr 1
simp only [Nat.succ_eq_add_one, Fin.val_fin_lt, Fin.coe_succAboveEmb, uncontractedListOrderPos]
rw [List.filter_map]
simp only [List.length_map]
congr
funext x
simp only [Function.comp_apply, Fin.succAbove, decide_eq_decide]
split
simp only [Fin.lt_def, Fin.coe_castSucc]
rename_i h
simp_all only [Fin.lt_def, Fin.coe_castSucc, not_lt, Fin.val_succ]
omega
/-- The equivalence between the positions of `c.uncontractedList` i.e. elements of
`Fin (c.uncontractedList).length` and the finite set `c.uncontracted` considered as a finite type.
-/
def uncontractedFinEquiv (c : WickContraction n) :
Fin (c.uncontractedList).length ≃ c.uncontracted where
toFun i := ⟨c.uncontractedList.get i, c.uncontractedList_get_mem_uncontracted i⟩
invFun i := ⟨List.indexOf i.1 c.uncontractedList,
List.indexOf_lt_length.mpr ((c.uncontractedList_mem_iff i.1).mpr i.2)⟩
left_inv i := by
ext
exact List.get_indexOf (uncontractedList_nodup c) _
right_inv i := by
ext
simp
@[simp]
lemma uncontractedList_getElem_uncontractedFinEquiv_symm (k : c.uncontracted) :
c.uncontractedList[(c.uncontractedFinEquiv.symm k).val] = k := by
simp [uncontractedFinEquiv]
lemma uncontractedFinEquiv_symm_eq_filter_length (k : c.uncontracted) :
(c.uncontractedFinEquiv.symm k).val =
(List.filter (fun i => i < k.val) c.uncontractedList).length := by
simp only [uncontractedFinEquiv, List.get_eq_getElem, Equiv.coe_fn_symm_mk]
rw [fin_list_sorted_indexOf_mem]
· simp
· exact uncontractedList_sorted c
· rw [uncontractedList_mem_iff]
exact k.2
lemma take_uncontractedFinEquiv_symm (k : c.uncontracted) :
c.uncontractedList.take (c.uncontractedFinEquiv.symm k).val =
c.uncontractedList.filter (fun i => i < k.val) := by
have hl := fin_list_sorted_split c.uncontractedList (uncontractedList_sorted c) k.val
conv_lhs =>
rhs
rw [hl]
rw [uncontractedFinEquiv_symm_eq_filter_length]
simp
/-- The equivalence between the type `Option c.uncontracted` for `WickContraction φs.length` and
`Option (Fin (c.uncontractedList.map φs.get).length)`, that is optional positions of
`c.uncontractedList.map φs.get` induced by `uncontractedFinEquiv`. -/
def uncontractedStatesEquiv (φs : List 𝓕.States) (c : WickContraction φs.length) :
Option c.uncontracted ≃ Option (Fin (c.uncontractedList.map φs.get).length) :=
Equiv.optionCongr (c.uncontractedFinEquiv.symm.trans (finCongr (by simp)))
@[simp]
lemma uncontractedStatesEquiv_none (φs : List 𝓕.States) (c : WickContraction φs.length) :
(uncontractedStatesEquiv φs c).toFun none = none := by
simp [uncontractedStatesEquiv]
lemma uncontractedStatesEquiv_list_sum [AddCommMonoid α] (φs : List 𝓕.States)
(c : WickContraction φs.length) (f : Option (Fin (c.uncontractedList.map φs.get).length) → α) :
∑ (i : Option (Fin (c.uncontractedList.map φs.get).length)), f i =
∑ (i : Option c.uncontracted), f (c.uncontractedStatesEquiv φs i) := by
rw [(c.uncontractedStatesEquiv φs).sum_comp]
lemma uncontractedList_extractEquiv_symm_some (c : WickContraction n) (i : Fin n.succ)
(k : c.uncontracted) : ((extractEquiv i).symm ⟨c, some k⟩).uncontractedList =
((c.uncontractedList).map i.succAboveEmb).eraseIdx (c.uncontractedFinEquiv.symm k) := by
rw [uncontractedList_eq_sort]
rw [uncontractedList_succAboveEmb_eraseIdx_eq_sort]
swap
simp only [Fin.is_lt]
congr
simp only [Nat.succ_eq_add_one, extractEquiv, Equiv.coe_fn_symm_mk,
uncontractedList_getElem_uncontractedFinEquiv_symm, Fin.succAboveEmb_apply]
rw [insert_some_uncontracted]
ext a
simp
lemma filter_uncontractedList (c : WickContraction n) (p : Fin n → Prop) [DecidablePred p] :
(c.uncontractedList.filter p) = (c.uncontracted.filter p).sort (· ≤ ·) := by
have h1 : (c.uncontractedList.filter p).Sorted (· ≤ ·) := by
apply List.Sorted.filter
exact uncontractedList_sorted c
have h2 : (c.uncontractedList.filter p).Nodup := by
refine List.Nodup.filter _ ?_
exact uncontractedList_nodup c
have h3 : (c.uncontractedList.filter p).toFinset = (c.uncontracted.filter p) := by
ext a
simp only [List.toFinset_filter, decide_eq_true_eq, Finset.mem_filter, List.mem_toFinset,
and_congr_left_iff]
rw [uncontractedList_mem_iff]
simp
have hx := (List.toFinset_sort (· ≤ ·) h2).mpr h1
rw [← hx, h3]
end WickContraction