/- 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 -/ import HepLean.PerturbationTheory.Wick.CreateAnnihilateSection import HepLean.PerturbationTheory.Wick.KoszulOrder /-! # Koszul signs and ordering for lists and algebras -/ namespace Wick open HepLean.List open FieldStatistic noncomputable section variable {𝓕 : Type} (q : 𝓕 β†’ FieldStatistic) (le : 𝓕 β†’ 𝓕 β†’ Prop) [DecidableRel le] /-- The element of the free algebra `FreeAlgebra β„‚ I` associated with a `List I`. -/ def ofList (l : List 𝓕) (x : β„‚) : FreeAlgebra β„‚ 𝓕 := FreeAlgebra.equivMonoidAlgebraFreeMonoid.symm (MonoidAlgebra.single l x) lemma ofList_pair (l r : List 𝓕) (x y : β„‚) : 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 lemma ofList_triple (la lb lc : List 𝓕) (xa xb xc : β„‚) : ofList (la ++ lb ++ lc) (xa * xb * xc) = ofList la xa * ofList lb xb * ofList lc xc := by rw [ofList_pair, ofList_pair] lemma ofList_triple_assoc (la lb lc : List 𝓕) (xa xb xc : β„‚) : 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)) lemma ofList_cons_eq_ofList (l : List 𝓕) (i : 𝓕) (x : β„‚) : 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 lemma ofList_singleton (i : 𝓕) : 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 lemma ofList_eq_smul_one (l : List 𝓕) (x : β„‚) : ofList l x = x β€’ ofList l 1 := by simp only [ofList] rw [← map_smul] simp lemma ofList_empty : ofList [] 1 = (1 : FreeAlgebra β„‚ 𝓕) := by simp only [ofList, EmbeddingLike.map_eq_one_iff] rfl lemma ofList_empty' : ofList [] x = x β€’ (1 : FreeAlgebra β„‚ 𝓕) := by rw [ofList_eq_smul_one, ofList_empty] lemma koszulOrder_ofList (l : List 𝓕) (x : β„‚) : koszulOrder q le (ofList l x) = (koszulSign q le l) β€’ ofList (List.insertionSort le l) x := by rw [ofList] rw [koszulOrder_single] change ofList (List.insertionSort le l) _ = _ rw [ofList_eq_smul_one] conv_rhs => rw [ofList_eq_smul_one] rw [smul_smul] 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 rw [koszulOrder_ofList] rw [smul_smul] rw [koszulSign_mul_self] simp /-- 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`. -/ def sumFiber (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] : FreeAlgebra β„‚ 𝓕 →ₐ[β„‚] FreeAlgebra β„‚ (Ξ£ i, f i) := FreeAlgebra.lift β„‚ fun i => βˆ‘ (j : f i), FreeAlgebra.ΞΉ β„‚ ⟨i, j⟩ lemma sumFiber_ΞΉ (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (i : 𝓕) : sumFiber f (FreeAlgebra.ΞΉ β„‚ i) = βˆ‘ (b : f i), FreeAlgebra.ΞΉ β„‚ ⟨i, b⟩ := by simp [sumFiber] /-- 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 `(φ₁⁰ + φ₁¹)(φ₂⁰ + Ο†β‚‚ΒΉ)(φ₃⁰ + φ₃¹)`. -/ def ofListLift (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (l : List 𝓕) (x : β„‚) : FreeAlgebra β„‚ (Ξ£ i, f i) := sumFiber f (ofList l x) lemma ofListLift_empty (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] : ofListLift f [] 1 = 1 := by simp only [ofListLift, EmbeddingLike.map_eq_one_iff] rw [ofList_empty] exact map_one (sumFiber f) lemma ofListLift_empty_smul (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (x : β„‚) : ofListLift f [] x = x β€’ 1 := by simp only [ofListLift, EmbeddingLike.map_eq_one_iff] rw [ofList_eq_smul_one] rw [ofList_empty] simp lemma ofListLift_cons (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (i : 𝓕) (r : List 𝓕) (x : β„‚) : 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 lemma ofListLift_singleton (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (i : 𝓕) (x : β„‚) : 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] lemma ofListLift_singleton_one (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (i : 𝓕) : 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 lemma ofListLift_cons_eq_ofListLift (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (i : 𝓕) (r : List 𝓕) (x : β„‚) : ofListLift f (i :: r) x = ofListLift f [i] 1 * ofListLift f r x := by rw [ofListLift_cons, ofListLift_singleton] simp only [one_smul] lemma ofListLift_expand (f : 𝓕 β†’ Type) [βˆ€ i, Fintype (f i)] (x : β„‚) : (l : List 𝓕) β†’ ofListLift f l x = βˆ‘ (a : CreateAnnihilateSect f l), ofList a.toList x | [] => by simp only [ofListLift, CreateAnnihilateSect, List.length_nil, List.get_eq_getElem, Finset.univ_unique, CreateAnnihilateSect.toList, Finset.sum_const, Finset.card_singleton, 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] conv_rhs => rw [← (CreateAnnihilateSect.extractEquiv ⟨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 lemma koszulOrder_ofListLift {f : 𝓕 β†’ Type} [βˆ€ i, Fintype (f i)] (l : List 𝓕) (x : β„‚) : koszulOrder (fun i => q i.fst) (fun i j => le i.1 j.1) (ofListLift f l x) = sumFiber f (koszulOrder q le (ofList l x)) := by rw [koszulOrder_ofList] rw [map_smul] change _ = _ β€’ ofListLift _ _ _ rw [ofListLift_expand] rw [map_sum] conv_lhs => rhs intro a rw [koszulOrder_ofList] rw [CreateAnnihilateSect.toList_koszulSign] rw [← Finset.smul_sum] apply congrArg conv_lhs => rhs intro n rw [← CreateAnnihilateSect.sort_toList] rw [ofListLift_expand] refine Fintype.sum_equiv ((HepLean.List.insertionSortEquiv le l).piCongr fun i => Equiv.cast ?_) _ _ ?_ congr 1 Β· rw [← HepLean.List.insertionSortEquiv_get] simp Β· intro x rfl 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) (ofListLift f l x) = koszulSign q le l β€’ ofListLift f (List.insertionSort le l) x := by rw [koszulOrder_ofListLift, koszulOrder_ofList, map_smul] rfl end end Wick