PhysLean/HepLean/PerturbationTheory/Wick/OfList.lean

206 lines
7.5 KiB
Text
Raw Normal View History

2024-12-19 14:25:09 +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-20 13:57:29 +00:00
import HepLean.PerturbationTheory.Wick.CreateAnnihilateSection
2024-12-19 14:25:09 +00:00
import HepLean.PerturbationTheory.Wick.KoszulOrder
/-!
# Koszul signs and ordering for lists and algebras
-/
namespace Wick
open HepLean.List
2024-12-20 13:11:22 +00:00
open FieldStatistic
2024-12-19 14:25:09 +00:00
noncomputable section
2024-12-20 13:11:22 +00:00
variable {𝓕 : Type} (q : 𝓕 → FieldStatistic) (le : 𝓕𝓕 → Prop) [DecidableRel le]
2024-12-19 15:40:04 +00:00
/-- The element of the free algebra `FreeAlgebra I` associated with a `List I`. -/
2024-12-20 13:11:22 +00:00
def ofList (l : List 𝓕) (x : ) : FreeAlgebra 𝓕 :=
2024-12-19 14:25:09 +00:00
FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x)
2024-12-20 13:11:22 +00:00
lemma ofList_pair (l r : List 𝓕) (x y : ) :
2024-12-19 14:25:09 +00:00
ofList (l ++ r) (x * y) = ofList l x * ofList r y := by
simp only [ofList, ← map_mul, MonoidAlgebra.single_mul_single, EmbeddingLike.apply_eq_iff_eq]
rfl
2024-12-20 13:11:22 +00:00
lemma ofList_triple (la lb lc : List 𝓕) (xa xb xc : ) :
2024-12-19 14:25:09 +00:00
ofList (la ++ lb ++ lc) (xa * xb * xc) = ofList la xa * ofList lb xb * ofList lc xc := by
rw [ofList_pair, ofList_pair]
2024-12-20 13:11:22 +00:00
lemma ofList_triple_assoc (la lb lc : List 𝓕) (xa xb xc : ) :
2024-12-19 14:25:09 +00:00
ofList (la ++ (lb ++ lc)) (xa * (xb * xc)) = ofList la xa * ofList lb xb * ofList lc xc := by
rw [ofList_pair, ofList_pair]
exact Eq.symm (mul_assoc (ofList la xa) (ofList lb xb) (ofList lc xc))
2024-12-20 13:11:22 +00:00
lemma ofList_cons_eq_ofList (l : List 𝓕) (i : 𝓕) (x : ) :
2024-12-19 14:25:09 +00:00
ofList (i :: l) x = ofList [i] 1 * ofList l x := by
simp only [ofList, ← map_mul, MonoidAlgebra.single_mul_single, one_mul,
EmbeddingLike.apply_eq_iff_eq]
rfl
2024-12-20 13:11:22 +00:00
lemma ofList_singleton (i : 𝓕) :
2024-12-19 14:25:09 +00:00
ofList [i] 1 = FreeAlgebra.ι i := by
simp only [ofList, FreeAlgebra.equivMonoidAlgebraFreeMonoid, MonoidAlgebra.of_apply,
MonoidAlgebra.single, AlgEquiv.ofAlgHom_symm_apply, MonoidAlgebra.lift_single, one_smul]
rfl
2024-12-20 13:11:22 +00:00
lemma ofList_eq_smul_one (l : List 𝓕) (x : ) :
2024-12-19 14:25:09 +00:00
ofList l x = x • ofList l 1 := by
simp only [ofList]
rw [← map_smul]
simp
2024-12-20 13:11:22 +00:00
lemma ofList_empty : ofList [] 1 = (1 : FreeAlgebra 𝓕) := by
2024-12-19 14:25:09 +00:00
simp only [ofList, EmbeddingLike.map_eq_one_iff]
rfl
2024-12-20 13:11:22 +00:00
lemma ofList_empty' : ofList [] x = x • (1 : FreeAlgebra 𝓕) := by
2024-12-19 14:25:09 +00:00
rw [ofList_eq_smul_one, ofList_empty]
2024-12-20 13:11:22 +00:00
lemma koszulOrder_ofList
(l : List 𝓕) (x : ) :
koszulOrder q le (ofList l x) = (koszulSign q le l) • ofList (List.insertionSort le l) x := by
2024-12-19 14:25:09 +00:00
rw [ofList]
rw [koszulOrder_single]
2024-12-20 13:11:22 +00:00
change ofList (List.insertionSort le l) _ = _
2024-12-19 14:25:09 +00:00
rw [ofList_eq_smul_one]
conv_rhs => rw [ofList_eq_smul_one]
rw [smul_smul]
2024-12-20 13:11:22 +00:00
lemma ofList_insertionSort_eq_koszulOrder (l : List 𝓕) (x : ) :
ofList (List.insertionSort le l) x = (koszulSign q le l) • koszulOrder q le (ofList l x) := by
2024-12-19 14:25:09 +00:00
rw [koszulOrder_ofList]
rw [smul_smul]
rw [koszulSign_mul_self]
simp
2024-12-19 15:40:04 +00:00
/-- The map of algebras from `FreeAlgebra I` to `FreeAlgebra (Σ i, f i)` taking
the monomial `i` to the sum of elements in `(Σ i, f i)` above `i`, i.e. the sum of the fiber
above `i`. -/
2024-12-20 13:11:22 +00:00
def sumFiber (f : 𝓕 → Type) [∀ i, Fintype (f i)] :
FreeAlgebra 𝓕 →ₐ[] FreeAlgebra (Σ i, f i) :=
2024-12-19 14:25:09 +00:00
FreeAlgebra.lift fun i => ∑ (j : f i), FreeAlgebra.ι ⟨i, j⟩
2024-12-20 13:11:22 +00:00
lemma sumFiber_ι (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) :
2024-12-19 14:25:09 +00:00
sumFiber f (FreeAlgebra.ι i) = ∑ (b : f i), FreeAlgebra.ι ⟨i, b⟩ := by
simp [sumFiber]
2024-12-19 15:40:04 +00:00
/-- Given a list `l : List I` the corresponding element of `FreeAlgebra (Σ i, f i)`
by mapping each `i : I` to the sum of it's fiber in `Σ i, f i` and taking the product of the
result.
For example, in terms of creation and annihlation opperators,
`[φ₁, φ₂, φ₃]` gets taken to `(φ₁⁰ + φ₁¹)(φ₂⁰ + φ₂¹)(φ₃⁰ + φ₃¹)`.
-/
2024-12-20 13:11:22 +00:00
def ofListLift (f : 𝓕 → Type) [∀ i, Fintype (f i)] (l : List 𝓕) (x : ) :
2024-12-19 14:25:09 +00:00
FreeAlgebra (Σ i, f i) :=
sumFiber f (ofList l x)
2024-12-20 13:11:22 +00:00
lemma ofListLift_empty (f : 𝓕 → Type) [∀ i, Fintype (f i)] :
2024-12-19 15:44:32 +00:00
ofListLift f [] 1 = 1 := by
2024-12-19 14:25:09 +00:00
simp only [ofListLift, EmbeddingLike.map_eq_one_iff]
rw [ofList_empty]
exact map_one (sumFiber f)
2024-12-20 13:11:22 +00:00
lemma ofListLift_empty_smul (f : 𝓕 → Type) [∀ i, Fintype (f i)] (x : ) :
2024-12-19 15:44:32 +00:00
ofListLift f [] x = x • 1 := by
2024-12-19 14:25:09 +00:00
simp only [ofListLift, EmbeddingLike.map_eq_one_iff]
rw [ofList_eq_smul_one]
rw [ofList_empty]
simp
2024-12-20 13:11:22 +00:00
lemma ofListLift_cons (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) (r : List 𝓕) (x : ) :
2024-12-19 14:25:09 +00:00
ofListLift f (i :: r) x = (∑ j : f i, FreeAlgebra.ι ⟨i, j⟩) * (ofListLift f r x) := by
rw [ofListLift, ofList_cons_eq_ofList, ofList_singleton, map_mul]
conv_lhs => lhs; rw [sumFiber]
rw [ofListLift]
simp
2024-12-20 13:11:22 +00:00
lemma ofListLift_singleton (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) (x : ) :
2024-12-19 14:25:09 +00:00
ofListLift f [i] x = ∑ j : f i, x • FreeAlgebra.ι ⟨i, j⟩ := by
simp only [ofListLift]
rw [ofList_eq_smul_one, ofList_singleton, map_smul]
rw [sumFiber_ι]
rw [Finset.smul_sum]
2024-12-20 13:11:22 +00:00
lemma ofListLift_singleton_one (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕) :
2024-12-19 14:25:09 +00:00
ofListLift f [i] 1 = ∑ j : f i, FreeAlgebra.ι ⟨i, j⟩ := by
simp only [ofListLift]
rw [ofList_eq_smul_one, ofList_singleton, map_smul]
rw [sumFiber_ι]
rw [Finset.smul_sum]
simp
2024-12-20 13:11:22 +00:00
lemma ofListLift_cons_eq_ofListLift (f : 𝓕 → Type) [∀ i, Fintype (f i)] (i : 𝓕)
(r : List 𝓕) (x : ) :
2024-12-19 14:25:09 +00:00
ofListLift f (i :: r) x = ofListLift f [i] 1 * ofListLift f r x := by
rw [ofListLift_cons, ofListLift_singleton]
simp only [one_smul]
2024-12-20 13:11:22 +00:00
lemma ofListLift_expand (f : 𝓕 → Type) [∀ i, Fintype (f i)] (x : ) :
2024-12-20 13:57:29 +00:00
(l : List 𝓕) → ofListLift f l x = ∑ (a : CreateAnnihilateSect f l), ofList a.toList x
2024-12-19 14:25:09 +00:00
| [] => by
2024-12-20 13:57:29 +00:00
simp only [ofListLift, CreateAnnihilateSect, List.length_nil, List.get_eq_getElem,
Finset.univ_unique, CreateAnnihilateSect.toList, Finset.sum_const, Finset.card_singleton,
2024-12-19 14:25:09 +00:00
one_smul]
rw [ofList_eq_smul_one, map_smul, ofList_empty, ofList_eq_smul_one, ofList_empty, map_one]
| i :: l => by
rw [ofListLift_cons, ofListLift_expand f x l]
2024-12-20 13:57:29 +00:00
conv_rhs => rw [← (CreateAnnihilateSect.extractEquiv
2024-12-19 14:25:09 +00:00
⟨0, by exact Nat.zero_lt_succ l.length⟩).symm.sum_comp (α := FreeAlgebra _)]
erw [Finset.sum_product]
rw [Finset.sum_mul]
conv_lhs =>
rhs
intro n
rw [Finset.mul_sum]
congr
funext j
congr
funext n
rw [← ofList_singleton, ← ofList_pair, one_mul]
rfl
2024-12-20 13:11:22 +00:00
lemma koszulOrder_ofListLift {f : 𝓕 → Type} [∀ i, Fintype (f i)]
(l : List 𝓕) (x : ) :
2024-12-20 13:34:49 +00:00
koszulOrder (fun i => q i.fst) (fun i j => le i.1 j.1) (ofListLift f l x) =
2024-12-20 13:11:22 +00:00
sumFiber f (koszulOrder q le (ofList l x)) := by
2024-12-19 14:25:09 +00:00
rw [koszulOrder_ofList]
rw [map_smul]
change _ = _ • ofListLift _ _ _
rw [ofListLift_expand]
rw [map_sum]
conv_lhs =>
rhs
intro a
rw [koszulOrder_ofList]
2024-12-20 13:57:29 +00:00
rw [CreateAnnihilateSect.toList_koszulSign]
2024-12-19 14:25:09 +00:00
rw [← Finset.smul_sum]
apply congrArg
conv_lhs =>
rhs
intro n
2024-12-20 13:57:29 +00:00
rw [← CreateAnnihilateSect.sort_toList]
2024-12-19 14:25:09 +00:00
rw [ofListLift_expand]
refine Fintype.sum_equiv
2024-12-20 13:11:22 +00:00
((HepLean.List.insertionSortEquiv le l).piCongr fun i => Equiv.cast ?_) _ _ ?_
2024-12-19 14:25:09 +00:00
congr 1
· rw [← HepLean.List.insertionSortEquiv_get]
simp
· intro x
rfl
2024-12-20 13:11:22 +00:00
lemma koszulOrder_ofListLift_eq_ofListLift {f : 𝓕 → Type} [∀ i, Fintype (f i)]
(l : List 𝓕) (x : ) : koszulOrder (fun i => q i.fst) (fun i j => le i.1 j.1)
2024-12-19 14:25:09 +00:00
(ofListLift f l x) =
2024-12-20 13:11:22 +00:00
koszulSign q le l • ofListLift f (List.insertionSort le l) x := by
2024-12-19 14:25:09 +00:00
rw [koszulOrder_ofListLift, koszulOrder_ofList, map_smul]
rfl
end
end Wick