PhysLean/HepLean/PerturbationTheory/Contractions/Basic.lean

293 lines
12 KiB
Text
Raw Normal View History

2024-12-15 12:42:50 +00:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Joseph Tooby-Smith
-/
2024-12-19 14:25:09 +00:00
import HepLean.PerturbationTheory.Wick.OperatorMap
import HepLean.Mathematics.Fin.Involutions
2024-12-15 12:42:50 +00:00
/-!
2024-12-20 15:21:13 +00:00
# Contractions of a list of fields
2024-12-15 12:42:50 +00:00
-/
namespace Wick
2024-12-17 07:15:47 +00:00
open HepLean.List
open HepLean.Fin
2024-12-20 13:33:39 +00:00
open FieldStatistic
2024-12-15 12:42:50 +00:00
2024-12-20 13:53:22 +00:00
variable {𝓕 : Type}
2025-01-03 15:13:16 +00:00
/-- Given a list of fields `φs`, the type of pairwise-contractions associated with `φs`
which have the list `φsᵤₙ` uncontracted. -/
inductive ContractionsAux : (φs : List 𝓕) → (φsᵤₙ : List 𝓕) → Type
2024-12-15 12:42:50 +00:00
| nil : ContractionsAux [] []
2025-01-03 15:13:16 +00:00
| cons {φs : List 𝓕} {φsᵤₙ : List 𝓕} {φ : 𝓕} (i : Option (Fin φsᵤₙ.length)) :
ContractionsAux φs φsᵤₙ → ContractionsAux (φ :: φs) (optionEraseZ φsᵤₙ φ i)
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- Given a list of fields `l`, the type of pairwise-contractions associated with `l`. -/
2025-01-03 15:13:16 +00:00
def Contractions (φs : List 𝓕) : Type := Σ aux, ContractionsAux φs aux
2024-12-15 12:42:50 +00:00
namespace Contractions
2024-12-20 13:53:22 +00:00
variable {l : List 𝓕} (c : Contractions l)
2024-12-15 12:42:50 +00:00
/-- The equivalence between `ContractionsAux` based on the propositionally equivalent
uncontracted lists. -/
def auxCongr : {φs : List 𝓕} → {φsᵤₙ φsᵤₙ' : List 𝓕} → (h : φsᵤₙ = φsᵤₙ') →
ContractionsAux φs φsᵤₙ ≃ ContractionsAux φs φsᵤₙ'
| _, _, _, Eq.refl _ => Equiv.refl _
lemma auxCongr_ext {φs : List 𝓕} {c c2 : Contractions φs} (h : c.1 = c2.1)
2025-01-05 17:00:36 +00:00
(hx : c.2 = auxCongr h.symm c2.2) : c = c2 := by
cases c
cases c2
2025-01-05 16:46:15 +00:00
simp only at h
subst h
2025-01-05 16:46:15 +00:00
simp only [auxCongr, Equiv.refl_apply] at hx
subst hx
rfl
2024-12-19 15:40:04 +00:00
/-- The list of uncontracted fields. -/
def uncontracted : List 𝓕 := c.1
2025-01-05 17:00:36 +00:00
lemma uncontracted_length_even_iff : {l : List 𝓕} → (c : Contractions l) →
Even l.length ↔ Even c.uncontracted.length
| [], ⟨[], ContractionsAux.nil⟩ => by
simp [uncontracted]
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ => by
simp only [List.length_cons, uncontracted, optionEraseZ]
rw [Nat.even_add_one, Nat.even_add_one]
rw [uncontracted_length_even_iff ⟨aux, c⟩]
rfl
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩=> by
simp only [List.length_cons, uncontracted, optionEraseZ_some_length]
rw [Nat.even_sub, Nat.even_add_one]
· simp only [Nat.not_even_iff_odd, Nat.not_even_one, iff_false]
rw [← Nat.not_even_iff_odd, ← Nat.not_even_iff_odd]
rw [uncontracted_length_even_iff ⟨aux, c⟩]
rfl
· refine Nat.one_le_iff_ne_zero.mpr (fun hn => ?_)
rw [hn] at n
exact Fin.elim0 n
2024-12-15 12:42:50 +00:00
2024-12-20 13:53:22 +00:00
lemma contractions_nil (a : Contractions ([] : List 𝓕)) : a = ⟨[], ContractionsAux.nil⟩ := by
2024-12-15 12:42:50 +00:00
cases a
rename_i aux c
cases c
rfl
/-- The embedding of the uncontracted fields into all fields. -/
2025-01-05 17:00:36 +00:00
def embedUncontracted {l : List 𝓕} (c : Contractions l) :
Fin c.uncontracted.length → Fin l.length :=
2025-01-03 15:13:16 +00:00
match l, c with
| [], ⟨[], ContractionsAux.nil⟩ => Fin.elim0
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ =>
Fin.cons ⟨0, Nat.zero_lt_succ φs.length⟩ (Fin.succ ∘ (embedUncontracted ⟨aux, c⟩))
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some n) c⟩ => by
let lc := embedUncontracted ⟨aux, c⟩
refine Fin.succ ∘ lc ∘ Fin.cast ?_ ∘ Fin.succAbove ⟨n, by
rw [uncontracted, optionEraseZ_some_length]
2025-01-03 15:13:16 +00:00
omega⟩
simp only [uncontracted, optionEraseZ_some_length]
2025-01-03 15:13:16 +00:00
have hq : aux.length ≠ 0 := by
by_contra hn
rw [hn] at n
exact Fin.elim0 n
omega
2025-01-05 17:00:36 +00:00
lemma embedUncontracted_injective {l : List 𝓕} (c : Contractions l) :
2025-01-03 15:13:16 +00:00
Function.Injective c.embedUncontracted := by
match l, c with
| [], ⟨[], ContractionsAux.nil⟩ =>
2025-01-05 16:46:15 +00:00
dsimp only [List.length_nil, embedUncontracted]
2025-01-03 15:13:16 +00:00
intro i
exact Fin.elim0 i
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) none c⟩ =>
2025-01-05 16:46:15 +00:00
dsimp only [List.length_cons, embedUncontracted, Fin.zero_eta]
2025-01-03 15:13:16 +00:00
refine Fin.cons_injective_iff.mpr ?_
apply And.intro
· simp only [Set.mem_range, Function.comp_apply, not_exists]
exact fun x => Fin.succ_ne_zero (embedUncontracted ⟨aux, c⟩ x)
· exact Function.Injective.comp (Fin.succ_injective φs.length)
(embedUncontracted_injective ⟨aux, c⟩)
2025-01-05 17:00:36 +00:00
| φ :: φs, ⟨_, .cons (φsᵤₙ := aux) (some i) c⟩ =>
dsimp only [List.length_cons, embedUncontracted]
refine Function.Injective.comp (Fin.succ_injective φs.length) ?hf
refine Function.Injective.comp (embedUncontracted_injective ⟨aux, c⟩) ?hf.hf
refine Function.Injective.comp (Fin.cast_injective (embedUncontracted.proof_5 φ φs aux i c))
Fin.succAbove_right_injective
2025-01-03 15:13:16 +00:00
2024-12-20 15:21:13 +00:00
/-- Establishes uniqueness of contractions for a single field, showing that any contraction
of a single field must be equivalent to the trivial contraction with no pairing. -/
2024-12-20 13:53:22 +00:00
lemma contractions_single {i : 𝓕} (a : Contractions [i]) : a =
2024-12-19 12:59:14 +00:00
⟨[i], ContractionsAux.cons none ContractionsAux.nil⟩ := by
2024-12-15 12:42:50 +00:00
cases a
rename_i aux c
cases c
2024-12-17 07:15:47 +00:00
rename_i aux' c'
cases c'
cases aux'
2024-12-19 12:59:14 +00:00
simp only [List.length_nil, optionEraseZ]
2024-12-17 07:15:47 +00:00
rename_i x
exact Fin.elim0 x
2024-12-20 15:21:13 +00:00
/--
This function provides an equivalence between the type of contractions for an empty list of fields
and the unit type, indicating that there is only one possible contraction for an empty list.
-/
2024-12-20 13:53:22 +00:00
def nilEquiv : Contractions ([] : List 𝓕) ≃ Unit where
2024-12-17 07:15:47 +00:00
toFun _ := ()
invFun _ := ⟨[], ContractionsAux.nil⟩
2024-12-19 12:59:14 +00:00
left_inv a := Eq.symm (contractions_nil a)
right_inv _ := rfl
2024-12-17 07:15:47 +00:00
2024-12-19 15:40:04 +00:00
/-- The equivalence between contractions of `a :: l` and contractions of
`Contractions l` paired with an optional element of `Fin (c.uncontracted).length` specifying
2024-12-19 15:40:04 +00:00
what `a` contracts with if any. -/
2025-01-03 15:13:16 +00:00
def consEquiv {φ : 𝓕} {φs : List 𝓕} :
Contractions (φ :: φs) ≃ (c : Contractions φs) × Option (Fin c.uncontracted.length) where
2024-12-15 12:42:50 +00:00
toFun c :=
match c with
| ⟨aux, c⟩ =>
match c with
2025-01-03 15:13:16 +00:00
| ContractionsAux.cons (φsᵤₙ := aux') i c => ⟨⟨aux', c⟩, i⟩
2024-12-15 12:42:50 +00:00
invFun ci :=
⟨(optionEraseZ (ci.fst.uncontracted) φ ci.2), ContractionsAux.cons (φ := φ) ci.2 ci.1.2⟩
2024-12-15 12:42:50 +00:00
left_inv c := by
match c with
| ⟨aux, c⟩ =>
match c with
2025-01-03 15:13:16 +00:00
| ContractionsAux.cons (φsᵤₙ := aux') i c => rfl
2024-12-15 12:42:50 +00:00
right_inv ci := by rfl
lemma consEquiv_ext {φs : List 𝓕} {c1 c2 : Contractions φs}
{n1 : Option (Fin c1.uncontracted.length)} {n2 : Option (Fin c2.uncontracted.length)}
(hc : c1 = c2) (hn : Option.map (finCongr (by rw [hc])) n1 = n2) :
2025-01-05 17:00:36 +00:00
(⟨c1, n1⟩ : (c : Contractions φs) × Option (Fin c.uncontracted.length)) = ⟨c2, n2⟩ := by
subst hc
subst hn
simp
2024-12-19 15:40:04 +00:00
/-- The type of contractions is decidable. -/
2025-01-03 15:13:16 +00:00
instance decidable : (φs : List 𝓕) → DecidableEq (Contractions φs)
2024-12-15 12:42:50 +00:00
| [] => fun a b =>
match a, b with
| ⟨_, a⟩, ⟨_, b⟩ =>
match a, b with
2024-12-19 12:59:14 +00:00
| ContractionsAux.nil, ContractionsAux.nil => isTrue rfl
2025-01-03 15:13:16 +00:00
| _ :: φs =>
haveI : DecidableEq (Contractions φs) := decidable φs
haveI : DecidableEq ((c : Contractions φs) × Option (Fin (c.uncontracted).length)) :=
2024-12-15 12:42:50 +00:00
Sigma.instDecidableEqSigma
2024-12-17 07:15:47 +00:00
Equiv.decidableEq consEquiv
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- The type of contractions is finite. -/
2025-01-03 15:13:16 +00:00
instance fintype : (φs : List 𝓕) → Fintype (Contractions φs)
2024-12-15 12:42:50 +00:00
| [] => {
elems := {⟨[], ContractionsAux.nil⟩}
complete := by
intro a
rw [Finset.mem_singleton]
exact contractions_nil a}
2025-01-03 15:13:16 +00:00
| φ :: φs =>
haveI : Fintype (Contractions φs) := fintype φs
haveI : Fintype ((c : Contractions φs) × Option (Fin (c.uncontracted).length)) :=
2024-12-15 12:42:50 +00:00
Sigma.instFintype
2024-12-17 07:15:47 +00:00
Fintype.ofEquiv _ consEquiv.symm
2024-12-15 12:42:50 +00:00
2024-12-20 15:21:13 +00:00
/-- A contraction is a full contraction if there normalizing list of fields is empty. -/
def IsFull : Prop := c.uncontracted = []
2024-12-20 15:21:13 +00:00
/-- Provides a decidable instance for determining if a contraction is full
(i.e., all fields are paired). -/
instance isFull_decidable : Decidable c.IsFull := by
have hn : c.IsFull ↔ c.uncontracted.length = 0 := by
2024-12-20 15:21:13 +00:00
simp [IsFull]
apply decidable_of_decidable_of_iff hn.symm
2024-12-19 15:40:04 +00:00
/-- A structure specifying when a type `I` and a map `f :I → Type` corresponds to
2025-01-13 23:59:30 +01:00
the splitting of a fields `φ` into a creation `n` and annihilation part `p`. -/
2024-12-20 13:53:22 +00:00
structure Splitting (f : 𝓕 → Type) [∀ i, Fintype (f i)]
2025-01-03 15:13:16 +00:00
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le] where
2024-12-19 15:40:04 +00:00
/-- The creation part of the fields. The label `n` corresponds to the fact that
2025-01-14 00:11:36 +01:00
in normal ordering these fields get pushed to the negative (left) direction. -/
2024-12-20 13:53:22 +00:00
𝓑n : 𝓕 → (Σ i, f i)
2025-01-14 00:11:36 +01:00
/-- The annihilation part of the fields. The label `p` corresponds to the fact that
in normal ordering these fields get pushed to the positive (right) direction. -/
2024-12-20 13:53:22 +00:00
𝓑p : 𝓕 → (Σ i, f i)
2025-01-13 23:59:30 +01:00
/-- The complex coefficient of creation part of a field `i`. This is usually `0` or `1`. -/
2024-12-20 13:53:22 +00:00
𝓧n : 𝓕
2025-01-14 00:11:36 +01:00
/-- The complex coefficient of annihilation part of a field `i`. This is usually `0` or `1`. -/
2024-12-20 13:53:22 +00:00
𝓧p : 𝓕
2024-12-19 14:25:09 +00:00
h𝓑 : ∀ i, ofListLift f [i] 1 = ofList [𝓑n i] (𝓧n i) + ofList [𝓑p i] (𝓧p i)
2025-01-03 15:13:16 +00:00
h𝓑n : ∀ i j, le (𝓑n i) j
h𝓑p : ∀ i j, le j (𝓑p i)
2024-12-15 12:42:50 +00:00
2024-12-19 15:40:04 +00:00
/-- In the static wick's theorem, the term associated with a contraction `c` containing
the contractions. -/
2024-12-20 13:53:22 +00:00
noncomputable def toCenterTerm (f : 𝓕 → Type) [∀ i, Fintype (f i)]
(q : 𝓕 → FieldStatistic)
2025-01-03 15:13:16 +00:00
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
2024-12-15 12:42:50 +00:00
{A : Type} [Semiring A] [Algebra A]
2024-12-19 15:40:04 +00:00
(F : FreeAlgebra (Σ i, f i) →ₐ[] A) :
2025-01-03 15:13:16 +00:00
{φs : List 𝓕} → (c : Contractions φs) → (S : Splitting f le) → A
2024-12-19 11:23:49 +00:00
| [], ⟨[], .nil⟩, _ => 1
2025-01-03 15:13:16 +00:00
| _ :: _, ⟨_, .cons (φsᵤₙ := aux') none c⟩, S => toCenterTerm f q le F ⟨aux', c⟩ S
| a :: _, ⟨_, .cons (φsᵤₙ := aux') (some n) c⟩, S => toCenterTerm f q le F ⟨aux', c⟩ S *
2024-12-17 07:15:47 +00:00
superCommuteCoef q [aux'.get n] (List.take (↑n) aux') •
2025-01-03 15:13:16 +00:00
F (((superCommute fun i => q i.fst) (ofList [S.𝓑p a] (S.𝓧p a)))
(ofListLift f [aux'.get n] 1))
2024-12-17 07:15:47 +00:00
2024-12-20 15:21:13 +00:00
/-- Shows that adding a field with no contractions (none) to an existing set of contractions
results in the same center term as the original contractions.
Physically, this represents that an uncontracted (free) field does not affect
the contraction structure of other fields in Wick's theorem. -/
2024-12-20 13:53:22 +00:00
lemma toCenterTerm_none (f : 𝓕 → Type) [∀ i, Fintype (f i)]
2025-01-03 15:13:16 +00:00
(q : 𝓕 → FieldStatistic) {φs : List 𝓕}
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
2024-12-17 07:15:47 +00:00
{A : Type} [Semiring A] [Algebra A]
2024-12-19 15:40:04 +00:00
(F : FreeAlgebra (Σ i, f i) →ₐ A)
2025-01-03 15:13:16 +00:00
(S : Splitting f le) (φ : 𝓕) (c : Contractions φs) :
toCenterTerm (φs := φ :: φs) f q le F (Contractions.consEquiv.symm ⟨c, none⟩) S =
toCenterTerm f q le F c S := by
2024-12-17 07:15:47 +00:00
rw [consEquiv]
2024-12-19 12:59:14 +00:00
simp only [Equiv.coe_fn_symm_mk]
2025-01-05 16:46:15 +00:00
dsimp only [toCenterTerm]
2024-12-17 07:15:47 +00:00
rfl
2024-12-15 12:42:50 +00:00
2024-12-20 15:21:13 +00:00
/-- Proves that the part of the term gained from Wick contractions is in
the center of the algebra. -/
2024-12-20 13:53:22 +00:00
lemma toCenterTerm_center (f : 𝓕 → Type) [∀ i, Fintype (f i)]
(q : 𝓕 → FieldStatistic)
2024-12-20 14:07:20 +00:00
(le : (Σ i, f i) → (Σ i, f i) → Prop) [DecidableRel le]
2024-12-17 07:15:47 +00:00
{A : Type} [Semiring A] [Algebra A]
2024-12-20 14:07:20 +00:00
(F : FreeAlgebra (Σ i, f i) →ₐ A) [OperatorMap (fun i => q i.1) le F] :
2025-01-03 15:13:16 +00:00
{φs : List 𝓕} → (c : Contractions φs) → (S : Splitting f le) →
2024-12-20 14:07:20 +00:00
(c.toCenterTerm f q le F S) ∈ Subalgebra.center A
2024-12-19 11:23:49 +00:00
| [], ⟨[], .nil⟩, _ => by
2025-01-05 16:46:15 +00:00
dsimp only [toCenterTerm]
2024-12-19 11:23:49 +00:00
exact Subalgebra.one_mem (Subalgebra.center A)
2025-01-03 15:13:16 +00:00
| _ :: _, ⟨_, .cons (φsᵤₙ := aux') none c⟩, S => by
2025-01-05 16:46:15 +00:00
dsimp only [toCenterTerm]
2024-12-20 14:07:20 +00:00
exact toCenterTerm_center f q le F ⟨aux', c⟩ S
2025-01-03 15:13:16 +00:00
| a :: _, ⟨_, .cons (φsᵤₙ := aux') (some n) c⟩, S => by
2025-01-05 16:46:15 +00:00
dsimp only [toCenterTerm, List.get_eq_getElem]
2024-12-19 11:23:49 +00:00
refine Subalgebra.mul_mem (Subalgebra.center A) ?hx ?hy
2024-12-20 14:07:20 +00:00
exact toCenterTerm_center f q le F ⟨aux', c⟩ S
2024-12-19 11:23:49 +00:00
apply Subalgebra.smul_mem
2024-12-19 14:25:09 +00:00
rw [ofListLift_expand]
2024-12-19 11:23:49 +00:00
rw [map_sum, map_sum]
refine Subalgebra.sum_mem (Subalgebra.center A) ?hy.hx.h
intro x _
2024-12-20 13:57:29 +00:00
simp only [CreateAnnihilateSect.toList]
2024-12-19 11:23:49 +00:00
rw [ofList_singleton]
2024-12-19 12:59:14 +00:00
exact OperatorMap.superCommute_ofList_singleton_ι_center (q := fun i => q i.1)
2024-12-20 14:07:20 +00:00
(le := le) F (S.𝓑p a) ⟨aux'[↑n], x.head⟩
2024-12-15 12:42:50 +00:00
end Contractions
end Wick