refactor: Rename States to FieldOps
This commit is contained in:
parent
171e80fc04
commit
8f41de5785
36 changed files with 946 additions and 946 deletions
|
@ -523,10 +523,10 @@ lemma prod_finset_eq_mul_fst_snd (c : WickContraction n) (a : c.1)
|
|||
|
||||
/-- 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) (φsΛ : WickContraction φs.length) :=
|
||||
def GradingCompliant (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) :=
|
||||
∀ (a : φsΛ.1), (𝓕 |>ₛ φs[φsΛ.fstFieldOfContract a]) = (𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||
|
||||
lemma gradingCompliant_congr {φs φs' : List 𝓕.States} (h : φs = φs')
|
||||
lemma gradingCompliant_congr {φs φs' : List 𝓕.FieldOp} (h : φs = φs')
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
GradingCompliant φs φsΛ ↔ GradingCompliant φs' (congr (by simp [h]) φsΛ) := by
|
||||
subst h
|
||||
|
|
|
@ -29,7 +29,7 @@ open HepLean.Fin
|
|||
`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 insertAndContract {φs : List 𝓕.States} (φ : 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
def insertAndContract {φs : List 𝓕.FieldOp} (φ : 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : Option φsΛ.uncontracted) :
|
||||
WickContraction (φs.insertIdx i φ).length :=
|
||||
congr (by simp) (φsΛ.insertAndContractNat i j)
|
||||
|
@ -38,7 +38,7 @@ def insertAndContract {φs : List 𝓕.States} (φ : 𝓕.States) (φsΛ : WickC
|
|||
scoped[WickContraction] notation φs "↩Λ" φ:max i:max j => insertAndContract φ φs i j
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_fstFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_fstFieldOfContract (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option φsΛ.uncontracted)
|
||||
(a : φsΛ.1) : (φsΛ ↩Λ φ i j).fstFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm (insertLift i j a)) =
|
||||
|
@ -46,7 +46,7 @@ lemma insertAndContract_fstFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.S
|
|||
simp [insertAndContract]
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_sndFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_sndFieldOfContract (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : Option (φsΛ.uncontracted))
|
||||
(a : φsΛ.1) : (φsΛ ↩Λ φ i j).sndFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm (insertLift i j a)) =
|
||||
|
@ -54,7 +54,7 @@ lemma insertAndContract_sndFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.S
|
|||
simp [insertAndContract]
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_fstFieldOfContract_some_incl (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(insertAndContract φ φsΛ i (some j)).fstFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
|
@ -92,7 +92,7 @@ lemma insertAndContract_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : L
|
|||
|
||||
-/
|
||||
@[simp]
|
||||
lemma insertAndContract_none_getDual?_self (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_none_getDual?_self (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
(φsΛ ↩Λ φ i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm i) = none := by
|
||||
simp only [Nat.succ_eq_add_one, insertAndContract, getDual?_congr, finCongr_apply, Fin.cast_trans,
|
||||
|
@ -100,27 +100,27 @@ lemma insertAndContract_none_getDual?_self (φ : 𝓕.States) (φs : List 𝓕.S
|
|||
have h1 := φsΛ.insertAndContractNat_none_getDual?_isNone i
|
||||
simpa using h1
|
||||
|
||||
lemma insertAndContract_isSome_getDual?_self (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_isSome_getDual?_self (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
((φsΛ ↩Λ φ i (some j)).getDual?
|
||||
(Fin.cast (insertIdx_length_fin φ φs i).symm i)).isSome := by
|
||||
simp [insertAndContract, getDual?_congr]
|
||||
|
||||
lemma insertAndContract_some_getDual?_self_not_none (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_some_getDual?_self_not_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
¬ ((φsΛ ↩Λ φ i (some j)).getDual?
|
||||
(Fin.cast (insertIdx_length_fin φ φs i).symm i)) = none := by
|
||||
simp [insertAndContract, getDual?_congr]
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_some_getDual?_self_eq (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_some_getDual?_self_eq (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
((φsΛ ↩Λ φ 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 [insertAndContract, getDual?_congr]
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_some_getDual?_some_eq (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_some_getDual?_some_eq (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
((φsΛ ↩Λ φ i (some j)).getDual?
|
||||
(Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j)))
|
||||
|
@ -131,14 +131,14 @@ lemma insertAndContract_some_getDual?_some_eq (φ : 𝓕.States) (φs : List
|
|||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_none_succAbove_getDual?_eq_none_iff (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_none_succAbove_getDual?_eq_none_iff (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length) :
|
||||
(φsΛ ↩Λ φ i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j)) = none ↔ φsΛ.getDual? j = none := by
|
||||
simp [insertAndContract, getDual?_congr]
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_some_succAbove_getDual?_eq_option (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_some_succAbove_getDual?_eq_option (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length)
|
||||
(k : φsΛ.uncontracted) (hkj : j ≠ k.1) :
|
||||
(φsΛ ↩Λ φ i (some k)).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
|
@ -150,7 +150,7 @@ lemma insertAndContract_some_succAbove_getDual?_eq_option (φ : 𝓕.States) (φ
|
|||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_none_succAbove_getDual?_isSome_iff (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_none_succAbove_getDual?_isSome_iff (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length) :
|
||||
((φsΛ ↩Λ φ i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j))).isSome ↔ (φsΛ.getDual? j).isSome := by
|
||||
|
@ -158,7 +158,7 @@ lemma insertAndContract_none_succAbove_getDual?_isSome_iff (φ : 𝓕.States) (
|
|||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_none_getDual?_get_eq (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_none_getDual?_get_eq (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : Fin φs.length)
|
||||
(h : ((φsΛ ↩Λ φ i none).getDual? (Fin.cast (insertIdx_length_fin φ φs i).symm
|
||||
(i.succAbove j))).isSome) :
|
||||
|
@ -169,7 +169,7 @@ lemma insertAndContract_none_getDual?_get_eq (φ : 𝓕.States) (φs : List 𝓕
|
|||
|
||||
/-........................................... -/
|
||||
@[simp]
|
||||
lemma insertAndContract_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_sndFieldOfContract_some_incl (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(φsΛ ↩Λ φ i (some j)).sndFieldOfContract
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
|
@ -201,7 +201,7 @@ lemma insertAndContract_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : L
|
|||
have hi : i.succAbove j ≠ i := Fin.succAbove_ne i j
|
||||
omega
|
||||
|
||||
lemma insertAndContract_none_prod_contractions (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_none_prod_contractions (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ)
|
||||
(f : (φsΛ ↩Λ φ i none).1 → M) [CommMonoid M] :
|
||||
∏ a, f a = ∏ (a : φsΛ.1), f (congrLift (insertIdx_length_fin φ φs i).symm
|
||||
|
@ -213,7 +213,7 @@ lemma insertAndContract_none_prod_contractions (φ : 𝓕.States) (φs : List
|
|||
erw [← e1.prod_comp]
|
||||
rfl
|
||||
|
||||
lemma insertAndContract_some_prod_contractions (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma insertAndContract_some_prod_contractions (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted)
|
||||
(f : (φsΛ ↩Λ φ i (some j)).1 → M) [CommMonoid M] :
|
||||
∏ a, f a = f (congrLift (insertIdx_length_fin φ φs i).symm
|
||||
|
@ -231,13 +231,13 @@ lemma insertAndContract_some_prod_contractions (φ : 𝓕.States) (φs : List
|
|||
|
||||
/-- 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 insertAndContractLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
def insertAndContractLiftFinset (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(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_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
lemma self_not_mem_insertAndContractLiftFinset (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) :
|
||||
Fin.cast (insertIdx_length_fin φ φs i).symm i ∉ insertAndContractLiftFinset φ i a := by
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractLiftFinset, Finset.mem_map_equiv, finCongr_symm,
|
||||
|
@ -246,7 +246,7 @@ lemma self_not_mem_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List
|
|||
intro x
|
||||
exact fun a => Fin.succAbove_ne i x
|
||||
|
||||
lemma succAbove_mem_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
lemma succAbove_mem_insertAndContractLiftFinset (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(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)
|
||||
∈ insertAndContractLiftFinset φ i a ↔ j ∈ a := by
|
||||
|
@ -261,7 +261,7 @@ lemma succAbove_mem_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List
|
|||
· intro h
|
||||
use j
|
||||
|
||||
lemma insert_fin_eq_self (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
lemma insert_fin_eq_self (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(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
|
||||
|
@ -276,7 +276,7 @@ lemma insert_fin_eq_self (φ : 𝓕.States) {φs : List 𝓕.States}
|
|||
use z
|
||||
rfl
|
||||
|
||||
lemma insertLift_sum (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
lemma insertLift_sum (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) :
|
||||
∑ c, f c =
|
||||
∑ (φsΛ : WickContraction φs.length), ∑ (k : Option φsΛ.uncontracted), f (φsΛ ↩Λ φ i k) := by
|
||||
|
@ -289,7 +289,7 @@ lemma insertLift_sum (φ : 𝓕.States) {φs : List 𝓕.States}
|
|||
## Uncontracted list
|
||||
|
||||
-/
|
||||
lemma insertAndContract_uncontractedList_none_map (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
lemma insertAndContract_uncontractedList_none_map (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
[φsΛ ↩Λ φ i none]ᵘᶜ = List.insertIdx (φsΛ.uncontractedListOrderPos i) φ [φsΛ]ᵘᶜ := by
|
||||
simp only [Nat.succ_eq_add_one, insertAndContract, uncontractedListGet]
|
||||
|
@ -305,14 +305,14 @@ lemma insertAndContract_uncontractedList_none_map (φ : 𝓕.States) {φs : List
|
|||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_uncontractedList_none_zero (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
lemma insertAndContract_uncontractedList_none_zero (φ : 𝓕.FieldOp) {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
[φsΛ ↩Λ φ 0 none]ᵘᶜ = φ :: [φsΛ]ᵘᶜ := by
|
||||
rw [insertAndContract_uncontractedList_none_map]
|
||||
simp [uncontractedListOrderPos]
|
||||
|
||||
open FieldStatistic in
|
||||
lemma stat_ofFinset_of_insertAndContractLiftFinset (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma stat_ofFinset_of_insertAndContractLiftFinset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) :
|
||||
(𝓕 |>ₛ ⟨(φs.insertIdx i φ).get, insertAndContractLiftFinset φ i a⟩) = 𝓕 |>ₛ ⟨φs.get, a⟩ := by
|
||||
simp only [ofFinset, Nat.succ_eq_add_one]
|
||||
|
|
|
@ -26,7 +26,7 @@ open FieldOpAlgebra
|
|||
/-- Given a Wick contraction `φsΛ` on `φs` and a Wick contraction `φsucΛ` on the uncontracted fields
|
||||
within `φsΛ`, a Wick contraction on `φs`consisting of the contractins in `φsΛ` and
|
||||
the contractions in `φsucΛ`. -/
|
||||
def join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
def join {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) : WickContraction φs.length :=
|
||||
⟨φsΛ.1 ∪ φsucΛ.1.map (Finset.mapEmbedding uncontractedListEmd).toEmbedding, by
|
||||
intro a ha
|
||||
|
@ -61,7 +61,7 @@ def join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
|||
rw [Finset.disjoint_map]
|
||||
exact φsucΛ.2.2 a ha b hb⟩
|
||||
|
||||
lemma join_congr {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma join_congr {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} {φsΛ' : WickContraction φs.length}
|
||||
(h1 : φsΛ = φsΛ') :
|
||||
join φsΛ φsucΛ = join φsΛ' (congr (by simp [h1]) φsucΛ) := by
|
||||
|
@ -70,11 +70,11 @@ lemma join_congr {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
|||
|
||||
/-- Given a contracting pair within `φsΛ` the corresponding contracting pair within
|
||||
`(join φsΛ φsucΛ)`. -/
|
||||
def joinLiftLeft {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
def joinLiftLeft {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : φsΛ.1 → (join φsΛ φsucΛ).1 :=
|
||||
fun a => ⟨a, by simp [join]⟩
|
||||
|
||||
lemma jointLiftLeft_injective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma jointLiftLeft_injective {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} :
|
||||
Function.Injective (@joinLiftLeft _ _ φsΛ φsucΛ) := by
|
||||
intro a b h
|
||||
|
@ -84,7 +84,7 @@ lemma jointLiftLeft_injective {φs : List 𝓕.States} {φsΛ : WickContraction
|
|||
|
||||
/-- Given a contracting pair within `φsucΛ` the corresponding contracting pair within
|
||||
`(join φsΛ φsucΛ)`. -/
|
||||
def joinLiftRight {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
def joinLiftRight {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : φsucΛ.1 → (join φsΛ φsucΛ).1 :=
|
||||
fun a => ⟨a.1.map uncontractedListEmd, by
|
||||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
|
@ -94,7 +94,7 @@ def joinLiftRight {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
|||
simp only [Finset.coe_mem, true_and]
|
||||
rfl⟩
|
||||
|
||||
lemma joinLiftRight_injective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma joinLiftRight_injective {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} :
|
||||
Function.Injective (@joinLiftRight _ _ φsΛ φsucΛ) := by
|
||||
intro a b h
|
||||
|
@ -103,7 +103,7 @@ lemma joinLiftRight_injective {φs : List 𝓕.States} {φsΛ : WickContraction
|
|||
simp only [Finset.map_inj] at h
|
||||
refine Subtype.eq h
|
||||
|
||||
lemma jointLiftLeft_disjoint_joinLiftRight {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma jointLiftLeft_disjoint_joinLiftRight {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} (a : φsΛ.1) (b : φsucΛ.1) :
|
||||
Disjoint (@joinLiftLeft _ _ _ φsucΛ a).1 (joinLiftRight b).1 := by
|
||||
simp only [joinLiftLeft, joinLiftRight]
|
||||
|
@ -111,7 +111,7 @@ lemma jointLiftLeft_disjoint_joinLiftRight {φs : List 𝓕.States} {φsΛ : Wic
|
|||
apply uncontractedListEmd_finset_disjoint_left
|
||||
exact a.2
|
||||
|
||||
lemma jointLiftLeft_neq_joinLiftRight {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma jointLiftLeft_neq_joinLiftRight {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} (a : φsΛ.1) (b : φsucΛ.1) :
|
||||
joinLiftLeft a ≠ joinLiftRight b := by
|
||||
by_contra hn
|
||||
|
@ -124,13 +124,13 @@ lemma jointLiftLeft_neq_joinLiftRight {φs : List 𝓕.States} {φsΛ : WickCont
|
|||
|
||||
/-- The map from contracted pairs of `φsΛ` and `φsucΛ` to contracted pairs in
|
||||
`(join φsΛ φsucΛ)`. -/
|
||||
def joinLift {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
def joinLift {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : φsΛ.1 ⊕ φsucΛ.1 → (join φsΛ φsucΛ).1 := fun a =>
|
||||
match a with
|
||||
| Sum.inl a => joinLiftLeft a
|
||||
| Sum.inr a => joinLiftRight a
|
||||
|
||||
lemma joinLift_injective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma joinLift_injective {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : Function.Injective (@joinLift _ _ φsΛ φsucΛ) := by
|
||||
intro a b h
|
||||
match a, b with
|
||||
|
@ -149,7 +149,7 @@ lemma joinLift_injective {φs : List 𝓕.States} {φsΛ : WickContraction φs.l
|
|||
have h1 := jointLiftLeft_neq_joinLiftRight b a
|
||||
simp_all
|
||||
|
||||
lemma joinLift_surjective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma joinLift_surjective {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : Function.Surjective (@joinLift _ _ φsΛ φsucΛ) := by
|
||||
intro a
|
||||
have ha2 := a.2
|
||||
|
@ -164,13 +164,13 @@ lemma joinLift_surjective {φs : List 𝓕.States} {φsΛ : WickContraction φs.
|
|||
refine Subtype.eq ?_
|
||||
exact ha3.2
|
||||
|
||||
lemma joinLift_bijective {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma joinLift_bijective {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} : Function.Bijective (@joinLift _ _ φsΛ φsucΛ) := by
|
||||
apply And.intro
|
||||
· exact joinLift_injective
|
||||
· exact joinLift_surjective
|
||||
|
||||
lemma prod_join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma prod_join {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (f : (join φsΛ φsucΛ).1 → M) [CommMonoid M]:
|
||||
∏ (a : (join φsΛ φsucΛ).1), f a = (∏ (a : φsΛ.1), f (joinLiftLeft a)) *
|
||||
∏ (a : φsucΛ.1), f (joinLiftRight a) := by
|
||||
|
@ -179,7 +179,7 @@ lemma prod_join {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
|||
simp only [Fintype.prod_sum_type, Finset.univ_eq_attach]
|
||||
rfl
|
||||
|
||||
lemma joinLiftLeft_or_joinLiftRight_of_mem_join {φs : List 𝓕.States}
|
||||
lemma joinLiftLeft_or_joinLiftRight_of_mem_join {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) {a : Finset (Fin φs.length)}
|
||||
(ha : a ∈ (join φsΛ φsucΛ).1) :
|
||||
|
@ -196,7 +196,7 @@ lemma joinLiftLeft_or_joinLiftRight_of_mem_join {φs : List 𝓕.States}
|
|||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma join_fstFieldOfContract_joinLiftRight {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_fstFieldOfContract_joinLiftRight {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsucΛ.1) :
|
||||
(join φsΛ φsucΛ).fstFieldOfContract (joinLiftRight a) =
|
||||
uncontractedListEmd (φsucΛ.fstFieldOfContract a) := by
|
||||
|
@ -207,7 +207,7 @@ lemma join_fstFieldOfContract_joinLiftRight {φs : List 𝓕.States} (φsΛ : Wi
|
|||
exact fstFieldOfContract_lt_sndFieldOfContract φsucΛ a
|
||||
|
||||
@[simp]
|
||||
lemma join_sndFieldOfContract_joinLiftRight {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_sndFieldOfContract_joinLiftRight {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsucΛ.1) :
|
||||
(join φsΛ φsucΛ).sndFieldOfContract (joinLiftRight a) =
|
||||
uncontractedListEmd (φsucΛ.sndFieldOfContract a) := by
|
||||
|
@ -218,7 +218,7 @@ lemma join_sndFieldOfContract_joinLiftRight {φs : List 𝓕.States} (φsΛ : Wi
|
|||
exact fstFieldOfContract_lt_sndFieldOfContract φsucΛ a
|
||||
|
||||
@[simp]
|
||||
lemma join_fstFieldOfContract_joinLiftLeft {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_fstFieldOfContract_joinLiftLeft {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsΛ.1) :
|
||||
(join φsΛ φsucΛ).fstFieldOfContract (joinLiftLeft a) =
|
||||
(φsΛ.fstFieldOfContract a) := by
|
||||
|
@ -228,7 +228,7 @@ lemma join_fstFieldOfContract_joinLiftLeft {φs : List 𝓕.States} (φsΛ : Wic
|
|||
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ a
|
||||
|
||||
@[simp]
|
||||
lemma join_sndFieldOfContract_joinLift {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_sndFieldOfContract_joinLift {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : φsΛ.1) :
|
||||
(join φsΛ φsucΛ).sndFieldOfContract (joinLiftLeft a) =
|
||||
(φsΛ.sndFieldOfContract a) := by
|
||||
|
@ -237,7 +237,7 @@ lemma join_sndFieldOfContract_joinLift {φs : List 𝓕.States} (φsΛ : WickCon
|
|||
· simp [joinLiftLeft]
|
||||
· exact fstFieldOfContract_lt_sndFieldOfContract φsΛ a
|
||||
|
||||
lemma mem_join_right_iff {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma mem_join_right_iff {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (a : Finset (Fin [φsΛ]ᵘᶜ.length)) :
|
||||
a ∈ φsucΛ.1 ↔ a.map uncontractedListEmd ∈ (join φsΛ φsucΛ).1 := by
|
||||
simp only [join, Finset.le_eq_subset, Finset.mem_union, Finset.mem_map,
|
||||
|
@ -257,7 +257,7 @@ lemma mem_join_right_iff {φs : List 𝓕.States} (φsΛ : WickContraction φs.l
|
|||
subst h2
|
||||
exact ha
|
||||
|
||||
lemma join_card {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma join_card {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{φsucΛ : WickContraction [φsΛ]ᵘᶜ.length} :
|
||||
(join φsΛ φsucΛ).1.card = φsΛ.1.card + φsucΛ.1.card := by
|
||||
simp only [join, Finset.le_eq_subset]
|
||||
|
@ -277,7 +277,7 @@ lemma join_card {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
|||
simp_all
|
||||
|
||||
@[simp]
|
||||
lemma empty_join {φs : List 𝓕.States} (φsΛ : WickContraction [empty (n := φs.length)]ᵘᶜ.length) :
|
||||
lemma empty_join {φs : List 𝓕.FieldOp} (φsΛ : WickContraction [empty (n := φs.length)]ᵘᶜ.length) :
|
||||
join empty φsΛ = congr (by simp) φsΛ := by
|
||||
apply Subtype.ext
|
||||
simp only [join, Finset.le_eq_subset, uncontractedListEmd_empty]
|
||||
|
@ -306,13 +306,13 @@ lemma empty_join {φs : List 𝓕.States} (φsΛ : WickContraction [empty (n :=
|
|||
simp
|
||||
|
||||
@[simp]
|
||||
lemma join_empty {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
lemma join_empty {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
join φsΛ empty = φsΛ := by
|
||||
apply Subtype.ext
|
||||
ext a
|
||||
simp [join, empty]
|
||||
|
||||
lemma join_timeContract {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_timeContract {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).timeContract = φsΛ.timeContract * φsucΛ.timeContract := by
|
||||
simp only [timeContract, List.get_eq_getElem]
|
||||
|
@ -322,7 +322,7 @@ lemma join_timeContract {φs : List 𝓕.States} (φsΛ : WickContraction φs.le
|
|||
funext a
|
||||
simp
|
||||
|
||||
lemma join_staticContract {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_staticContract {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).staticContract = φsΛ.staticContract * φsucΛ.staticContract := by
|
||||
simp only [staticContract, List.get_eq_getElem]
|
||||
|
@ -332,7 +332,7 @@ lemma join_staticContract {φs : List 𝓕.States} (φsΛ : WickContraction φs.
|
|||
funext a
|
||||
simp
|
||||
|
||||
lemma mem_join_uncontracted_of_mem_right_uncontracted {φs : List 𝓕.States}
|
||||
lemma mem_join_uncontracted_of_mem_right_uncontracted {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length)
|
||||
(ha : i ∈ φsucΛ.uncontracted) :
|
||||
|
@ -352,7 +352,7 @@ lemma mem_join_uncontracted_of_mem_right_uncontracted {φs : List 𝓕.States}
|
|||
rw [mem_uncontracted_iff_not_contracted] at ha
|
||||
exact ha p hp
|
||||
|
||||
lemma exists_mem_left_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.States}
|
||||
lemma exists_mem_left_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin φs.length)
|
||||
(ha : i ∈ (join φsΛ φsucΛ).uncontracted) :
|
||||
|
@ -364,7 +364,7 @@ lemma exists_mem_left_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.Sta
|
|||
intro p hp
|
||||
simp_all
|
||||
|
||||
lemma exists_mem_right_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.States}
|
||||
lemma exists_mem_right_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin φs.length)
|
||||
(hi : i ∈ (join φsΛ φsucΛ).uncontracted) :
|
||||
|
@ -385,7 +385,7 @@ lemma exists_mem_right_uncontracted_of_mem_join_uncontracted {φs : List 𝓕.St
|
|||
rw [Finset.mapEmbedding_apply])
|
||||
simpa using hip
|
||||
|
||||
lemma join_uncontractedList {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_uncontractedList {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).uncontractedList = List.map uncontractedListEmd φsucΛ.uncontractedList := by
|
||||
rw [uncontractedList_eq_sort]
|
||||
|
@ -404,7 +404,7 @@ lemma join_uncontractedList {φs : List 𝓕.States} (φsΛ : WickContraction φ
|
|||
· intro a b h
|
||||
exact uncontractedListEmd_strictMono h
|
||||
|
||||
lemma join_uncontractedList_get {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_uncontractedList_get {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
((join φsΛ φsucΛ).uncontractedList).get =
|
||||
φsΛ.uncontractedListEmd ∘ (φsucΛ.uncontractedList).get ∘
|
||||
|
@ -417,7 +417,7 @@ lemma join_uncontractedList_get {φs : List 𝓕.States} (φsΛ : WickContractio
|
|||
ext i
|
||||
simp
|
||||
|
||||
lemma join_uncontractedListGet {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_uncontractedListGet {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).uncontractedListGet = φsucΛ.uncontractedListGet := by
|
||||
simp only [uncontractedListGet, join_uncontractedList, List.map_map, List.map_inj_left,
|
||||
|
@ -428,7 +428,7 @@ lemma join_uncontractedListGet {φs : List 𝓕.States} (φsΛ : WickContraction
|
|||
Function.Embedding.coe_subtype]
|
||||
rfl
|
||||
|
||||
lemma join_uncontractedListEmb {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_uncontractedListEmb {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).uncontractedListEmd =
|
||||
((finCongr (congrArg List.length (join_uncontractedListGet _ _))).toEmbedding.trans
|
||||
|
@ -439,7 +439,7 @@ lemma join_uncontractedListEmb {φs : List 𝓕.States} (φsΛ : WickContraction
|
|||
rw [join_uncontractedList_get]
|
||||
rfl
|
||||
|
||||
lemma join_assoc {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_assoc {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (φsucΛ' : WickContraction [φsΛ.join φsucΛ]ᵘᶜ.length) :
|
||||
join (join φsΛ φsucΛ) (φsucΛ') = join φsΛ (join φsucΛ (congr
|
||||
(congrArg List.length (join_uncontractedListGet _ _)) φsucΛ')) := by
|
||||
|
@ -490,7 +490,7 @@ lemma join_assoc {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
|||
change Finset.map (Equiv.refl _).toEmbedding a = _
|
||||
simp only [Equiv.refl_toEmbedding, Finset.map_refl]
|
||||
|
||||
lemma join_getDual?_apply_uncontractedListEmb_eq_none_iff {φs : List 𝓕.States}
|
||||
lemma join_getDual?_apply_uncontractedListEmb_eq_none_iff {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).getDual? (uncontractedListEmd i) = none
|
||||
|
@ -506,7 +506,7 @@ lemma join_getDual?_apply_uncontractedListEmb_eq_none_iff {φs : List 𝓕.State
|
|||
· intro h
|
||||
exact mem_join_uncontracted_of_mem_right_uncontracted φsΛ φsucΛ i h
|
||||
|
||||
lemma join_getDual?_apply_uncontractedListEmb_isSome_iff {φs : List 𝓕.States}
|
||||
lemma join_getDual?_apply_uncontractedListEmb_isSome_iff {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length) :
|
||||
((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)).isSome
|
||||
|
@ -514,7 +514,7 @@ lemma join_getDual?_apply_uncontractedListEmb_isSome_iff {φs : List 𝓕.States
|
|||
rw [← Decidable.not_iff_not]
|
||||
simp [join_getDual?_apply_uncontractedListEmb_eq_none_iff]
|
||||
|
||||
lemma join_getDual?_apply_uncontractedListEmb_some {φs : List 𝓕.States}
|
||||
lemma join_getDual?_apply_uncontractedListEmb_some {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length)
|
||||
(hi :((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)).isSome) :
|
||||
|
@ -532,7 +532,7 @@ lemma join_getDual?_apply_uncontractedListEmb_some {φs : List 𝓕.States}
|
|||
simp
|
||||
|
||||
@[simp]
|
||||
lemma join_getDual?_apply_uncontractedListEmb {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_getDual?_apply_uncontractedListEmb {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i : Fin [φsΛ]ᵘᶜ.length) :
|
||||
((join φsΛ φsucΛ).getDual? (uncontractedListEmd i)) =
|
||||
Option.map uncontractedListEmd (φsucΛ.getDual? i) := by
|
||||
|
@ -556,7 +556,7 @@ lemma join_getDual?_apply_uncontractedListEmb {φs : List 𝓕.States} (φsΛ :
|
|||
|
||||
section
|
||||
|
||||
variable {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
variable {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
|
||||
lemma join_sub_quot (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
join (subContraction S ha) (quotContraction S ha) = φsΛ := by
|
||||
|
@ -590,7 +590,7 @@ end
|
|||
open FieldStatistic
|
||||
|
||||
@[simp]
|
||||
lemma join_singleton_getDual?_left {φs : List 𝓕.States}
|
||||
lemma join_singleton_getDual?_left {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).getDual? i = some j := by
|
||||
|
@ -598,7 +598,7 @@ lemma join_singleton_getDual?_left {φs : List 𝓕.States}
|
|||
simp [singleton, join]
|
||||
|
||||
@[simp]
|
||||
lemma join_singleton_getDual?_right {φs : List 𝓕.States}
|
||||
lemma join_singleton_getDual?_right {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).getDual? j = some i := by
|
||||
|
@ -609,12 +609,12 @@ lemma join_singleton_getDual?_right {φs : List 𝓕.States}
|
|||
exact Finset.pair_comm j i
|
||||
|
||||
|
||||
lemma exists_contraction_pair_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma exists_contraction_pair_of_card_ge_zero {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(h : 0 < φsΛ.1.card) :
|
||||
∃ a, a ∈ φsΛ.1 := by
|
||||
simpa using h
|
||||
|
||||
lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(h : 0 < φsΛ.1.card) (hc : φsΛ.GradingCompliant) :
|
||||
∃ (i j : Fin φs.length) (h : i < j) (φsucΛ : WickContraction [singleton h]ᵘᶜ.length),
|
||||
φsΛ = join (singleton h) φsucΛ ∧ (𝓕 |>ₛ φs[i]) = (𝓕 |>ₛ φs[j])
|
||||
|
@ -645,7 +645,7 @@ lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : Wi
|
|||
simp only [subContraction, Finset.card_singleton, id_eq, eq_mpr_eq_cast] at h1
|
||||
omega
|
||||
|
||||
lemma join_not_gradingCompliant_of_left_not_gradingCompliant {φs : List 𝓕.States}
|
||||
lemma join_not_gradingCompliant_of_left_not_gradingCompliant {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) (φsucΛ : WickContraction [φsΛ]ᵘᶜ.length)
|
||||
(hc : ¬ φsΛ.GradingCompliant) : ¬ (join φsΛ φsucΛ).GradingCompliant := by
|
||||
simp_all only [GradingCompliant, Fin.getElem_fin, Subtype.forall, not_forall]
|
||||
|
|
|
@ -32,16 +32,16 @@ def signFinset (c : WickContraction n) (i1 i2 : Fin n) : Finset (Fin n) :=
|
|||
the sign associated with `φsΛ` is sign corresponding to the number
|
||||
of fermionic-fermionic exchanges one must do to put elements in contracted pairs
|
||||
of `φsΛ` next to each other. -/
|
||||
def sign (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) : ℂ :=
|
||||
def sign (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) : ℂ :=
|
||||
∏ (a : φsΛ.1), 𝓢(𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a],
|
||||
𝓕 |>ₛ ⟨φs.get, φsΛ.signFinset (φsΛ.fstFieldOfContract a) (φsΛ.sndFieldOfContract a)⟩)
|
||||
|
||||
lemma sign_empty (φs : List 𝓕.States) :
|
||||
lemma sign_empty (φs : List 𝓕.FieldOp) :
|
||||
sign φs empty = 1 := by
|
||||
rw [sign]
|
||||
simp [empty]
|
||||
|
||||
lemma sign_congr {φs φs' : List 𝓕.States} (h : φs = φs') (φsΛ : WickContraction φs.length) :
|
||||
lemma sign_congr {φs φs' : List 𝓕.FieldOp} (h : φs = φs') (φsΛ : WickContraction φs.length) :
|
||||
sign φs' (congr (by simp [h]) φsΛ) = sign φs φsΛ := by
|
||||
subst h
|
||||
rfl
|
||||
|
|
|
@ -19,7 +19,7 @@ variable {n : ℕ} (c : WickContraction n)
|
|||
open HepLean.List
|
||||
open FieldStatistic
|
||||
|
||||
lemma signFinset_insertAndContract_none (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signFinset_insertAndContract_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (i1 i2 : Fin φs.length) :
|
||||
(φsΛ ↩Λ φ i none).signFinset (finCongr (insertIdx_length_fin φ φs i).symm
|
||||
|
@ -81,14 +81,14 @@ lemma signFinset_insertAndContract_none (φ : 𝓕.States) (φs : List 𝓕.Stat
|
|||
|
||||
For each contracted pair `{a1, a2}` in `φsΛ` if `a1 < a2` such that `i` is within the range
|
||||
`a1 < i < a2` we pick up a sign equal to `𝓢(φ, φs[a2])`. -/
|
||||
def signInsertNone (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
def signInsertNone (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) : ℂ :=
|
||||
∏ (a : φsΛ.1),
|
||||
if i.succAbove (φsΛ.fstFieldOfContract a) < i ∧ i < i.succAbove (φsΛ.sndFieldOfContract a) then
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||
else 1
|
||||
|
||||
lemma sign_insert_none (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
lemma sign_insert_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) :
|
||||
(φsΛ ↩Λ φ i none).sign = (φsΛ.signInsertNone φ φs i) * φsΛ.sign := by
|
||||
rw [sign]
|
||||
|
@ -109,7 +109,7 @@ lemma sign_insert_none (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : Wick
|
|||
simp
|
||||
· rw [stat_ofFinset_of_insertAndContractLiftFinset]
|
||||
|
||||
lemma signInsertNone_eq_mul_fst_snd (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertNone_eq_mul_fst_snd (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
φsΛ.signInsertNone φ φs i = ∏ (a : φsΛ.1),
|
||||
(if i.succAbove (φsΛ.fstFieldOfContract a) < i then
|
||||
|
@ -143,7 +143,7 @@ lemma signInsertNone_eq_mul_fst_snd (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
have hx := (Fin.succAbove_lt_succAbove_iff (p := i)).mpr hn
|
||||
omega
|
||||
|
||||
lemma signInsertNone_eq_prod_prod (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertNone_eq_prod_prod (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
|
||||
φsΛ.signInsertNone φ φs i = ∏ (a : φsΛ.1), ∏ (x : a),
|
||||
(if i.succAbove x < i then 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs[x.1]) else 1) := by
|
||||
|
@ -158,12 +158,12 @@ lemma signInsertNone_eq_prod_prod (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
erw [hG a]
|
||||
rfl
|
||||
|
||||
lemma sign_insert_none_zero (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) :
|
||||
lemma sign_insert_none_zero (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) :
|
||||
(φsΛ ↩Λ φ 0 none).sign = φsΛ.sign := by
|
||||
rw [sign_insert_none]
|
||||
simp [signInsertNone]
|
||||
|
||||
lemma signInsertNone_eq_prod_getDual?_Some (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertNone_eq_prod_getDual?_Some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
|
||||
φsΛ.signInsertNone φ φs i = ∏ (x : Fin φs.length),
|
||||
if (φsΛ.getDual? x).isSome then
|
||||
|
@ -190,7 +190,7 @@ lemma signInsertNone_eq_prod_getDual?_Some (φ : 𝓕.States) (φs : List 𝓕.S
|
|||
rfl
|
||||
exact hG
|
||||
|
||||
lemma signInsertNone_eq_filter_map (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertNone_eq_filter_map (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
|
||||
φsΛ.signInsertNone φ φs i =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ((List.filter (fun x => (φsΛ.getDual? x).isSome ∧ i.succAbove x < i)
|
||||
|
@ -216,7 +216,7 @@ lemma signInsertNone_eq_filter_map (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
|
||||
/-- The change in sign when inserting a field `φ` at `i` into `φsΛ` is equal
|
||||
to the sign got by moving `φ` through each field `φ₀…φᵢ₋₁` which has a dual. -/
|
||||
lemma signInsertNone_eq_filterset (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertNone_eq_filterset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (hG : GradingCompliant φs φsΛ) :
|
||||
φsΛ.signInsertNone φ φs i = 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, Finset.univ.filter
|
||||
(fun x => (φsΛ.getDual? x).isSome ∧ i.succAbove x < i)⟩) := by
|
||||
|
|
|
@ -27,7 +27,7 @@ open FieldStatistic
|
|||
-/
|
||||
|
||||
|
||||
lemma stat_ofFinset_eq_one_of_gradingCompliant (φs : List 𝓕.States)
|
||||
lemma stat_ofFinset_eq_one_of_gradingCompliant (φs : List 𝓕.FieldOp)
|
||||
(a : Finset (Fin φs.length)) (φsΛ : WickContraction φs.length) (hg : GradingCompliant φs φsΛ)
|
||||
(hnon : ∀ i, φsΛ.getDual? i = none → i ∉ a)
|
||||
(hsom : ∀ i, (h : (φsΛ.getDual? i).isSome) → i ∈ a → (φsΛ.getDual? i).get h ∈ a) :
|
||||
|
@ -64,7 +64,7 @@ lemma stat_ofFinset_eq_one_of_gradingCompliant (φs : List 𝓕.States)
|
|||
rfl
|
||||
|
||||
|
||||
lemma signFinset_insertAndContract_some (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signFinset_insertAndContract_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (i1 i2 : Fin φs.length)
|
||||
(j : φsΛ.uncontracted) :
|
||||
(φsΛ ↩Λ φ i (some j)).signFinset (finCongr (insertIdx_length_fin φ φs i).symm
|
||||
|
@ -206,7 +206,7 @@ lemma signFinset_insertAndContract_some (φ : 𝓕.States) (φs : List 𝓕.Stat
|
|||
inserting `φ` into `φs` at position `i` and contracting it with `j : c.uncontracted`
|
||||
coming from contractions other then the `i` and `j` contraction but which
|
||||
are effected by this new contraction. -/
|
||||
def signInsertSomeProd (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
def signInsertSomeProd (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) : ℂ :=
|
||||
∏ (a : φsΛ.1),
|
||||
if i.succAbove (φsΛ.fstFieldOfContract a) < i ∧ i < i.succAbove (φsΛ.sndFieldOfContract a) ∧
|
||||
|
@ -223,7 +223,7 @@ def signInsertSomeProd (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : Wick
|
|||
and an `i : Fin φs.length.succ`, the change in sign of the contraction associated with
|
||||
inserting `φ` into `φs` at position `i` and contracting it with `j : c.uncontracted`
|
||||
coming from putting `i` next to `j`. -/
|
||||
def signInsertSomeCoef (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
def signInsertSomeCoef (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) : ℂ :=
|
||||
let a : (φsΛ ↩Λ φ i (some j)).1 := congrLift (insertIdx_length_fin φ φs i).symm
|
||||
⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩;
|
||||
|
@ -235,11 +235,11 @@ def signInsertSomeCoef (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : Wick
|
|||
/-- Given a Wick contraction `φsΛ` associated with a list of states `φs`
|
||||
and an `i : Fin φs.length.succ`, the change in sign of the contraction associated with
|
||||
inserting `φ` into `φs` at position `i` and contracting it with `j : c.uncontracted`. -/
|
||||
def signInsertSome (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
def signInsertSome (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) : ℂ :=
|
||||
signInsertSomeCoef φ φs φsΛ i j * signInsertSomeProd φ φs φsΛ i j
|
||||
|
||||
lemma sign_insert_some (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
lemma sign_insert_some (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(φsΛ ↩Λ φ i (some j)).sign = (φsΛ.signInsertSome φ φs i j) * φsΛ.sign := by
|
||||
rw [sign]
|
||||
|
@ -280,7 +280,7 @@ lemma sign_insert_some (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : Wick
|
|||
rw [stat_ofFinset_of_insertAndContractLiftFinset]
|
||||
simp_all
|
||||
|
||||
lemma signInsertSomeProd_eq_one_if (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSomeProd_eq_one_if (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted)
|
||||
(hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1])) :
|
||||
φsΛ.signInsertSomeProd φ φs i j =
|
||||
|
@ -313,7 +313,7 @@ lemma signInsertSomeProd_eq_one_if (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
implies_true, and_true]
|
||||
omega
|
||||
|
||||
lemma signInsertSomeProd_eq_prod_prod (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSomeProd_eq_prod_prod (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1]))
|
||||
(hg : GradingCompliant φs φsΛ) :
|
||||
|
@ -347,7 +347,7 @@ lemma signInsertSomeProd_eq_prod_prod (φ : 𝓕.States) (φs : List 𝓕.States
|
|||
· omega
|
||||
simp [hφj]
|
||||
|
||||
lemma signInsertSomeProd_eq_prod_fin (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSomeProd_eq_prod_fin (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1]))
|
||||
(hg : GradingCompliant φs φsΛ) :
|
||||
|
@ -380,7 +380,7 @@ lemma signInsertSomeProd_eq_prod_fin (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
simp only [hφj, Fin.getElem_fin]
|
||||
exact hg
|
||||
|
||||
lemma signInsertSomeProd_eq_list (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSomeProd_eq_list (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1]))
|
||||
(hg : GradingCompliant φs φsΛ) :
|
||||
|
@ -414,7 +414,7 @@ lemma signInsertSomeProd_eq_list (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
simp only [hφj, Fin.getElem_fin]
|
||||
exact hg
|
||||
|
||||
lemma signInsertSomeProd_eq_finset (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSomeProd_eq_finset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1]))
|
||||
(hg : GradingCompliant φs φsΛ) :
|
||||
|
@ -442,7 +442,7 @@ lemma signInsertSomeProd_eq_finset (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
simp only [hφj, Fin.getElem_fin]
|
||||
exact hg
|
||||
|
||||
lemma signInsertSomeCoef_if (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
lemma signInsertSomeCoef_if (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1])) :
|
||||
φsΛ.signInsertSomeCoef φ φs i j =
|
||||
if i < i.succAbove j then
|
||||
|
@ -462,7 +462,7 @@ lemma signInsertSomeCoef_if (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ :
|
|||
· simp [hφj]
|
||||
|
||||
lemma stat_signFinset_insert_some_self_fst
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : WickContraction φs.length)
|
||||
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length)
|
||||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(𝓕 |>ₛ ⟨(φs.insertIdx i φ).get,
|
||||
(signFinset (φsΛ ↩Λ φ i (some j)) (finCongr (insertIdx_length_fin φ φs i).symm i)
|
||||
|
@ -538,7 +538,7 @@ lemma stat_signFinset_insert_some_self_fst
|
|||
have hy2 := hy2 h
|
||||
simpa [Option.get_map] using hy2
|
||||
|
||||
lemma stat_signFinset_insert_some_self_snd (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma stat_signFinset_insert_some_self_snd (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(𝓕 |>ₛ ⟨(φs.insertIdx i φ).get,
|
||||
(signFinset (φsΛ ↩Λ φ i (some j))
|
||||
|
@ -618,7 +618,7 @@ lemma stat_signFinset_insert_some_self_snd (φ : 𝓕.States) (φs : List 𝓕.S
|
|||
simp only [Option.get_map, Function.comp_apply, Fin.coe_cast, Fin.val_fin_lt]
|
||||
exact Fin.succAbove_lt_succAbove_iff.mpr hy2
|
||||
|
||||
lemma signInsertSomeCoef_eq_finset (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSomeCoef_eq_finset (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted)
|
||||
(hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1])) : φsΛ.signInsertSomeCoef φ φs i j =
|
||||
if i < i.succAbove j then
|
||||
|
@ -637,7 +637,7 @@ lemma signInsertSomeCoef_eq_finset (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
contracting it with `k` (`k < i`) is equal
|
||||
to the sign got by moving `φ` through each field `φ₀…φᵢ₋₁`
|
||||
multiplied by the sign got moving `φ` through each uncontracted field `φ₀…φₖ`. -/
|
||||
lemma signInsertSome_mul_filter_contracted_of_lt (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSome_mul_filter_contracted_of_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||
(hk : i.succAbove k < i) (hg : GradingCompliant φs φsΛ ∧ (𝓕 |>ₛ φ) = 𝓕 |>ₛ φs[k.1]) :
|
||||
signInsertSome φ φs φsΛ i k *
|
||||
|
@ -744,7 +744,7 @@ lemma signInsertSome_mul_filter_contracted_of_lt (φ : 𝓕.States) (φs : List
|
|||
contracting it with `k` (`i < k`) is equal
|
||||
to the sign got by moving `φ` through each field `φ₀…φᵢ₋₁`
|
||||
multiplied by the sign got moving `φ` through each uncontracted field `φ₀…φₖ₋₁`. -/
|
||||
lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma signInsertSome_mul_filter_contracted_of_not_lt (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||
(hk : ¬ i.succAbove k < i) (hg : GradingCompliant φs φsΛ ∧ (𝓕 |>ₛ φ) = 𝓕 |>ₛ φs[k.1]) :
|
||||
signInsertSome φ φs φsΛ i k *
|
||||
|
|
|
@ -20,7 +20,7 @@ open FieldOpAlgebra
|
|||
|
||||
open FieldStatistic
|
||||
|
||||
lemma stat_signFinset_right {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma stat_signFinset_right {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (i j : Fin [φsΛ]ᵘᶜ.length) :
|
||||
(𝓕 |>ₛ ⟨[φsΛ]ᵘᶜ.get, φsucΛ.signFinset i j⟩) =
|
||||
(𝓕 |>ₛ ⟨φs.get, (φsucΛ.signFinset i j).map uncontractedListEmd⟩) := by
|
||||
|
@ -32,7 +32,7 @@ lemma stat_signFinset_right {φs : List 𝓕.States} (φsΛ : WickContraction φ
|
|||
intro i j h
|
||||
exact uncontractedListEmd_strictMono h
|
||||
|
||||
lemma signFinset_right_map_uncontractedListEmd_eq_filter {φs : List 𝓕.States}
|
||||
lemma signFinset_right_map_uncontractedListEmd_eq_filter {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) (φsucΛ : WickContraction [φsΛ]ᵘᶜ.length)
|
||||
(i j : Fin [φsΛ]ᵘᶜ.length) : (φsucΛ.signFinset i j).map uncontractedListEmd =
|
||||
((join φsΛ φsucΛ).signFinset (uncontractedListEmd i) (uncontractedListEmd j)).filter
|
||||
|
@ -90,7 +90,7 @@ lemma signFinset_right_map_uncontractedListEmd_eq_filter {φs : List 𝓕.States
|
|||
exact hl
|
||||
exact fun _ _ h => uncontractedListEmd_strictMono h
|
||||
|
||||
lemma sign_right_eq_prod_mul_prod {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma sign_right_eq_prod_mul_prod {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
φsucΛ.sign = (∏ a, 𝓢(𝓕|>ₛ [φsΛ]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get,
|
||||
((join φsΛ φsucΛ).signFinset (uncontractedListEmd (φsucΛ.fstFieldOfContract a))
|
||||
|
@ -107,7 +107,7 @@ lemma sign_right_eq_prod_mul_prod {φs : List 𝓕.States} (φsΛ : WickContract
|
|||
rw [stat_signFinset_right, signFinset_right_map_uncontractedListEmd_eq_filter]
|
||||
rw [ofFinset_filter]
|
||||
|
||||
lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.States}
|
||||
lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).signFinset i j =
|
||||
|
@ -149,7 +149,7 @@ lemma join_singleton_signFinset_eq_filter {φs : List 𝓕.States}
|
|||
· simp only [Bool.not_eq_true, Option.not_isSome, Option.isNone_iff_eq_none] at h2'
|
||||
simp [h2']
|
||||
|
||||
lemma join_singleton_left_signFinset_eq_filter {φs : List 𝓕.States}
|
||||
lemma join_singleton_left_signFinset_eq_filter {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(𝓕 |>ₛ ⟨φs.get, (singleton h).signFinset i j⟩)
|
||||
|
@ -166,7 +166,7 @@ lemma join_singleton_left_signFinset_eq_filter {φs : List 𝓕.States}
|
|||
|
||||
/-- The difference in sign between `φsucΛ.sign` and the direct contribution of `φsucΛ` to
|
||||
`(join (singleton h) φsucΛ)`. -/
|
||||
def joinSignRightExtra {φs : List 𝓕.States}
|
||||
def joinSignRightExtra {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) : ℂ :=
|
||||
∏ a, 𝓢(𝓕|>ₛ [singleton h]ᵘᶜ[φsucΛ.sndFieldOfContract a], 𝓕|>ₛ ⟨φs.get,
|
||||
|
@ -176,7 +176,7 @@ def joinSignRightExtra {φs : List 𝓕.States}
|
|||
|
||||
/-- The difference in sign between `(singleton h).sign` and the direct contribution of
|
||||
`(singleton h)` to `(join (singleton h) φsucΛ)`. -/
|
||||
def joinSignLeftExtra {φs : List 𝓕.States}
|
||||
def joinSignLeftExtra {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) : ℂ :=
|
||||
𝓢(𝓕 |>ₛ φs[j], (𝓕 |>ₛ ⟨φs.get, ((singleton h).signFinset i j).filter (fun c =>
|
||||
|
@ -184,7 +184,7 @@ def joinSignLeftExtra {φs : List 𝓕.States}
|
|||
((h1 : ((join (singleton h) φsucΛ).getDual? c).isSome) →
|
||||
(((join (singleton h) φsucΛ).getDual? c).get h1) < i)))⟩))
|
||||
|
||||
lemma join_singleton_sign_left {φs : List 𝓕.States}
|
||||
lemma join_singleton_sign_left {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(singleton h).sign = 𝓢(𝓕 |>ₛ φs[j],
|
||||
|
@ -194,7 +194,7 @@ lemma join_singleton_sign_left {φs : List 𝓕.States}
|
|||
rw [map_mul]
|
||||
rfl
|
||||
|
||||
lemma join_singleton_sign_right {φs : List 𝓕.States}
|
||||
lemma join_singleton_sign_right {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
φsucΛ.sign =
|
||||
|
@ -206,7 +206,7 @@ lemma join_singleton_sign_right {φs : List 𝓕.States}
|
|||
rfl
|
||||
|
||||
|
||||
lemma joinSignRightExtra_eq_i_j_finset_eq_if {φs : List 𝓕.States}
|
||||
lemma joinSignRightExtra_eq_i_j_finset_eq_if {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j)
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
joinSignRightExtra h φsucΛ = ∏ a,
|
||||
|
@ -298,7 +298,7 @@ lemma joinSignRightExtra_eq_i_j_finset_eq_if {φs : List 𝓕.States}
|
|||
Option.get_some, forall_const, false_or, true_and]
|
||||
omega
|
||||
|
||||
lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
||||
lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j) (hs : (𝓕 |>ₛ φs[i]) = (𝓕 |>ₛ φs[j]))
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
joinSignLeftExtra h φsucΛ = joinSignRightExtra h φsucΛ := by
|
||||
|
@ -380,7 +380,7 @@ lemma joinSignLeftExtra_eq_joinSignRightExtra {φs : List 𝓕.States}
|
|||
simp only [Finset.disjoint_singleton_right, Finset.mem_singleton]
|
||||
exact Fin.ne_of_lt h
|
||||
|
||||
lemma join_sign_singleton {φs : List 𝓕.States}
|
||||
lemma join_sign_singleton {φs : List 𝓕.FieldOp}
|
||||
{i j : Fin φs.length} (h : i < j) (hs : (𝓕 |>ₛ φs[i]) = (𝓕 |>ₛ φs[j]))
|
||||
(φsucΛ : WickContraction [singleton h]ᵘᶜ.length) :
|
||||
(join (singleton h) φsucΛ).sign = (singleton h).sign * φsucΛ.sign := by
|
||||
|
@ -401,7 +401,7 @@ lemma join_sign_singleton {φs : List 𝓕.States}
|
|||
· funext a
|
||||
simp
|
||||
|
||||
lemma join_sign_induction {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_sign_induction {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (hc : φsΛ.GradingCompliant) :
|
||||
(n : ℕ) → (hn : φsΛ.1.card = n) →
|
||||
(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign
|
||||
|
@ -428,12 +428,12 @@ lemma join_sign_induction {φs : List 𝓕.States} (φsΛ : WickContraction φs.
|
|||
apply sign_congr
|
||||
exact join_uncontractedListGet (singleton hij) φsucΛ'
|
||||
|
||||
lemma join_sign {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_sign {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) (hc : φsΛ.GradingCompliant) :
|
||||
(join φsΛ φsucΛ).sign = φsΛ.sign * φsucΛ.sign := by
|
||||
exact join_sign_induction φsΛ φsucΛ hc (φsΛ).1.card rfl
|
||||
|
||||
lemma join_sign_timeContract {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_sign_timeContract {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
(join φsΛ φsucΛ).sign • (join φsΛ φsucΛ).timeContract.1 =
|
||||
(φsΛ.sign • φsΛ.timeContract.1) * (φsucΛ.sign • φsucΛ.timeContract.1) := by
|
||||
|
|
|
@ -46,7 +46,7 @@ lemma of_singleton_eq {i j : Fin n} (hij : i < j) (a : (singleton hij).1) :
|
|||
rw [@mem_singleton_iff] at ha2
|
||||
exact Subtype.coe_eq_of_eq_mk ha2
|
||||
|
||||
lemma singleton_prod {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j)
|
||||
lemma singleton_prod {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j)
|
||||
(f : (singleton hij).1 → M) [CommMonoid M] :
|
||||
∏ a, f a = f ⟨{i,j}, mem_singleton hij⟩:= by
|
||||
simp [singleton, of_singleton_eq]
|
||||
|
@ -67,7 +67,7 @@ lemma singleton_sndFieldOfContract {i j : Fin n} (hij : i < j) :
|
|||
· simp
|
||||
· exact hij
|
||||
|
||||
lemma singleton_sign_expand {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j) :
|
||||
lemma singleton_sign_expand {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j) :
|
||||
(singleton hij).sign = 𝓢(𝓕 |>ₛ φs[j], 𝓕 |>ₛ ⟨φs.get, (singleton hij).signFinset i j⟩) := by
|
||||
rw [sign, singleton_prod]
|
||||
simp
|
||||
|
@ -79,7 +79,7 @@ lemma singleton_getDual?_eq_none_iff_neq {i j : Fin n} (hij : i < j) (a : Fin n)
|
|||
simp only [singleton, Finset.mem_singleton, forall_eq, Finset.mem_insert, not_or, ne_eq]
|
||||
omega
|
||||
|
||||
lemma singleton_uncontractedEmd_neq_left {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j)
|
||||
lemma singleton_uncontractedEmd_neq_left {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j)
|
||||
(a : Fin [singleton hij]ᵘᶜ.length) :
|
||||
(singleton hij).uncontractedListEmd a ≠ i := by
|
||||
by_contra hn
|
||||
|
@ -90,7 +90,7 @@ lemma singleton_uncontractedEmd_neq_left {φs : List 𝓕.States} {i j : Fin φs
|
|||
simp [singleton]
|
||||
simp_all
|
||||
|
||||
lemma singleton_uncontractedEmd_neq_right {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j)
|
||||
lemma singleton_uncontractedEmd_neq_right {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j)
|
||||
(a : Fin [singleton hij]ᵘᶜ.length) :
|
||||
(singleton hij).uncontractedListEmd a ≠ j := by
|
||||
by_contra hn
|
||||
|
@ -111,7 +111,7 @@ lemma mem_signFinset {i j : Fin n} (hij : i < j) (a : Fin n) :
|
|||
apply Or.inl
|
||||
omega
|
||||
|
||||
lemma subContraction_singleton_eq_singleton {φs : List 𝓕.States}
|
||||
lemma subContraction_singleton_eq_singleton {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(a : φsΛ.1) : φsΛ.subContraction {a.1} (by simp) =
|
||||
singleton (φsΛ.fstFieldOfContract_lt_sndFieldOfContract a) := by
|
||||
|
@ -119,13 +119,13 @@ lemma subContraction_singleton_eq_singleton {φs : List 𝓕.States}
|
|||
simp only [subContraction, singleton, Finset.singleton_inj]
|
||||
exact finset_eq_fstFieldOfContract_sndFieldOfContract φsΛ a
|
||||
|
||||
lemma singleton_timeContract {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j) :
|
||||
lemma singleton_timeContract {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j) :
|
||||
(singleton hij).timeContract =
|
||||
⟨FieldOpAlgebra.timeContract φs[i] φs[j], timeContract_mem_center _ _⟩ := by
|
||||
rw [timeContract, singleton_prod]
|
||||
simp
|
||||
|
||||
lemma singleton_staticContract {φs : List 𝓕.States} {i j : Fin φs.length} (hij : i < j) :
|
||||
lemma singleton_staticContract {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j) :
|
||||
(singleton hij).staticContract.1 =
|
||||
[anPart φs[i], ofFieldOp φs[j]]ₛ := by
|
||||
rw [staticContract, singleton_prod]
|
||||
|
|
|
@ -21,7 +21,7 @@ open FieldOpAlgebra
|
|||
/-- Given a Wick contraction `φsΛ` 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 staticContract {φs : List 𝓕.States}
|
||||
noncomputable def staticContract {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
Subalgebra.center ℂ 𝓕.FieldOpAlgebra :=
|
||||
∏ (a : φsΛ.1), ⟨[anPart (φs.get (φsΛ.fstFieldOfContract a)),
|
||||
|
@ -29,7 +29,7 @@ noncomputable def staticContract {φs : List 𝓕.States}
|
|||
superCommute_anPart_ofFieldOp_mem_center _ _⟩
|
||||
|
||||
@[simp]
|
||||
lemma staticContract_insertAndContract_none (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma staticContract_insertAndContract_none (φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
(φsΛ ↩Λ φ i none).staticContract = φsΛ.staticContract := by
|
||||
rw [staticContract, insertAndContract_none_prod_contractions]
|
||||
|
@ -46,7 +46,7 @@ This follows from the fact that `(φsΛ ↩Λ φ i (some j))` has one more contr
|
|||
corresponding to `φ` contracted with `φⱼ`. The order depends on whether we insert `φ` before
|
||||
or after `φⱼ`. -/
|
||||
lemma staticContract_insertAndContract_some
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(φsΛ ↩Λ φ i (some j)).staticContract =
|
||||
(if i < i.succAbove j then
|
||||
|
@ -67,16 +67,16 @@ lemma staticContract_insertAndContract_some
|
|||
open FieldStatistic
|
||||
|
||||
lemma staticConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||
(hik : i < i.succAbove k) :
|
||||
(φsΛ ↩Λ φ i (some k)).staticContract =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x < k))⟩)
|
||||
• (contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedStatesEquiv φs φsΛ) (some k)) *
|
||||
• (contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedFieldOpEquiv φs φsΛ) (some k)) *
|
||||
φsΛ.staticContract) := by
|
||||
rw [staticContract_insertAndContract_some]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
||||
contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem,
|
||||
Algebra.smul_mul_assoc, uncontractedListGet]
|
||||
|
@ -97,7 +97,7 @@ lemma staticConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
|
|||
rw [h1]
|
||||
simp only [exchangeSign_mul_self]
|
||||
|
||||
lemma staticContract_of_not_gradingCompliant (φs : List 𝓕.States)
|
||||
lemma staticContract_of_not_gradingCompliant (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (h : ¬ GradingCompliant φs φsΛ) :
|
||||
φsΛ.staticContract = 0 := by
|
||||
rw [staticContract]
|
||||
|
|
|
@ -16,7 +16,7 @@ open FieldSpecification
|
|||
variable {𝓕 : FieldSpecification}
|
||||
|
||||
namespace WickContraction
|
||||
variable {n : ℕ} {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
variable {n : ℕ} {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
open HepLean.List
|
||||
open FieldOpAlgebra
|
||||
|
||||
|
|
|
@ -20,16 +20,16 @@ open FieldOpAlgebra
|
|||
|
||||
/-- The condition on a Wick contraction which is true iff and only if every contraction
|
||||
is between two fields of equal time. -/
|
||||
def EqTimeOnly {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) : Prop :=
|
||||
def EqTimeOnly {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : Prop :=
|
||||
∀ (i j), {i, j} ∈ φsΛ.1 → timeOrderRel φs[i] φs[j]
|
||||
noncomputable section
|
||||
|
||||
instance {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
instance {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
Decidable (EqTimeOnly φsΛ) :=
|
||||
inferInstanceAs (Decidable (∀ (i j), {i, j} ∈ φsΛ.1 → timeOrderRel φs[i] φs[j]))
|
||||
|
||||
namespace EqTimeOnly
|
||||
variable {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
variable {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
|
||||
lemma timeOrderRel_of_eqTimeOnly_pair {i j : Fin φs.length} (h : {i, j} ∈ φsΛ.1)
|
||||
(hc : EqTimeOnly φsΛ) :
|
||||
|
@ -48,7 +48,7 @@ lemma timeOrderRel_both_of_eqTimeOnly {i j : Fin φs.length} (h : {i, j} ∈ φs
|
|||
rw [@Finset.pair_comm]
|
||||
exact h
|
||||
|
||||
lemma eqTimeOnly_iff_forall_finset {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
lemma eqTimeOnly_iff_forall_finset {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
φsΛ.EqTimeOnly ↔ ∀ (a : φsΛ.1),
|
||||
timeOrderRel (φs[φsΛ.fstFieldOfContract a]) (φs[φsΛ.sndFieldOfContract a])
|
||||
∧ timeOrderRel (φs[φsΛ.sndFieldOfContract a]) (φs[φsΛ.fstFieldOfContract a]) := by
|
||||
|
@ -92,7 +92,7 @@ lemma eqTimeOnly_iff_forall_finset {φs : List 𝓕.States} (φsΛ : WickContrac
|
|||
simp_all
|
||||
|
||||
@[simp]
|
||||
lemma empty_mem {φs : List 𝓕.States} : empty (n := φs.length).EqTimeOnly := by
|
||||
lemma empty_mem {φs : List 𝓕.FieldOp} : empty (n := φs.length).EqTimeOnly := by
|
||||
rw [eqTimeOnly_iff_forall_finset]
|
||||
simp [empty]
|
||||
|
||||
|
@ -109,12 +109,12 @@ lemma staticContract_eq_timeContract_of_eqTimeOnly (h : φsΛ.EqTimeOnly) :
|
|||
exact a.2
|
||||
exact h
|
||||
|
||||
lemma eqTimeOnly_congr {φs φs' : List 𝓕.States} (h : φs = φs') (φsΛ : WickContraction φs.length) :
|
||||
lemma eqTimeOnly_congr {φs φs' : List 𝓕.FieldOp} (h : φs = φs') (φsΛ : WickContraction φs.length) :
|
||||
(congr (by simp [h]) φsΛ).EqTimeOnly (φs := φs') ↔ φsΛ.EqTimeOnly := by
|
||||
subst h
|
||||
simp
|
||||
|
||||
lemma quotContraction_eqTimeOnly {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma quotContraction_eqTimeOnly {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
(h : φsΛ.EqTimeOnly) (S : Finset (Finset (Fin φs.length))) (ha : S ⊆ φsΛ.1) :
|
||||
(φsΛ.quotContraction S ha).EqTimeOnly := by
|
||||
rw [eqTimeOnly_iff_forall_finset]
|
||||
|
@ -127,7 +127,7 @@ lemma quotContraction_eqTimeOnly {φs : List 𝓕.States} {φsΛ : WickContracti
|
|||
rw [eqTimeOnly_iff_forall_finset] at h
|
||||
apply h
|
||||
|
||||
lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(h : 0 < φsΛ.1.card) (h1 : φsΛ.EqTimeOnly) :
|
||||
∃ (i j : Fin φs.length) (h : i < j) (φsucΛ : WickContraction [singleton h]ᵘᶜ.length),
|
||||
φsΛ = join (singleton h) φsucΛ ∧ (timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i])
|
||||
|
@ -161,7 +161,7 @@ lemma exists_join_singleton_of_card_ge_zero {φs : List 𝓕.States} (φsΛ : Wi
|
|||
simp only [subContraction, Finset.card_singleton, id_eq, eq_mpr_eq_cast] at h1
|
||||
omega
|
||||
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction {φs : List 𝓕.States}
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(hl : φsΛ.EqTimeOnly) (a b: 𝓕.FieldOpAlgebra) : (n : ℕ) → (hn : φsΛ.1.card = n) →
|
||||
𝓣(a * φsΛ.timeContract.1 * b) = φsΛ.timeContract.1 * 𝓣(a * b)
|
||||
|
@ -184,13 +184,13 @@ lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction {φs : List 𝓕.St
|
|||
simp_all only [Nat.succ_eq_add_one, Fin.getElem_fin, add_left_inj]
|
||||
simp_all
|
||||
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid {φs : List 𝓕.States}
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_mid {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(hl : φsΛ.EqTimeOnly) (a b : 𝓕.FieldOpAlgebra) :
|
||||
𝓣(a * φsΛ.timeContract.1 * b) = φsΛ.timeContract.1 * 𝓣(a * b) := by
|
||||
exact timeOrder_timeContract_mul_of_eqTimeOnly_mid_induction φsΛ hl a b φsΛ.1.card rfl
|
||||
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_left {φs : List 𝓕.States}
|
||||
lemma timeOrder_timeContract_mul_of_eqTimeOnly_left {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(hl : φsΛ.EqTimeOnly) (b : 𝓕.FieldOpAlgebra) :
|
||||
𝓣(φsΛ.timeContract.1 * b) = φsΛ.timeContract.1 * 𝓣(b) := by
|
||||
|
@ -199,7 +199,7 @@ lemma timeOrder_timeContract_mul_of_eqTimeOnly_left {φs : List 𝓕.States}
|
|||
rw [timeOrder_timeContract_mul_of_eqTimeOnly_mid φsΛ hl]
|
||||
simp
|
||||
|
||||
lemma exists_join_singleton_of_not_eqTimeOnly {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma exists_join_singleton_of_not_eqTimeOnly {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(h1 : ¬ φsΛ.EqTimeOnly) :
|
||||
∃ (i j : Fin φs.length) (h : i < j) (φsucΛ : WickContraction [singleton h]ᵘᶜ.length),
|
||||
φsΛ = join (singleton h) φsucΛ ∧ (¬ timeOrderRel φs[i] φs[j] ∨ ¬ timeOrderRel φs[j] φs[i]) := by
|
||||
|
@ -224,7 +224,7 @@ lemma exists_join_singleton_of_not_eqTimeOnly {φs : List 𝓕.States} (φsΛ :
|
|||
· simp_all [h1]
|
||||
· simp_all [h1]
|
||||
|
||||
lemma timeOrder_timeContract_of_not_eqTimeOnly {φs : List 𝓕.States}
|
||||
lemma timeOrder_timeContract_of_not_eqTimeOnly {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length)
|
||||
(hl : ¬ φsΛ.EqTimeOnly) : 𝓣(φsΛ.timeContract.1) = 0 := by
|
||||
obtain ⟨i, j, hij, φsucΛ, rfl, hr⟩ := exists_join_singleton_of_not_eqTimeOnly φsΛ hl
|
||||
|
@ -238,7 +238,7 @@ lemma timeOrder_timeContract_of_not_eqTimeOnly {φs : List 𝓕.States}
|
|||
intro h
|
||||
simp_all
|
||||
|
||||
lemma timeOrder_staticContract_of_not_mem {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma timeOrder_staticContract_of_not_mem {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(hl : ¬ φsΛ.EqTimeOnly) : 𝓣(φsΛ.staticContract.1) = 0 := by
|
||||
obtain ⟨i, j, hij, φsucΛ, rfl, hr⟩ := exists_join_singleton_of_not_eqTimeOnly φsΛ hl
|
||||
rw [join_staticContract]
|
||||
|
@ -254,17 +254,17 @@ end EqTimeOnly
|
|||
|
||||
/-- The condition on a Wick contraction which is true if it has at least one contraction
|
||||
which is between two equal time fields. -/
|
||||
def HaveEqTime {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) : Prop :=
|
||||
def HaveEqTime {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) : Prop :=
|
||||
∃ (i j : Fin φs.length) (h : {i, j} ∈ φsΛ.1),
|
||||
timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i]
|
||||
|
||||
noncomputable instance {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
noncomputable instance {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
Decidable (HaveEqTime φsΛ) :=
|
||||
inferInstanceAs (Decidable (∃ (i j : Fin φs.length)
|
||||
(h : ({i, j} : Finset (Fin φs.length)) ∈ φsΛ.1),
|
||||
timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i]))
|
||||
|
||||
lemma haveEqTime_iff_finset {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
lemma haveEqTime_iff_finset {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
HaveEqTime φsΛ ↔ ∃ (a : Finset (Fin φs.length)) (h : a ∈ φsΛ.1),
|
||||
timeOrderRel φs[φsΛ.fstFieldOfContract ⟨a, h⟩] φs[φsΛ.sndFieldOfContract ⟨a, h⟩]
|
||||
∧ timeOrderRel φs[φsΛ.sndFieldOfContract ⟨a, h⟩] φs[φsΛ.fstFieldOfContract ⟨a, h⟩] := by
|
||||
|
@ -297,20 +297,20 @@ lemma haveEqTime_iff_finset {φs : List 𝓕.States} (φsΛ : WickContraction φ
|
|||
exact h1
|
||||
|
||||
@[simp]
|
||||
lemma empty_not_haveEqTime {φs : List 𝓕.States} :
|
||||
lemma empty_not_haveEqTime {φs : List 𝓕.FieldOp} :
|
||||
¬ HaveEqTime (empty : WickContraction φs.length) := by
|
||||
rw [haveEqTime_iff_finset]
|
||||
simp [empty]
|
||||
|
||||
/-- Given a Wick contraction the subset of contracted pairs between eqaul time fields. -/
|
||||
def eqTimeContractSet {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
def eqTimeContractSet {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
Finset (Finset (Fin φs.length)) :=
|
||||
Finset.univ.filter (fun a =>
|
||||
a ∈ φsΛ.1 ∧ ∀ (h : a ∈ φsΛ.1),
|
||||
timeOrderRel φs[φsΛ.fstFieldOfContract ⟨a, h⟩] φs[φsΛ.sndFieldOfContract ⟨a, h⟩]
|
||||
∧ timeOrderRel φs[φsΛ.sndFieldOfContract ⟨a, h⟩] φs[φsΛ.fstFieldOfContract ⟨a, h⟩])
|
||||
|
||||
lemma eqTimeContractSet_subset {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
lemma eqTimeContractSet_subset {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
eqTimeContractSet φsΛ ⊆ φsΛ.1 := by
|
||||
simp only [eqTimeContractSet, Fin.getElem_fin]
|
||||
intro a
|
||||
|
@ -318,12 +318,12 @@ lemma eqTimeContractSet_subset {φs : List 𝓕.States} (φsΛ : WickContraction
|
|||
intro h _
|
||||
exact h
|
||||
|
||||
lemma mem_of_mem_eqTimeContractSet{φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma mem_of_mem_eqTimeContractSet{φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{a : Finset (Fin φs.length)} (h : a ∈ eqTimeContractSet φsΛ) : a ∈ φsΛ.1 := by
|
||||
simp only [eqTimeContractSet, Fin.getElem_fin, Finset.mem_filter, Finset.mem_univ, true_and] at h
|
||||
exact h.1
|
||||
|
||||
lemma join_eqTimeContractSet {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_eqTimeContractSet {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
eqTimeContractSet (join φsΛ φsucΛ) = φsΛ.eqTimeContractSet ∪
|
||||
φsucΛ.eqTimeContractSet.map (Finset.mapEmbedding uncontractedListEmd).toEmbedding := by
|
||||
|
@ -378,7 +378,7 @@ lemma join_eqTimeContractSet {φs : List 𝓕.States} (φsΛ : WickContraction
|
|||
join_sndFieldOfContract_joinLiftRight]
|
||||
simpa using h2
|
||||
|
||||
lemma eqTimeContractSet_of_not_haveEqTime {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma eqTimeContractSet_of_not_haveEqTime {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
(h : ¬ HaveEqTime φsΛ) : eqTimeContractSet φsΛ = ∅ := by
|
||||
ext a
|
||||
simp only [Finset.not_mem_empty, iff_false]
|
||||
|
@ -389,7 +389,7 @@ lemma eqTimeContractSet_of_not_haveEqTime {φs : List 𝓕.States} {φsΛ : Wick
|
|||
have h2 := hn.2 hn.1
|
||||
simp_all
|
||||
|
||||
lemma eqTimeContractSet_of_mem_eqTimeOnly {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma eqTimeContractSet_of_mem_eqTimeOnly {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
(h : φsΛ.EqTimeOnly) : eqTimeContractSet φsΛ = φsΛ.1 := by
|
||||
ext a
|
||||
simp only [eqTimeContractSet, Fin.getElem_fin, Finset.mem_filter, Finset.mem_univ, true_and,
|
||||
|
@ -397,7 +397,7 @@ lemma eqTimeContractSet_of_mem_eqTimeOnly {φs : List 𝓕.States} {φsΛ : Wick
|
|||
rw [EqTimeOnly.eqTimeOnly_iff_forall_finset] at h
|
||||
exact fun h_1 => h ⟨a, h_1⟩
|
||||
|
||||
lemma subContraction_eqTimeContractSet_eqTimeOnly {φs : List 𝓕.States}
|
||||
lemma subContraction_eqTimeContractSet_eqTimeOnly {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
(φsΛ.subContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ)).EqTimeOnly := by
|
||||
rw [EqTimeOnly.eqTimeOnly_iff_forall_finset]
|
||||
|
@ -408,7 +408,7 @@ lemma subContraction_eqTimeContractSet_eqTimeOnly {φs : List 𝓕.States}
|
|||
simp only [subContraction_fstFieldOfContract, Fin.getElem_fin, subContraction_sndFieldOfContract]
|
||||
exact ha2.2 ha2.1
|
||||
|
||||
lemma pair_mem_eqTimeContractSet_iff {φs : List 𝓕.States} {i j : Fin φs.length}
|
||||
lemma pair_mem_eqTimeContractSet_iff {φs : List 𝓕.FieldOp} {i j : Fin φs.length}
|
||||
(φsΛ : WickContraction φs.length) (h : {i, j} ∈ φsΛ.1) :
|
||||
{i, j} ∈ φsΛ.eqTimeContractSet ↔ timeOrderRel φs[i] φs[j] ∧ timeOrderRel φs[j] φs[i] := by
|
||||
simp only [eqTimeContractSet, Fin.getElem_fin, Finset.mem_filter, Finset.mem_univ, true_and]
|
||||
|
@ -434,7 +434,7 @@ lemma pair_mem_eqTimeContractSet_iff {φs : List 𝓕.States} {i j : Fin φs.len
|
|||
simp_all only [and_self]
|
||||
|
||||
lemma subContraction_eqTimeContractSet_not_empty_of_haveEqTime
|
||||
{φs : List 𝓕.States} (φsΛ : WickContraction φs.length) (h : HaveEqTime φsΛ) :
|
||||
{φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) (h : HaveEqTime φsΛ) :
|
||||
φsΛ.subContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ) ≠ empty := by
|
||||
simp only [ne_eq]
|
||||
erw [Subtype.eq_iff]
|
||||
|
@ -448,7 +448,7 @@ lemma subContraction_eqTimeContractSet_not_empty_of_haveEqTime
|
|||
simp_all only [Fin.getElem_fin, and_self]
|
||||
exact h1
|
||||
|
||||
lemma quotContraction_eqTimeContractSet_not_haveEqTime {φs : List 𝓕.States}
|
||||
lemma quotContraction_eqTimeContractSet_not_haveEqTime {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
¬ HaveEqTime (φsΛ.quotContraction (eqTimeContractSet φsΛ) (eqTimeContractSet_subset φsΛ)) := by
|
||||
rw [haveEqTime_iff_finset]
|
||||
|
@ -468,7 +468,7 @@ lemma quotContraction_eqTimeContractSet_not_haveEqTime {φs : List 𝓕.States}
|
|||
obtain ⟨h, h1⟩ := hn''
|
||||
simp_all
|
||||
|
||||
lemma join_haveEqTime_of_eqTimeOnly_nonEmpty {φs : List 𝓕.States} (φsΛ : WickContraction φs.length)
|
||||
lemma join_haveEqTime_of_eqTimeOnly_nonEmpty {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length)
|
||||
(h1 : φsΛ.EqTimeOnly) (h2 : φsΛ ≠ empty)
|
||||
(φsucΛ : WickContraction [φsΛ]ᵘᶜ.length) :
|
||||
HaveEqTime (join φsΛ φsucΛ) := by
|
||||
|
@ -483,7 +483,7 @@ lemma join_haveEqTime_of_eqTimeOnly_nonEmpty {φs : List 𝓕.States} (φsΛ : W
|
|||
rw [Finset.pair_comm]
|
||||
exact h
|
||||
|
||||
lemma hasEqTimeEquiv_ext_sigma {φs : List 𝓕.States} {x1 x2 :
|
||||
lemma hasEqTimeEquiv_ext_sigma {φs : List 𝓕.FieldOp} {x1 x2 :
|
||||
Σ (φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly ∧ φsΛ ≠ empty}),
|
||||
{φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ HaveEqTime φssucΛ}}
|
||||
(h1 : x1.1.1 = x2.1.1) (h2 : x1.2.1 = congr (by simp [h1]) x2.2.1) : x1 = x2 := by
|
||||
|
@ -497,7 +497,7 @@ lemma hasEqTimeEquiv_ext_sigma {φs : List 𝓕.States} {x1 x2 :
|
|||
/-- The equivalence which seperates a Wick contraction which has an equal time contraction
|
||||
into a non-empty contraction only between equal-time fields and a Wick contraction which
|
||||
does not have equal time contractions. -/
|
||||
def hasEqTimeEquiv (φs : List 𝓕.States) :
|
||||
def hasEqTimeEquiv (φs : List 𝓕.FieldOp) :
|
||||
{φsΛ : WickContraction φs.length // HaveEqTime φsΛ} ≃
|
||||
Σ (φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly ∧ φsΛ ≠ empty}),
|
||||
{φssucΛ : WickContraction [φsΛ.1]ᵘᶜ.length // ¬ HaveEqTime φssucΛ} where
|
||||
|
@ -538,7 +538,7 @@ def hasEqTimeEquiv (φs : List 𝓕.States) :
|
|||
rw [uncontractedListEmd_congr hs]
|
||||
rw [Finset.map_map]
|
||||
|
||||
lemma sum_haveEqTime (φs : List 𝓕.States)
|
||||
lemma sum_haveEqTime (φs : List 𝓕.FieldOp)
|
||||
(f : WickContraction φs.length → M) [AddCommMonoid M]:
|
||||
∑ (φsΛ : {φsΛ : WickContraction φs.length // HaveEqTime φsΛ}), f φsΛ =
|
||||
∑ (φsΛ : {φsΛ : WickContraction φs.length // φsΛ.EqTimeOnly ∧ φsΛ ≠ empty}),
|
||||
|
|
|
@ -21,7 +21,7 @@ open FieldOpAlgebra
|
|||
/-- Given a Wick contraction `φsΛ` 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 {φs : List 𝓕.States}
|
||||
noncomputable def timeContract {φs : List 𝓕.FieldOp}
|
||||
(φsΛ : WickContraction φs.length) :
|
||||
Subalgebra.center ℂ 𝓕.FieldOpAlgebra :=
|
||||
∏ (a : φsΛ.1), ⟨FieldOpAlgebra.timeContract
|
||||
|
@ -35,7 +35,7 @@ This result follows from the fact that `timeContract` only depends on contracted
|
|||
and `(φsΛ ↩Λ φ i none)` has the 'same' contracted pairs as `φsΛ`. -/
|
||||
@[simp]
|
||||
lemma timeContract_insertAndContract_none
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
(φsΛ ↩Λ φ i none).timeContract = φsΛ.timeContract := by
|
||||
rw [timeContract, insertAndContract_none_prod_contractions]
|
||||
|
@ -52,7 +52,7 @@ This follows from the fact that `(φsΛ ↩Λ φ i (some j))` has one more contr
|
|||
corresponding to `φ` contracted with `φⱼ`. The order depends on whether we insert `φ` before
|
||||
or after `φⱼ`. -/
|
||||
lemma timeConract_insertAndContract_some
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(φsΛ ↩Λ φ i (some j)).timeContract =
|
||||
(if i < i.succAbove j then
|
||||
|
@ -71,7 +71,7 @@ lemma timeConract_insertAndContract_some
|
|||
simp
|
||||
|
||||
@[simp]
|
||||
lemma timeContract_empty (φs : List 𝓕.States) :
|
||||
lemma timeContract_empty (φs : List 𝓕.FieldOp) :
|
||||
(@empty φs.length).timeContract = 1 := by
|
||||
rw [timeContract, empty]
|
||||
simp
|
||||
|
@ -79,16 +79,16 @@ lemma timeContract_empty (φs : List 𝓕.States) :
|
|||
open FieldStatistic
|
||||
|
||||
lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||
(ht : 𝓕.timeOrderRel φ φs[k.1]) (hik : i < i.succAbove k) :
|
||||
(φsΛ ↩Λ φ i (some k)).timeContract =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x < k))⟩)
|
||||
• (contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedStatesEquiv φs φsΛ) (some k)) *
|
||||
• (contractStateAtIndex φ [φsΛ]ᵘᶜ ((uncontractedFieldOpEquiv φs φsΛ) (some k)) *
|
||||
φsΛ.timeContract) := by
|
||||
rw [timeConract_insertAndContract_some]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
||||
contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem,
|
||||
Algebra.smul_mul_assoc, uncontractedListGet]
|
||||
|
@ -112,16 +112,16 @@ lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_lt
|
|||
· exact ht
|
||||
|
||||
lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ : 𝓕.FieldOp) (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (k : φsΛ.uncontracted)
|
||||
(ht : ¬ 𝓕.timeOrderRel φs[k.1] φ) (hik : ¬ i < i.succAbove k) :
|
||||
(φsΛ ↩Λ φ i (some k)).timeContract =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ ⟨φs.get, (φsΛ.uncontracted.filter (fun x => x ≤ k))⟩)
|
||||
• (contractStateAtIndex φ [φsΛ]ᵘᶜ
|
||||
((uncontractedStatesEquiv φs φsΛ) (some k)) * φsΛ.timeContract) := by
|
||||
((uncontractedFieldOpEquiv φs φsΛ) (some k)) * φsΛ.timeContract) := by
|
||||
rw [timeConract_insertAndContract_some]
|
||||
simp only [Nat.succ_eq_add_one, Fin.getElem_fin, ite_mul, instCommGroup.eq_1,
|
||||
contractStateAtIndex, uncontractedStatesEquiv, Equiv.optionCongr_apply,
|
||||
contractStateAtIndex, uncontractedFieldOpEquiv, Equiv.optionCongr_apply,
|
||||
Equiv.coe_trans, Option.map_some', Function.comp_apply, finCongr_apply, Fin.coe_cast,
|
||||
List.getElem_map, uncontractedList_getElem_uncontractedIndexEquiv_symm, List.get_eq_getElem,
|
||||
Algebra.smul_mul_assoc, uncontractedListGet]
|
||||
|
@ -163,7 +163,7 @@ lemma timeConract_insertAndContract_some_eq_mul_contractStateAtIndex_not_lt
|
|||
simp_all only [Fin.getElem_fin, Nat.succ_eq_add_one, not_lt, false_or]
|
||||
exact ht
|
||||
|
||||
lemma timeContract_of_not_gradingCompliant (φs : List 𝓕.States)
|
||||
lemma timeContract_of_not_gradingCompliant (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (h : ¬ GradingCompliant φs φsΛ) :
|
||||
φsΛ.timeContract = 0 := by
|
||||
rw [timeContract]
|
||||
|
|
|
@ -316,41 +316,41 @@ lemma take_uncontractedIndexEquiv_symm (k : c.uncontracted) :
|
|||
|
||||
/-- Given a Wick Contraction `φsΛ` for a list of states `φs`. The list of uncontracted
|
||||
states in `φs`. -/
|
||||
def uncontractedListGet {φs : List 𝓕.States} (φsΛ : WickContraction φs.length) :
|
||||
List 𝓕.States := φsΛ.uncontractedList.map φs.get
|
||||
def uncontractedListGet {φs : List 𝓕.FieldOp} (φsΛ : WickContraction φs.length) :
|
||||
List 𝓕.FieldOp := φsΛ.uncontractedList.map φs.get
|
||||
|
||||
@[inherit_doc uncontractedListGet]
|
||||
scoped[WickContraction] notation "[" φsΛ "]ᵘᶜ" => uncontractedListGet φsΛ
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedListGet_empty {φs : List 𝓕.States} :
|
||||
lemma uncontractedListGet_empty {φs : List 𝓕.FieldOp} :
|
||||
(empty (n := φs.length)).uncontractedListGet = φs := by
|
||||
simp [uncontractedListGet]
|
||||
|
||||
/-!
|
||||
|
||||
## uncontractedStatesEquiv
|
||||
## uncontractedFieldOpEquiv
|
||||
|
||||
-/
|
||||
|
||||
/-- 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 `uncontractedIndexEquiv`. -/
|
||||
def uncontractedStatesEquiv (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) :
|
||||
def uncontractedFieldOpEquiv (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) :
|
||||
Option φsΛ.uncontracted ≃ Option (Fin [φsΛ]ᵘᶜ.length) :=
|
||||
Equiv.optionCongr (φsΛ.uncontractedIndexEquiv.symm.trans
|
||||
(finCongr (by simp [uncontractedListGet])))
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedStatesEquiv_none (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) :
|
||||
(uncontractedStatesEquiv φs φsΛ).toFun none = none := by
|
||||
simp [uncontractedStatesEquiv]
|
||||
lemma uncontractedFieldOpEquiv_none (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) :
|
||||
(uncontractedFieldOpEquiv φs φsΛ).toFun none = none := by
|
||||
simp [uncontractedFieldOpEquiv]
|
||||
|
||||
lemma uncontractedStatesEquiv_list_sum [AddCommMonoid α] (φs : List 𝓕.States)
|
||||
lemma uncontractedFieldOpEquiv_list_sum [AddCommMonoid α] (φs : List 𝓕.FieldOp)
|
||||
(φsΛ : WickContraction φs.length) (f : Option (Fin [φsΛ]ᵘᶜ.length) → α) :
|
||||
∑ (i : Option (Fin [φsΛ]ᵘᶜ.length)), f i =
|
||||
∑ (i : Option φsΛ.uncontracted), f (φsΛ.uncontractedStatesEquiv φs i) := by
|
||||
rw [(φsΛ.uncontractedStatesEquiv φs).sum_comp]
|
||||
∑ (i : Option φsΛ.uncontracted), f (φsΛ.uncontractedFieldOpEquiv φs i) := by
|
||||
rw [(φsΛ.uncontractedFieldOpEquiv φs).sum_comp]
|
||||
|
||||
/-!
|
||||
|
||||
|
@ -359,34 +359,34 @@ lemma uncontractedStatesEquiv_list_sum [AddCommMonoid α] (φs : List 𝓕.State
|
|||
-/
|
||||
|
||||
/-- The embedding of `Fin [φsΛ]ᵘᶜ.length` into `Fin φs.length`. -/
|
||||
def uncontractedListEmd {φs : List 𝓕.States} {φsΛ : WickContraction φs.length} :
|
||||
def uncontractedListEmd {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length} :
|
||||
Fin [φsΛ]ᵘᶜ.length ↪ Fin φs.length := ((finCongr (by simp [uncontractedListGet])).trans
|
||||
φsΛ.uncontractedIndexEquiv).toEmbedding.trans
|
||||
(Function.Embedding.subtype fun x => x ∈ φsΛ.uncontracted)
|
||||
|
||||
lemma uncontractedListEmd_congr {φs : List 𝓕.States} {φsΛ φsΛ' : WickContraction φs.length}
|
||||
lemma uncontractedListEmd_congr {φs : List 𝓕.FieldOp} {φsΛ φsΛ' : WickContraction φs.length}
|
||||
(h : φsΛ = φsΛ') : φsΛ.uncontractedListEmd =
|
||||
(finCongr (by simp [h])).toEmbedding.trans φsΛ'.uncontractedListEmd := by
|
||||
subst h
|
||||
rfl
|
||||
|
||||
lemma uncontractedListEmd_toFun_eq_get (φs : List 𝓕.States) (φsΛ : WickContraction φs.length) :
|
||||
lemma uncontractedListEmd_toFun_eq_get (φs : List 𝓕.FieldOp) (φsΛ : WickContraction φs.length) :
|
||||
(uncontractedListEmd (φsΛ := φsΛ)).toFun =
|
||||
φsΛ.uncontractedList.get ∘ (finCongr (by simp [uncontractedListGet])) := by
|
||||
rfl
|
||||
|
||||
lemma uncontractedListEmd_strictMono {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma uncontractedListEmd_strictMono {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
{i j : Fin [φsΛ]ᵘᶜ.length} (h : i < j) : uncontractedListEmd i < uncontractedListEmd j := by
|
||||
simp only [uncontractedListEmd, uncontractedIndexEquiv, List.get_eq_getElem,
|
||||
Equiv.trans_toEmbedding, Function.Embedding.trans_apply, Equiv.coe_toEmbedding, finCongr_apply,
|
||||
Equiv.coe_fn_mk, Fin.coe_cast, Function.Embedding.coe_subtype]
|
||||
exact List.Sorted.get_strictMono φsΛ.uncontractedList_sorted_lt h
|
||||
|
||||
lemma uncontractedListEmd_mem_uncontracted {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma uncontractedListEmd_mem_uncontracted {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
(i : Fin [φsΛ]ᵘᶜ.length) : uncontractedListEmd i ∈ φsΛ.uncontracted := by
|
||||
simp [uncontractedListEmd]
|
||||
|
||||
lemma uncontractedListEmd_surjective_mem_uncontracted {φs : List 𝓕.States}
|
||||
lemma uncontractedListEmd_surjective_mem_uncontracted {φs : List 𝓕.FieldOp}
|
||||
{φsΛ : WickContraction φs.length} (i : Fin φs.length) (hi : i ∈ φsΛ.uncontracted) :
|
||||
∃ j, φsΛ.uncontractedListEmd j = i := by
|
||||
simp only [uncontractedListEmd, Equiv.trans_toEmbedding, Function.Embedding.trans_apply,
|
||||
|
@ -401,7 +401,7 @@ lemma uncontractedListEmd_surjective_mem_uncontracted {φs : List 𝓕.States}
|
|||
rw [hj]
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedListEmd_finset_disjoint_left {φs : List 𝓕.States}
|
||||
lemma uncontractedListEmd_finset_disjoint_left {φs : List 𝓕.FieldOp}
|
||||
{φsΛ : WickContraction φs.length} (a : Finset (Fin [φsΛ]ᵘᶜ.length))
|
||||
(b : Finset (Fin φs.length)) (hb : b ∈ φsΛ.1) : Disjoint (a.map uncontractedListEmd) b := by
|
||||
rw [Finset.disjoint_left]
|
||||
|
@ -413,7 +413,7 @@ lemma uncontractedListEmd_finset_disjoint_left {φs : List 𝓕.States}
|
|||
rw [mem_uncontracted_iff_not_contracted] at h1
|
||||
exact h1 b hb
|
||||
|
||||
lemma uncontractedListEmd_finset_not_mem {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma uncontractedListEmd_finset_not_mem {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
(a : Finset (Fin [φsΛ]ᵘᶜ.length)) :
|
||||
a.map uncontractedListEmd ∉ φsΛ.1 := by
|
||||
by_contra hn
|
||||
|
@ -424,13 +424,13 @@ lemma uncontractedListEmd_finset_not_mem {φs : List 𝓕.States} {φsΛ : WickC
|
|||
simp at h2
|
||||
|
||||
@[simp]
|
||||
lemma getElem_uncontractedListEmd {φs : List 𝓕.States} {φsΛ : WickContraction φs.length}
|
||||
lemma getElem_uncontractedListEmd {φs : List 𝓕.FieldOp} {φsΛ : WickContraction φs.length}
|
||||
(k : Fin [φsΛ]ᵘᶜ.length) : φs[(uncontractedListEmd k).1] = [φsΛ]ᵘᶜ[k.1] := by
|
||||
simp only [uncontractedListGet, List.getElem_map, List.get_eq_getElem]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
lemma uncontractedListEmd_empty {φs : List 𝓕.States} :
|
||||
lemma uncontractedListEmd_empty {φs : List 𝓕.FieldOp} :
|
||||
(empty (n := φs.length)).uncontractedListEmd = (finCongr (by simp)).toEmbedding := by
|
||||
ext x
|
||||
simp [uncontractedListEmd, uncontractedIndexEquiv]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue