PhysLean/HepLean/PerturbationTheory/WickContraction/Singleton.lean

135 lines
4.7 KiB
Text
Raw Normal View History

2025-01-31 16:02:02 +00:00
/-
Copyright (c) 2025 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
import HepLean.PerturbationTheory.WickContraction.SubContraction
/-!
# Singleton of contractions
-/
open FieldSpecification
variable {𝓕 : FieldSpecification}
namespace WickContraction
variable {n : } (c : WickContraction n)
open HepLean.List
open FieldOpAlgebra
open FieldStatistic
2025-02-03 06:13:13 +00:00
/-- The Wick contraction formed from a single ordered pair. -/
2025-01-31 16:02:02 +00:00
def singleton {i j : Fin n} (hij : i < j) : WickContraction n :=
⟨{{i, j}}, by
intro i hi
2025-02-03 05:39:48 +00:00
simp only [Finset.mem_singleton] at hi
2025-01-31 16:02:02 +00:00
subst hi
rw [@Finset.card_eq_two]
use i, j
2025-02-03 05:39:48 +00:00
simp only [ne_eq, and_true]
2025-02-03 06:13:13 +00:00
omega, by
2025-01-31 16:02:02 +00:00
intro i hi j hj
simp_all⟩
lemma mem_singleton {i j : Fin n} (hij : i < j) :
2025-02-03 06:13:13 +00:00
{i, j} ∈ (singleton hij).1 := by
2025-01-31 16:02:02 +00:00
simp [singleton]
lemma mem_singleton_iff {i j : Fin n} (hij : i < j) {a : Finset (Fin n)} :
2025-02-03 06:13:13 +00:00
a ∈ (singleton hij).1 ↔ a = {i, j} := by
2025-01-31 16:02:02 +00:00
simp [singleton]
2025-02-03 06:13:13 +00:00
lemma of_singleton_eq {i j : Fin n} (hij : i < j) (a : (singleton hij).1) :
2025-01-31 16:02:02 +00:00
a = ⟨{i, j}, mem_singleton hij⟩ := by
have ha2 := a.2
rw [@mem_singleton_iff] at ha2
exact Subtype.coe_eq_of_eq_mk ha2
2025-02-03 11:28:14 +00:00
lemma singleton_prod {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j)
2025-02-03 06:13:13 +00:00
(f : (singleton hij).1 → M) [CommMonoid M] :
2025-01-31 16:02:02 +00:00
∏ a, f a = f ⟨{i,j}, mem_singleton hij⟩:= by
2025-02-03 06:13:13 +00:00
simp [singleton, of_singleton_eq]
2025-01-31 16:02:02 +00:00
@[simp]
lemma singleton_fstFieldOfContract {i j : Fin n} (hij : i < j) :
(singleton hij).fstFieldOfContract ⟨{i, j}, mem_singleton hij⟩ = i := by
refine eq_fstFieldOfContract_of_mem (singleton hij) ⟨{i, j}, mem_singleton hij⟩ i j ?_ ?_ ?_
· simp
· simp
· exact hij
@[simp]
lemma singleton_sndFieldOfContract {i j : Fin n} (hij : i < j) :
(singleton hij).sndFieldOfContract ⟨{i, j}, mem_singleton hij⟩ = j := by
refine eq_sndFieldOfContract_of_mem (singleton hij) ⟨{i, j}, mem_singleton hij⟩ i j ?_ ?_ ?_
· simp
· simp
· exact hij
2025-02-03 11:28:14 +00:00
lemma singleton_sign_expand {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j) :
2025-01-31 16:02:02 +00:00
(singleton hij).sign = 𝓢(𝓕 |>ₛ φs[j], 𝓕 |>ₛ ⟨φs.get, (singleton hij).signFinset i j⟩) := by
rw [sign, singleton_prod]
simp
lemma singleton_getDual?_eq_none_iff_neq {i j : Fin n} (hij : i < j) (a : Fin n) :
(singleton hij).getDual? a = none ↔ (i ≠ a ∧ j ≠ a) := by
rw [getDual?_eq_none_iff_mem_uncontracted]
rw [mem_uncontracted_iff_not_contracted]
2025-02-03 05:39:48 +00:00
simp only [singleton, Finset.mem_singleton, forall_eq, Finset.mem_insert, not_or, ne_eq]
2025-01-31 16:02:02 +00:00
omega
2025-02-03 11:28:14 +00:00
lemma singleton_uncontractedEmd_neq_left {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j)
2025-02-03 06:13:13 +00:00
(a : Fin [singleton hij]ᵘᶜ.length) :
(singleton hij).uncontractedListEmd a ≠ i := by
2025-01-31 16:02:02 +00:00
by_contra hn
have h1 : (singleton hij).uncontractedListEmd a ∈ (singleton hij).uncontracted := by
simp [uncontractedListEmd]
2025-02-03 06:13:13 +00:00
have h2 : i ∉ (singleton hij).uncontracted := by
2025-01-31 16:02:02 +00:00
rw [mem_uncontracted_iff_not_contracted]
simp [singleton]
simp_all
2025-02-03 11:28:14 +00:00
lemma singleton_uncontractedEmd_neq_right {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j)
2025-02-03 06:13:13 +00:00
(a : Fin [singleton hij]ᵘᶜ.length) :
(singleton hij).uncontractedListEmd a ≠ j := by
2025-01-31 16:02:02 +00:00
by_contra hn
have h1 : (singleton hij).uncontractedListEmd a ∈ (singleton hij).uncontracted := by
simp [uncontractedListEmd]
2025-02-03 06:13:13 +00:00
have h2 : j ∉ (singleton hij).uncontracted := by
2025-01-31 16:02:02 +00:00
rw [mem_uncontracted_iff_not_contracted]
simp [singleton]
simp_all
@[simp]
lemma mem_signFinset {i j : Fin n} (hij : i < j) (a : Fin n) :
a ∈ (singleton hij).signFinset i j ↔ i < a ∧ a < j := by
2025-02-03 05:39:48 +00:00
simp only [signFinset, Finset.mem_filter, Finset.mem_univ, true_and, and_congr_right_iff,
and_iff_left_iff_imp]
2025-01-31 16:02:02 +00:00
intro h1 h2
rw [@singleton_getDual?_eq_none_iff_neq]
apply Or.inl
omega
2025-02-03 11:28:14 +00:00
lemma subContraction_singleton_eq_singleton {φs : List 𝓕.FieldOp}
2025-01-31 16:02:02 +00:00
(φsΛ : WickContraction φs.length)
(a : φsΛ.1) : φsΛ.subContraction {a.1} (by simp) =
singleton (φsΛ.fstFieldOfContract_lt_sndFieldOfContract a) := by
apply Subtype.ext
2025-02-03 05:39:48 +00:00
simp only [subContraction, singleton, Finset.singleton_inj]
2025-01-31 16:02:02 +00:00
exact finset_eq_fstFieldOfContract_sndFieldOfContract φsΛ a
2025-02-03 11:28:14 +00:00
lemma singleton_timeContract {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j) :
2025-02-01 11:51:06 +00:00
(singleton hij).timeContract =
⟨FieldOpAlgebra.timeContract φs[i] φs[j], timeContract_mem_center _ _⟩ := by
rw [timeContract, singleton_prod]
simp
2025-02-03 11:28:14 +00:00
lemma singleton_staticContract {φs : List 𝓕.FieldOp} {i j : Fin φs.length} (hij : i < j) :
2025-02-01 11:51:06 +00:00
(singleton hij).staticContract.1 =
[anPart φs[i], ofFieldOp φs[j]]ₛ := by
rw [staticContract, singleton_prod]
simp
2025-01-31 16:02:02 +00:00
end WickContraction