/- 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.FieldStatistics.ExchangeSign /-! # Field statistics of a finite set. -/ namespace FieldStatistic variable {๐“• : Type} /-- The field statistic associated with a map `f : Fin n โ†’ ๐“•` (usually `.get` of a list) and a finite set of elements of `Fin n`. -/ def ofFinset {n : โ„•} (q : ๐“• โ†’ FieldStatistic) (f : Fin n โ†’ ๐“•) (a : Finset (Fin n)) : FieldStatistic := ofList q ((a.sort (ยท โ‰ค ยท)).map f) @[simp] lemma ofFinset_empty (q : ๐“• โ†’ FieldStatistic) (f : Fin n โ†’ ๐“•) : ofFinset q f โˆ… = 1 := by simp only [ofFinset, Finset.sort_empty, List.map_nil, ofList_empty] rfl lemma ofFinset_singleton {n : โ„•} (q : ๐“• โ†’ FieldStatistic) (f : Fin n โ†’ ๐“•) (i : Fin n) : ofFinset q f {i} = q (f i) := by simp [ofFinset] lemma ofFinset_finset_map {n m : โ„•} (q : ๐“• โ†’ FieldStatistic) (i : Fin m โ†’ Fin n) (hi : Function.Injective i) (f : Fin n โ†’ ๐“•) (a : Finset (Fin m)) : ofFinset q (f โˆ˜ i) a = ofFinset q f (a.map โŸจi, hiโŸฉ) := by simp only [ofFinset] apply FieldStatistic.ofList_perm rw [โ† List.map_map] refine List.Perm.map f ?_ apply List.perm_of_nodup_nodup_toFinset_eq ยท refine (List.nodup_map_iff_inj_on ?_).mpr ?_ exact Finset.sort_nodup (fun x1 x2 => x1 โ‰ค x2) a simp only [Finset.mem_sort] intro x hx y hy exact fun a => hi a ยท exact Finset.sort_nodup (fun x1 x2 => x1 โ‰ค x2) (Finset.map { toFun := i, inj' := hi } a) ยท ext a simp lemma ofFinset_insert (q : ๐“• โ†’ FieldStatistic) (ฯ†s : List ๐“•) (a : Finset (Fin ฯ†s.length)) (i : Fin ฯ†s.length) (h : i โˆ‰ a) : ofFinset q ฯ†s.get (Insert.insert i a) = (q ฯ†s[i]) * ofFinset q ฯ†s.get a := by simp only [ofFinset, instCommGroup, Fin.getElem_fin] rw [โ† ofList_cons_eq_mul] have h1 : (ฯ†s[โ†‘i] :: List.map ฯ†s.get (Finset.sort (fun x1 x2 => x1 โ‰ค x2) a)) = List.map ฯ†s.get (i :: Finset.sort (fun x1 x2 => x1 โ‰ค x2) a) := by simp erw [h1] apply ofList_perm refine List.Perm.map ฯ†s.get ?_ refine (List.perm_ext_iff_of_nodup ?_ ?_).mpr ?_ ยท exact Finset.sort_nodup (fun x1 x2 => x1 โ‰ค x2) (Insert.insert i a) ยท simp only [List.nodup_cons, Finset.mem_sort, Finset.sort_nodup, and_true] exact h intro a simp lemma ofFinset_erase (q : ๐“• โ†’ FieldStatistic) (ฯ†s : List ๐“•) (a : Finset (Fin ฯ†s.length)) (i : Fin ฯ†s.length) (h : i โˆˆ a) : ofFinset q ฯ†s.get (a.erase i) = (q ฯ†s[i]) * ofFinset q ฯ†s.get a := by have ha : a = Insert.insert i (a.erase i) := by exact Eq.symm (Finset.insert_erase h) conv_rhs => rw [ha] rw [ofFinset_insert] rw [โ† mul_assoc] simp only [instCommGroup, Fin.getElem_fin, mul_self, one_mul] simp lemma ofFinset_eq_prod (q : ๐“• โ†’ FieldStatistic) (ฯ†s : List ๐“•) (a : Finset (Fin ฯ†s.length)) : ofFinset q ฯ†s.get a = โˆ (i : Fin ฯ†s.length), if i โˆˆ a then (q ฯ†s[i]) else 1 := by rw [ofFinset] rw [ofList_map_eq_finset_prod] congr funext i simp only [Finset.mem_sort, Fin.getElem_fin] exact Finset.sort_nodup (fun x1 x2 => x1 โ‰ค x2) a lemma ofFinset_union (q : ๐“• โ†’ FieldStatistic) (ฯ†s : List ๐“•) (a b : Finset (Fin ฯ†s.length)) : ofFinset q ฯ†s.get a * ofFinset q ฯ†s.get b = ofFinset q ฯ†s.get ((a โˆช b) \ (a โˆฉ b)) := by rw [ofFinset_eq_prod, ofFinset_eq_prod, ofFinset_eq_prod] rw [โ† Finset.prod_mul_distrib] congr funext x simp only [instCommGroup, Fin.getElem_fin, mul_ite, ite_mul, mul_self, one_mul, mul_one, Finset.mem_sdiff, Finset.mem_union, Finset.mem_inter, not_and] split ยท rename_i h simp [h] ยท rename_i h simp [h] lemma ofFinset_union_disjoint (q : ๐“• โ†’ FieldStatistic) (ฯ†s : List ๐“•) (a b : Finset (Fin ฯ†s.length)) (h : Disjoint a b) : ofFinset q ฯ†s.get a * ofFinset q ฯ†s.get b = ofFinset q ฯ†s.get (a โˆช b) := by rw [ofFinset_union, Finset.disjoint_iff_inter_eq_empty.mp h] simp lemma ofFinset_filter_mul_neg (q : ๐“• โ†’ FieldStatistic) (ฯ†s : List ๐“•) (a : Finset (Fin ฯ†s.length)) (p : Fin ฯ†s.length โ†’ Prop) [DecidablePred p] : ofFinset q ฯ†s.get (Finset.filter p a) * ofFinset q ฯ†s.get (Finset.filter (fun i => ยฌ p i) a) = ofFinset q ฯ†s.get a := by rw [ofFinset_union_disjoint] congr exact Finset.filter_union_filter_neg_eq p a exact Finset.disjoint_filter_filter_neg a a p lemma ofFinset_filter (q : ๐“• โ†’ FieldStatistic) (ฯ†s : List ๐“•) (a : Finset (Fin ฯ†s.length)) (p : Fin ฯ†s.length โ†’ Prop) [DecidablePred p] : ofFinset q ฯ†s.get (Finset.filter p a) = ofFinset q ฯ†s.get (Finset.filter (fun i => ยฌ p i) a) * ofFinset q ฯ†s.get a := by rw [โ† ofFinset_filter_mul_neg q ฯ†s a p] conv_rhs => rhs rw [mul_comm] rw [โ† mul_assoc] simp end FieldStatistic