refactor: Lint
This commit is contained in:
parent
2490535569
commit
c7bd59c981
10 changed files with 120 additions and 90 deletions
|
@ -85,6 +85,7 @@ lemma ofStateAlgebra_ofState (φ : 𝓕.States) :
|
|||
Roughly `[φ1, φ2]` gets sent to `(φ1ᶜ+ φ1ᵃ) * (φ2ᶜ+ φ2ᵃ)` etc. -/
|
||||
def ofStateList (φs : List 𝓕.States) : CrAnAlgebra 𝓕 := (List.map ofState φs).prod
|
||||
|
||||
/-- Coercion from `List 𝓕.States` to `CrAnAlgebra 𝓕` through `ofStateList`. -/
|
||||
instance : Coe (List 𝓕.States) (CrAnAlgebra 𝓕) := ⟨ofStateList⟩
|
||||
|
||||
@[simp]
|
||||
|
|
|
@ -36,6 +36,7 @@ def normalOrder : CrAnAlgebra 𝓕 →ₗ[ℂ] CrAnAlgebra 𝓕 :=
|
|||
Basis.constr ofCrAnListBasis ℂ fun φs =>
|
||||
normalOrderSign φs • ofCrAnList (normalOrderList φs)
|
||||
|
||||
@[inherit_doc normalOrder]
|
||||
scoped[FieldSpecification.CrAnAlgebra] notation "𝓝(" a ")" => normalOrder a
|
||||
|
||||
lemma normalOrder_ofCrAnList (φs : List 𝓕.CrAnStates) :
|
||||
|
|
|
@ -406,7 +406,7 @@ lemma superCommute_ofCrAnList_ofStateList_cons (φ : 𝓕.States) (φs : List
|
|||
|
||||
/--
|
||||
Within the creation and annihilation algebra, we have that
|
||||
`[φᶜᵃs, φᶜᵃ₀ … φᶜᵃₙ]ₛca = ∑ i, sᵢ • φᶜᵃs₀ … φᶜᵃᵢ₋₁ * [φᶜᵃs, φᶜᵃᵢ]ₛca * φᶜᵃᵢ₊₁ … φᶜᵃₙ`
|
||||
`[φᶜᵃs, φᶜᵃ₀ … φᶜᵃₙ]ₛca = ∑ i, sᵢ • φᶜᵃs₀ … φᶜᵃᵢ₋₁ * [φᶜᵃs, φᶜᵃᵢ]ₛca * φᶜᵃᵢ₊₁ … φᶜᵃₙ`
|
||||
where `sᵢ` is the exchange sign for `φᶜᵃs` and `φᶜᵃs₀ … φᶜᵃᵢ₋₁`.
|
||||
-/
|
||||
lemma superCommute_ofCrAnList_ofCrAnList_eq_sum (φs : List 𝓕.CrAnStates) :
|
||||
|
|
|
@ -20,7 +20,7 @@ is a generalization of the operator algebra of a field theory with field specifi
|
|||
It is an algebra `A` with a map `crAnF` from the creation and annihilation free algebra
|
||||
satisfying a number of conditions with respect to super commutators.
|
||||
The true operator algebra of a field theory with field specification `𝓕`is an
|
||||
example of a proto-operator algebra.-/
|
||||
example of a proto-operator algebra. -/
|
||||
structure ProtoOperatorAlgebra where
|
||||
/-- The algebra representing the operator algebra. -/
|
||||
A : Type
|
||||
|
|
|
@ -29,12 +29,13 @@ 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 𝓕.States} (φ : 𝓕.States) (φ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)
|
||||
|
||||
scoped[WickContraction] notation φs "↩Λ" φ:max i:max j => insertAndContract φ φs i j
|
||||
@[inherit_doc insertAndContract]
|
||||
scoped[WickContraction] notation φs "↩Λ" φ:max i:max j => insertAndContract φ φs i j
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContract_fstFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
|
@ -56,14 +57,16 @@ lemma insertAndContract_sndFieldOfContract (φ : 𝓕.States) (φs : List 𝓕.S
|
|||
lemma insertAndContract_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ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 simp [insertAndContractNat]⟩) =
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
simp [insertAndContractNat]⟩) =
|
||||
if i < i.succAbove j.1 then
|
||||
finCongr (insertIdx_length_fin φ φs i).symm i else
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) := by
|
||||
split
|
||||
· rename_i h
|
||||
refine (insertAndContract φ φsΛ i (some j)).eq_fstFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
simp [insertAndContractNat]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
|
@ -72,7 +75,8 @@ lemma insertAndContract_fstFieldOfContract_some_incl (φ : 𝓕.States) (φs : L
|
|||
simp_all
|
||||
· rename_i h
|
||||
refine (insertAndContract φ φsΛ i (some j)).eq_fstFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
simp [insertAndContractNat]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
|
||||
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
|
@ -141,7 +145,8 @@ lemma insertAndContract_some_succAbove_getDual?_eq_option (φ : 𝓕.States) (φ
|
|||
(i.succAbove j)) = Option.map (Fin.cast (insertIdx_length_fin φ φs i).symm ∘ i.succAbove)
|
||||
(φsΛ.getDual? j) := by
|
||||
simp only [Nat.succ_eq_add_one, insertAndContract, getDual?_congr, finCongr_apply, Fin.cast_trans,
|
||||
Fin.cast_eq_self, ne_eq, hkj, not_false_eq_true, insertAndContractNat_some_getDual?_of_neq, Option.map_map]
|
||||
Fin.cast_eq_self, ne_eq, hkj, not_false_eq_true, insertAndContractNat_some_getDual?_of_neq,
|
||||
Option.map_map]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
|
@ -167,14 +172,16 @@ lemma insertAndContract_none_getDual?_get_eq (φ : 𝓕.States) (φs : List 𝓕
|
|||
lemma insertAndContract_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φ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 simp [insertAndContractNat]⟩) =
|
||||
(congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
simp [insertAndContractNat]⟩) =
|
||||
if i < i.succAbove j.1 then
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j.1) else
|
||||
finCongr (insertIdx_length_fin φ φs i).symm i := by
|
||||
split
|
||||
· rename_i h
|
||||
refine (φsΛ ↩Λ φ i (some j)).eq_sndFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
simp [insertAndContractNat]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm i) (j :=
|
||||
finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j)) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
|
@ -183,7 +190,8 @@ lemma insertAndContract_sndFieldOfContract_some_incl (φ : 𝓕.States) (φs : L
|
|||
simp_all
|
||||
· rename_i h
|
||||
refine (φsΛ ↩Λ φ i (some j)).eq_sndFieldOfContract_of_mem
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by simp [insertAndContractNat]⟩)
|
||||
(a := congrLift (insertIdx_length_fin φ φs i).symm ⟨{i, i.succAbove j}, by
|
||||
simp [insertAndContractNat]⟩)
|
||||
(i := finCongr (insertIdx_length_fin φ φs i).symm (i.succAbove j))
|
||||
(j := finCongr (insertIdx_length_fin φ φs i).symm i) ?_ ?_ ?_
|
||||
· simp [congrLift]
|
||||
|
@ -240,8 +248,8 @@ lemma self_not_mem_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List
|
|||
|
||||
lemma succAbove_mem_insertAndContractLiftFinset (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(i : Fin φs.length.succ) (a : Finset (Fin φs.length)) (j : Fin φs.length) :
|
||||
Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j) ∈ insertAndContractLiftFinset φ i a ↔
|
||||
j ∈ a := by
|
||||
Fin.cast (insertIdx_length_fin φ φs i).symm (i.succAbove j)
|
||||
∈ insertAndContractLiftFinset φ i a ↔ j ∈ a := by
|
||||
simp only [insertAndContractLiftFinset, Finset.mem_map_equiv, finCongr_symm, finCongr_apply,
|
||||
Fin.cast_trans, Fin.cast_eq_self]
|
||||
simp only [Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
|
@ -268,7 +276,6 @@ lemma insert_fin_eq_self (φ : 𝓕.States) {φs : List 𝓕.States}
|
|||
use z
|
||||
rfl
|
||||
|
||||
|
||||
lemma insertLift_sum (φ : 𝓕.States) {φs : List 𝓕.States}
|
||||
(i : Fin φs.length.succ) [AddCommMonoid M] (f : WickContraction (φs.insertIdx i φ).length → M) :
|
||||
∑ c, f c =
|
||||
|
|
|
@ -123,8 +123,9 @@ def insertAndContractNat (c : WickContraction n) (i : Fin n.succ) (j : Option (c
|
|||
rw [Finset.mapEmbedding_apply, Finset.mapEmbedding_apply, Finset.disjoint_map]
|
||||
exact c.2.2 a' ha' b' hb'
|
||||
|
||||
lemma insertAndContractNat_of_isSome (c : WickContraction n) (i : Fin n.succ) (j : Option c.uncontracted)
|
||||
(hj : j.isSome) : (insertAndContractNat c i j).1 = Insert.insert {i, i.succAbove (j.get hj)}
|
||||
lemma insertAndContractNat_of_isSome (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Option c.uncontracted) (hj : j.isSome) :
|
||||
(insertAndContractNat c i j).1 = Insert.insert {i, i.succAbove (j.get hj)}
|
||||
(Finset.map (Finset.mapEmbedding i.succAboveEmb).toEmbedding c.1) := by
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset]
|
||||
rw [@Option.isSome_iff_exists] at hj
|
||||
|
@ -153,12 +154,14 @@ lemma self_mem_uncontracted_of_insertAndContractNat_none (c : WickContraction n)
|
|||
· exact Fin.ne_succAbove i y
|
||||
|
||||
@[simp]
|
||||
lemma self_not_mem_uncontracted_of_insertAndContractNat_some (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) : i ∉ (insertAndContractNat c i (some j)).uncontracted := by
|
||||
lemma self_not_mem_uncontracted_of_insertAndContractNat_some (c : WickContraction n)
|
||||
(i : Fin n.succ) (j : c.uncontracted) :
|
||||
i ∉ (insertAndContractNat c i (some j)).uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp [insertAndContractNat]
|
||||
|
||||
lemma insertAndContractNat_succAbove_mem_uncontracted_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
|
||||
lemma insertAndContractNat_succAbove_mem_uncontracted_iff (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Fin n) :
|
||||
(i.succAbove j) ∈ (insertAndContractNat c i none).uncontracted ↔ j ∈ c.uncontracted := by
|
||||
rw [mem_uncontracted_iff_not_contracted, mem_uncontracted_iff_not_contracted]
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset, Finset.mem_map,
|
||||
|
@ -191,8 +194,8 @@ lemma insertAndContractNat_succAbove_mem_uncontracted_iff (c : WickContraction n
|
|||
(fun a => hp'.2 (i.succAbove_right_injective a))
|
||||
|
||||
@[simp]
|
||||
lemma mem_uncontracted_insertAndContractNat_none_iff (c : WickContraction n) (i : Fin n.succ) (k : Fin n.succ) :
|
||||
k ∈ (insertAndContractNat c i none).uncontracted ↔
|
||||
lemma mem_uncontracted_insertAndContractNat_none_iff (c : WickContraction n) (i : Fin n.succ)
|
||||
(k : Fin n.succ) : k ∈ (insertAndContractNat c i none).uncontracted ↔
|
||||
k = i ∨ ∃ j, k = i.succAbove j ∧ j ∈ c.uncontracted := by
|
||||
by_cases hi : k = i
|
||||
· subst hi
|
||||
|
@ -201,7 +204,8 @@ lemma mem_uncontracted_insertAndContractNat_none_iff (c : WickContraction n) (i
|
|||
obtain ⟨z, hk⟩ := hi
|
||||
subst hk
|
||||
have hn : ¬ i.succAbove z = i := Fin.succAbove_ne i z
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat_succAbove_mem_uncontracted_iff, hn, false_or]
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat_succAbove_mem_uncontracted_iff, hn,
|
||||
false_or]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
exact ⟨z, rfl, h⟩
|
||||
|
@ -212,7 +216,8 @@ lemma mem_uncontracted_insertAndContractNat_none_iff (c : WickContraction n) (i
|
|||
exact hk.2
|
||||
|
||||
lemma insertAndContractNat_none_uncontracted (c : WickContraction n) (i : Fin n.succ) :
|
||||
(insertAndContractNat c i none).uncontracted = Insert.insert i (c.uncontracted.map i.succAboveEmb) := by
|
||||
(insertAndContractNat c i none).uncontracted =
|
||||
Insert.insert i (c.uncontracted.map i.succAboveEmb) := by
|
||||
ext a
|
||||
simp only [Nat.succ_eq_add_one, mem_uncontracted_insertAndContractNat_none_iff, Finset.mem_insert,
|
||||
Finset.mem_map, Fin.succAboveEmb_apply]
|
||||
|
@ -255,8 +260,8 @@ lemma mem_uncontracted_insertAndContractNat_some_iff (c : WickContraction n) (i
|
|||
∃ z, k = i.succAbove z ∧ z ∈ c.uncontracted ∧ z ≠ j := by
|
||||
by_cases hki : k = i
|
||||
· subst hki
|
||||
simp only [Nat.succ_eq_add_one, self_not_mem_uncontracted_of_insertAndContractNat_some, ne_eq, false_iff,
|
||||
not_exists, not_and, Decidable.not_not]
|
||||
simp only [Nat.succ_eq_add_one, self_not_mem_uncontracted_of_insertAndContractNat_some, ne_eq,
|
||||
false_iff, not_exists, not_and, Decidable.not_not]
|
||||
exact fun x hx => False.elim (Fin.ne_succAbove k x hx)
|
||||
· simp only [Nat.succ_eq_add_one, ← Fin.exists_succAbove_eq_iff] at hki
|
||||
obtain ⟨z, hk⟩ := hki
|
||||
|
@ -279,9 +284,9 @@ lemma mem_uncontracted_insertAndContractNat_some_iff (c : WickContraction n) (i
|
|||
rw [mem_uncontracted_iff_not_contracted]
|
||||
intro p hp
|
||||
rw [mem_uncontracted_iff_not_contracted] at h
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset, Finset.mem_insert,
|
||||
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
|
||||
not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] at h
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset,
|
||||
Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp,
|
||||
Finset.mem_singleton, not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] at h
|
||||
have hc := h.2 p hp
|
||||
rw [Finset.mapEmbedding_apply] at hc
|
||||
exact (Finset.mem_map' (i.succAboveEmb)).mpr.mt hc
|
||||
|
@ -290,20 +295,22 @@ lemma mem_uncontracted_insertAndContractNat_some_iff (c : WickContraction n) (i
|
|||
rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)] at hz'1
|
||||
subst hz'1
|
||||
rw [mem_uncontracted_iff_not_contracted]
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset, Finset.mem_insert,
|
||||
Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp, Finset.mem_singleton,
|
||||
not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂]
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset,
|
||||
Finset.mem_insert, Finset.mem_map, RelEmbedding.coe_toEmbedding, forall_eq_or_imp,
|
||||
Finset.mem_singleton, not_or, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂]
|
||||
apply And.intro
|
||||
· rw [Function.Injective.eq_iff (Fin.succAbove_right_injective)]
|
||||
exact And.intro (Fin.succAbove_ne i z) (fun a => hjz a.symm)
|
||||
· rw [mem_uncontracted_iff_not_contracted] at hz'
|
||||
exact fun a ha hc => hz'.1 a ha ((Finset.mem_map' (i.succAboveEmb)).mp hc)
|
||||
|
||||
lemma insertAndContractNat_some_uncontracted (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted) :
|
||||
(insertAndContractNat c i (some j)).uncontracted = (c.uncontracted.erase j).map i.succAboveEmb := by
|
||||
lemma insertAndContractNat_some_uncontracted (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) :
|
||||
(insertAndContractNat c i (some j)).uncontracted =
|
||||
(c.uncontracted.erase j).map i.succAboveEmb := by
|
||||
ext a
|
||||
simp only [Nat.succ_eq_add_one, mem_uncontracted_insertAndContractNat_some_iff, ne_eq, Finset.map_erase,
|
||||
Fin.succAboveEmb_apply, Finset.mem_erase, Finset.mem_map]
|
||||
simp only [Nat.succ_eq_add_one, mem_uncontracted_insertAndContractNat_some_iff, ne_eq,
|
||||
Finset.map_erase, Fin.succAboveEmb_apply, Finset.mem_erase, Finset.mem_map]
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
obtain ⟨z, h1, h2, h3⟩ := h
|
||||
|
@ -338,20 +345,22 @@ lemma insertAndContractNat_none_getDual?_isNone (c : WickContraction n) (i : Fin
|
|||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_succAbove_getDual?_eq_none_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
|
||||
lemma insertAndContractNat_succAbove_getDual?_eq_none_iff (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Fin n) :
|
||||
(insertAndContractNat c i none).getDual? (i.succAbove j) = none ↔ c.getDual? j = none := by
|
||||
have h1 := insertAndContractNat_succAbove_mem_uncontracted_iff c i j
|
||||
simpa [uncontracted] using h1
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_succAbove_getDual?_isSome_iff (c : WickContraction n) (i : Fin n.succ) (j : Fin n) :
|
||||
lemma insertAndContractNat_succAbove_getDual?_isSome_iff (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Fin n) :
|
||||
((insertAndContractNat c i none).getDual? (i.succAbove j)).isSome ↔ (c.getDual? j).isSome := by
|
||||
rw [← not_iff_not]
|
||||
simp
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_succAbove_getDual?_get (c : WickContraction n) (i : Fin n.succ) (j : Fin n)
|
||||
(h : ((insertAndContractNat c i none).getDual? (i.succAbove j)).isSome) :
|
||||
lemma insertAndContractNat_succAbove_getDual?_get (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Fin n) (h : ((insertAndContractNat c i none).getDual? (i.succAbove j)).isSome) :
|
||||
((insertAndContractNat c i none).getDual? (i.succAbove j)).get h =
|
||||
i.succAbove ((c.getDual? j).get (by simpa using h)) := by
|
||||
have h1 : (insertAndContractNat c i none).getDual? (i.succAbove j) = some
|
||||
|
@ -366,14 +375,15 @@ lemma insertAndContractNat_succAbove_getDual?_get (c : WickContraction n) (i : F
|
|||
exact Option.get_of_mem h h1
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_some_getDual?_eq (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted) :
|
||||
lemma insertAndContractNat_some_getDual?_eq (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) :
|
||||
(insertAndContractNat c i (some j)).getDual? i = some (i.succAbove j) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp [insertAndContractNat]
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_some_getDual?_neq_none (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
|
||||
(k : Fin n) (hkj : k ≠ j.1) :
|
||||
lemma insertAndContractNat_some_getDual?_neq_none (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) (k : Fin n) (hkj : k ≠ j.1) :
|
||||
(insertAndContractNat c i (some j)).getDual? (i.succAbove k) = none ↔ c.getDual? k = none := by
|
||||
apply Iff.intro
|
||||
· intro h
|
||||
|
@ -394,9 +404,10 @@ lemma insertAndContractNat_some_getDual?_neq_none (c : WickContraction n) (i : F
|
|||
simpa [uncontracted, -mem_uncontracted_insertAndContractNat_some_iff, ne_eq] using hk
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_some_getDual?_neq_isSome (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
|
||||
(k : Fin n) (hkj : k ≠ j.1) :
|
||||
((insertAndContractNat c i (some j)).getDual? (i.succAbove k)).isSome ↔ (c.getDual? k).isSome := by
|
||||
lemma insertAndContractNat_some_getDual?_neq_isSome (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) (k : Fin n) (hkj : k ≠ j.1) :
|
||||
((insertAndContractNat c i (some j)).getDual? (i.succAbove k)).isSome ↔
|
||||
(c.getDual? k).isSome := by
|
||||
rw [← not_iff_not]
|
||||
simp [hkj]
|
||||
|
||||
|
@ -409,8 +420,8 @@ lemma insertAndContractNat_some_getDual?_neq_isSome_get (c : WickContraction n)
|
|||
have h1 : ((insertAndContractNat c i (some j)).getDual? (i.succAbove k))
|
||||
= some (i.succAbove ((c.getDual? k).get (by simpa [hkj] using h))) := by
|
||||
rw [getDual?_eq_some_iff_mem]
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset, Finset.mem_insert, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding]
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset, Finset.mem_insert,
|
||||
Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Or.inr
|
||||
use { k, ((c.getDual? k).get (by simpa [hkj] using h))}
|
||||
simp only [self_getDual?_get_mem, true_and]
|
||||
|
@ -419,8 +430,8 @@ lemma insertAndContractNat_some_getDual?_neq_isSome_get (c : WickContraction n)
|
|||
exact Option.get_of_mem h h1
|
||||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_some_getDual?_of_neq (c : WickContraction n) (i : Fin n.succ) (j : c.uncontracted)
|
||||
(k : Fin n) (hkj : k ≠ j.1) :
|
||||
lemma insertAndContractNat_some_getDual?_of_neq (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : c.uncontracted) (k : Fin n) (hkj : k ≠ j.1) :
|
||||
(insertAndContractNat c i (some j)).getDual? (i.succAbove k) =
|
||||
Option.map i.succAbove (c.getDual? k) := by
|
||||
by_cases h : (c.getDual? k).isSome
|
||||
|
@ -444,8 +455,8 @@ lemma insertAndContractNat_some_getDual?_of_neq (c : WickContraction n) (i : Fin
|
|||
|
||||
-/
|
||||
@[simp]
|
||||
lemma insertAndContractNat_erase (c : WickContraction n) (i : Fin n.succ) (j : Option (c.uncontracted)) :
|
||||
erase (insertAndContractNat c i j) i = c := by
|
||||
lemma insertAndContractNat_erase (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Option c.uncontracted) : erase (insertAndContractNat c i j) i = c := by
|
||||
refine Subtype.eq ?_
|
||||
simp only [erase, Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset]
|
||||
conv_rhs => rw [c.eq_filter_mem_self]
|
||||
|
@ -489,12 +500,13 @@ lemma insertAndContractNat_erase (c : WickContraction n) (i : Fin n.succ) (j : O
|
|||
simp only [ha, true_and]
|
||||
rw [Finset.mapEmbedding_apply]
|
||||
|
||||
lemma insertAndContractNat_getDualErase (c : WickContraction n) (i : Fin n.succ) (j : Option c.uncontracted) :
|
||||
(insertAndContractNat c i j).getDualErase i =
|
||||
lemma insertAndContractNat_getDualErase (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Option c.uncontracted) : (insertAndContractNat c i j).getDualErase i =
|
||||
uncontractedCongr (c := c) (c' := (c.insertAndContractNat i j).erase i) (by simp) j := by
|
||||
match n with
|
||||
| 0 =>
|
||||
simp only [insertAndContractNat, Nat.succ_eq_add_one, Nat.reduceAdd, Finset.le_eq_subset, getDualErase]
|
||||
simp only [insertAndContractNat, Nat.succ_eq_add_one, Nat.reduceAdd, Finset.le_eq_subset,
|
||||
getDualErase]
|
||||
fin_cases j
|
||||
simp
|
||||
| Nat.succ n =>
|
||||
|
@ -503,8 +515,9 @@ lemma insertAndContractNat_getDualErase (c : WickContraction n) (i : Fin n.succ)
|
|||
have hi := insertAndContractNat_none_getDual?_isNone c i
|
||||
simp [getDualErase, hi]
|
||||
| some j =>
|
||||
simp only [Nat.succ_eq_add_one, getDualErase, insertAndContractNat_some_getDual?_eq, Option.isSome_some,
|
||||
↓reduceDIte, Option.get_some, predAboveI_succAbove, uncontractedCongr_some, Option.some.injEq]
|
||||
simp only [Nat.succ_eq_add_one, getDualErase, insertAndContractNat_some_getDual?_eq,
|
||||
Option.isSome_some, ↓reduceDIte, Option.get_some, predAboveI_succAbove,
|
||||
uncontractedCongr_some, Option.some.injEq]
|
||||
rfl
|
||||
|
||||
@[simp]
|
||||
|
@ -513,7 +526,8 @@ lemma erase_insert (c : WickContraction n.succ) (i : Fin n.succ) :
|
|||
match n with
|
||||
| 0 =>
|
||||
apply Subtype.eq
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, insertAndContractNat, getDualErase, Finset.le_eq_subset]
|
||||
simp only [Nat.succ_eq_add_one, Nat.reduceAdd, insertAndContractNat, getDualErase,
|
||||
Finset.le_eq_subset]
|
||||
ext a
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Iff.intro
|
||||
|
@ -565,8 +579,8 @@ lemma erase_insert (c : WickContraction n.succ) (i : Fin n.succ) :
|
|||
rfl
|
||||
simp only [Nat.succ_eq_add_one, ne_eq, self_neq_getDual?_get, not_false_eq_true]
|
||||
exact (getDualErase_isSome_iff_getDual?_isSome c i).mpr hi
|
||||
· simp only [Nat.succ_eq_add_one, insertAndContractNat, getDualErase, hi, Bool.false_eq_true, ↓reduceDIte,
|
||||
Finset.le_eq_subset]
|
||||
· simp only [Nat.succ_eq_add_one, insertAndContractNat, getDualErase, hi, Bool.false_eq_true,
|
||||
↓reduceDIte, Finset.le_eq_subset]
|
||||
ext a
|
||||
simp only [Finset.mem_map, RelEmbedding.coe_toEmbedding]
|
||||
apply Iff.intro
|
||||
|
@ -626,7 +640,8 @@ lemma insertLift_none_bijective {c : WickContraction n} (i : Fin n.succ) :
|
|||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_fstFieldOfContract (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Option (c.uncontracted)) (a : c.1) : (c.insertAndContractNat i j).fstFieldOfContract (insertLift i j a) =
|
||||
(j : Option (c.uncontracted)) (a : c.1) :
|
||||
(c.insertAndContractNat i j).fstFieldOfContract (insertLift i j a) =
|
||||
i.succAbove (c.fstFieldOfContract a) := by
|
||||
refine (c.insertAndContractNat i j).eq_fstFieldOfContract_of_mem (a := (insertLift i j a))
|
||||
(i := i.succAbove (c.fstFieldOfContract a)) (j := i.succAbove (c.sndFieldOfContract a)) ?_ ?_ ?_
|
||||
|
@ -641,7 +656,8 @@ lemma insertAndContractNat_fstFieldOfContract (c : WickContraction n) (i : Fin n
|
|||
|
||||
@[simp]
|
||||
lemma insertAndContractNat_sndFieldOfContract (c : WickContraction n) (i : Fin n.succ)
|
||||
(j : Option (c.uncontracted)) (a : c.1) : (c.insertAndContractNat i j).sndFieldOfContract (insertLift i j a) =
|
||||
(j : Option (c.uncontracted)) (a : c.1) :
|
||||
(c.insertAndContractNat i j).sndFieldOfContract (insertLift i j a) =
|
||||
i.succAbove (c.sndFieldOfContract a) := by
|
||||
refine (c.insertAndContractNat i j).eq_sndFieldOfContract_of_mem (a := (insertLift i j a))
|
||||
(i := i.succAbove (c.fstFieldOfContract a)) (j := i.succAbove (c.sndFieldOfContract a)) ?_ ?_ ?_
|
||||
|
@ -701,8 +717,8 @@ lemma insertLiftSome_surjective {c : WickContraction n} (i : Fin n.succ) (j : c.
|
|||
Function.Surjective (insertLiftSome i j) := by
|
||||
intro a
|
||||
have ha := a.2
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset, Finset.mem_insert, Finset.mem_map,
|
||||
RelEmbedding.coe_toEmbedding] at ha
|
||||
simp only [Nat.succ_eq_add_one, insertAndContractNat, Finset.le_eq_subset, Finset.mem_insert,
|
||||
Finset.mem_map, RelEmbedding.coe_toEmbedding] at ha
|
||||
rcases ha with ha | ha
|
||||
· use Sum.inl ()
|
||||
exact Subtype.eq ha.symm
|
||||
|
|
|
@ -41,12 +41,12 @@ lemma signFinset_insertAndContract_none (φ : 𝓕.States) (φs : List 𝓕.Stat
|
|||
rcases insert_fin_eq_self φ i k with hk | hk
|
||||
· subst hk
|
||||
conv_lhs => simp only [Nat.succ_eq_add_one, signFinset, finCongr_apply, Finset.mem_filter,
|
||||
Finset.mem_univ,
|
||||
insertAndContract_none_getDual?_self, Option.isSome_none, Bool.false_eq_true, IsEmpty.forall_iff,
|
||||
or_self, and_true, true_and]
|
||||
Finset.mem_univ, insertAndContract_none_getDual?_self, Option.isSome_none, Bool.false_eq_true,
|
||||
IsEmpty.forall_iff, or_self, and_true, true_and]
|
||||
by_cases h : i.succAbove i1 < i ∧ i < i.succAbove i2
|
||||
· simp [h, Fin.lt_def]
|
||||
· simp only [Nat.succ_eq_add_one, h, ↓reduceIte, self_not_mem_insertAndContractLiftFinset, iff_false]
|
||||
· simp only [Nat.succ_eq_add_one, h, ↓reduceIte, self_not_mem_insertAndContractLiftFinset,
|
||||
iff_false]
|
||||
rw [Fin.lt_def, Fin.lt_def] at h ⊢
|
||||
simp_all
|
||||
· obtain ⟨k, hk⟩ := hk
|
||||
|
@ -69,8 +69,8 @@ lemma signFinset_insertAndContract_none (φ : 𝓕.States) (φs : List 𝓕.Stat
|
|||
rw [h1]
|
||||
rw [succAbove_mem_insertAndContractLiftFinset]
|
||||
simp only [Nat.succ_eq_add_one, signFinset, finCongr_apply, Finset.mem_filter, Finset.mem_univ,
|
||||
insertAndContract_none_succAbove_getDual?_eq_none_iff, insertAndContract_none_succAbove_getDual?_isSome_iff,
|
||||
insertAndContract_none_getDual?_get_eq, true_and]
|
||||
insertAndContract_none_succAbove_getDual?_eq_none_iff, true_and,
|
||||
insertAndContract_none_succAbove_getDual?_isSome_iff, insertAndContract_none_getDual?_get_eq]
|
||||
rw [Fin.lt_def, Fin.lt_def, Fin.lt_def, Fin.lt_def]
|
||||
simp only [Fin.coe_cast, Fin.val_fin_lt]
|
||||
rw [Fin.succAbove_lt_succAbove_iff, Fin.succAbove_lt_succAbove_iff]
|
||||
|
@ -115,8 +115,8 @@ lemma stat_ofFinset_of_insertAndContractLiftFinset (φ : 𝓕.States) (φs : Lis
|
|||
Function.comp_apply, finCongr_apply]
|
||||
rcases insert_fin_eq_self φ i b with hk | hk
|
||||
· subst hk
|
||||
simp only [Nat.succ_eq_add_one, self_not_mem_insertAndContractLiftFinset, iff_false, not_exists,
|
||||
not_and]
|
||||
simp only [Nat.succ_eq_add_one, self_not_mem_insertAndContractLiftFinset, iff_false,
|
||||
not_exists, not_and]
|
||||
intro x hx
|
||||
refine Fin.ne_of_val_ne ?h.inl.h
|
||||
simp only [Fin.coe_cast, ne_eq]
|
||||
|
@ -347,9 +347,9 @@ lemma sign_insert_none (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : Wick
|
|||
rw [insertAndContract_none_prod_contractions]
|
||||
congr
|
||||
funext a
|
||||
simp only [instCommGroup.eq_1, Nat.succ_eq_add_one, insertAndContract_sndFieldOfContract, finCongr_apply,
|
||||
Fin.getElem_fin, Fin.coe_cast, insertIdx_getElem_fin, insertAndContract_fstFieldOfContract, ite_mul,
|
||||
one_mul]
|
||||
simp only [instCommGroup, Nat.succ_eq_add_one, insertAndContract_sndFieldOfContract,
|
||||
finCongr_apply, Fin.getElem_fin, Fin.coe_cast, insertIdx_getElem_fin,
|
||||
insertAndContract_fstFieldOfContract, ite_mul, one_mul]
|
||||
erw [signFinset_insertAndContract_none]
|
||||
split
|
||||
· rw [ofFinset_insert]
|
||||
|
@ -363,9 +363,11 @@ lemma sign_insert_none (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : Wick
|
|||
lemma signInsertNone_eq_mul_fst_snd (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
φsΛ.signInsertNone φ φs i = ∏ (a : φsΛ.1),
|
||||
(if i.succAbove (φsΛ.fstFieldOfContract a) < i then 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||
(if i.succAbove (φsΛ.fstFieldOfContract a) < i then
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||
else 1) *
|
||||
(if i.succAbove (φsΛ.sndFieldOfContract a) < i then 𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||
(if i.succAbove (φsΛ.sndFieldOfContract a) < i then
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ φs[φsΛ.sndFieldOfContract a])
|
||||
else 1) := by
|
||||
rw [signInsertNone]
|
||||
congr
|
||||
|
@ -389,8 +391,7 @@ lemma signInsertNone_eq_mul_fst_snd (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
rw [if_neg]
|
||||
simp only [mul_one]
|
||||
have hn := fstFieldOfContract_lt_sndFieldOfContract φsΛ a
|
||||
have hx : i.succAbove (φsΛ.fstFieldOfContract a) < i.succAbove (φsΛ.sndFieldOfContract a) := by
|
||||
exact Fin.succAbove_lt_succAbove_iff.mpr hn
|
||||
have hx := (Fin.succAbove_lt_succAbove_iff (p := i)).mpr hn
|
||||
omega
|
||||
|
||||
lemma signInsertNone_eq_prod_prod (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
|
@ -530,9 +531,9 @@ lemma sign_insert_some (φ : 𝓕.States) (φs : List 𝓕.States) (φsΛ : Wick
|
|||
rw [insertAndContract_some_prod_contractions]
|
||||
congr
|
||||
funext a
|
||||
simp only [instCommGroup.eq_1, Nat.succ_eq_add_one, insertAndContract_sndFieldOfContract, finCongr_apply,
|
||||
Fin.getElem_fin, Fin.coe_cast, insertIdx_getElem_fin, insertAndContract_fstFieldOfContract, not_lt,
|
||||
ite_mul, one_mul]
|
||||
simp only [instCommGroup, Nat.succ_eq_add_one, insertAndContract_sndFieldOfContract,
|
||||
finCongr_apply, Fin.getElem_fin, Fin.coe_cast, insertIdx_getElem_fin,
|
||||
insertAndContract_fstFieldOfContract, not_lt, ite_mul, one_mul]
|
||||
erw [signFinset_insertAndContract_some]
|
||||
split
|
||||
· rename_i h
|
||||
|
@ -625,8 +626,7 @@ lemma signInsertSomeProd_eq_prod_prod (φ : 𝓕.States) (φs : List 𝓕.States
|
|||
have ha := fstFieldOfContract_lt_sndFieldOfContract φsΛ a
|
||||
apply And.intro
|
||||
· intro hi
|
||||
have hx : i.succAbove (φsΛ.fstFieldOfContract a) < i.succAbove (φsΛ.sndFieldOfContract a) := by
|
||||
exact Fin.succAbove_lt_succAbove_iff.mpr ha
|
||||
have hx := (Fin.succAbove_lt_succAbove_iff (p := i)).mpr ha
|
||||
omega
|
||||
· omega
|
||||
simp [hφj]
|
||||
|
@ -669,8 +669,9 @@ lemma signInsertSomeProd_eq_list (φ : 𝓕.States) (φs : List 𝓕.States)
|
|||
(i : Fin φs.length.succ) (j : φsΛ.uncontracted) (hφj : (𝓕 |>ₛ φ) = (𝓕 |>ₛ φs[j.1]))
|
||||
(hg : GradingCompliant φs φsΛ) :
|
||||
φsΛ.signInsertSomeProd φ φs i j =
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (List.filter (fun x => (φsΛ.getDual? x).isSome ∧ ∀ (h : (φsΛ.getDual? x).isSome),
|
||||
x < j ∧ (i.succAbove x < i ∧ i < i.succAbove ((φsΛ.getDual? x).get h)
|
||||
𝓢(𝓕 |>ₛ φ, 𝓕 |>ₛ (List.filter (fun x => (φsΛ.getDual? x).isSome ∧
|
||||
∀ (h : (φsΛ.getDual? x).isSome), x < j ∧ (i.succAbove x < i ∧
|
||||
i < i.succAbove ((φsΛ.getDual? x).get h)
|
||||
∨ j < ((φsΛ.getDual? x).get h) ∧ ¬ i.succAbove x < i))
|
||||
(List.finRange φs.length)).map φs.get) := by
|
||||
rw [signInsertSomeProd_eq_prod_fin]
|
||||
|
|
|
@ -34,7 +34,8 @@ noncomputable def timeContract (𝓞 : 𝓕.ProtoOperatorAlgebra) {φs : List
|
|||
This result follows from the fact that `timeContract` only depends on contracted pairs,
|
||||
and `(φsΛ ↩Λ φ i none)` has the 'same' contracted pairs as `φsΛ`. -/
|
||||
@[simp]
|
||||
lemma timeContract_insertAndContract_none (𝓞 : 𝓕.ProtoOperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma timeContract_insertAndContract_none (𝓞 : 𝓕.ProtoOperatorAlgebra)
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) :
|
||||
(φsΛ ↩Λ φ i none).timeContract 𝓞 = φsΛ.timeContract 𝓞 := by
|
||||
rw [timeContract, insertAndContract_none_prod_contractions]
|
||||
|
@ -42,7 +43,8 @@ lemma timeContract_insertAndContract_none (𝓞 : 𝓕.ProtoOperatorAlgebra) (φ
|
|||
ext a
|
||||
simp
|
||||
|
||||
lemma timeConract_insertAndContract_some (𝓞 : 𝓕.ProtoOperatorAlgebra) (φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
lemma timeConract_insertAndContract_some (𝓞 : 𝓕.ProtoOperatorAlgebra)
|
||||
(φ : 𝓕.States) (φs : List 𝓕.States)
|
||||
(φsΛ : WickContraction φs.length) (i : Fin φs.length.succ) (j : φsΛ.uncontracted) :
|
||||
(φsΛ ↩Λ φ i (some j)).timeContract 𝓞 =
|
||||
(if i < i.succAbove j then
|
||||
|
|
|
@ -287,10 +287,13 @@ 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
|
||||
|
||||
scoped[WickContraction] notation "[" φsΛ "]ᵘᶜ" => uncontractedListGet φsΛ
|
||||
@[inherit_doc uncontractedListGet]
|
||||
scoped[WickContraction] notation "[" φsΛ "]ᵘᶜ" => uncontractedListGet φsΛ
|
||||
/-!
|
||||
|
||||
## uncontractedStatesEquiv
|
||||
|
|
|
@ -263,7 +263,7 @@ lemma wick_term_some_eq_wick_term_optionEraseZ (φ : 𝓕.States) (φs : List
|
|||
exact hg'
|
||||
|
||||
/--
|
||||
Given a Wick contraction `φsΛ` of `φs = φ₀φ₁…φₙ` and an `i`, we have that
|
||||
Given a Wick contraction `φsΛ` of `φs = φ₀φ₁…φₙ` and an `i`, we have that
|
||||
`(φsΛ.sign • φsΛ.timeContract 𝓞) * 𝓞.crAnF (φ * 𝓝([φsΛ]ᵘᶜ))`
|
||||
is equal to the product of
|
||||
- the exchange sign of `φ` and `φ₀φ₁…φᵢ₋₁`,
|
||||
|
@ -310,7 +310,6 @@ lemma wick_term_cons_eq_sum_wick_term (φ : 𝓕.States) (φs : List 𝓕.States
|
|||
rw [@Subalgebra.mem_center_iff] at ht
|
||||
rw [ht]
|
||||
|
||||
|
||||
/-!
|
||||
|
||||
## Wick's theorem
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue