/- 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 Mathlib.Algebra.FreeAlgebra import Mathlib.Algebra.Lie.OfAssociative import Mathlib.Analysis.Complex.Basic import HepLean.Mathematics.List.InsertIdx /-! # Field statistics Basic properties related to whether a field, or list of fields, is bosonic or fermionic. -/ /-- A field can either be bosonic or fermionic in nature. That is to say, they can either have Bose-Einstein statistics or Fermi-Dirac statistics. -/ inductive FieldStatistic : Type where | bosonic : FieldStatistic | fermionic : FieldStatistic deriving DecidableEq namespace FieldStatistic variable {𝓕 : Type} /-- Field statistics form a commuative group equivalent to `β„€β‚‚`. -/ @[simp] instance : CommGroup FieldStatistic where one := bosonic mul a b := match a, b with | bosonic, bosonic => bosonic | bosonic, fermionic => fermionic | fermionic, bosonic => fermionic | fermionic, fermionic => bosonic inv a := a mul_assoc a b c := by cases a <;> cases b <;> cases c <;> dsimp [HMul.hMul] one_mul a := by cases a <;> dsimp [HMul.hMul] mul_one a := by cases a <;> dsimp [HMul.hMul] inv_mul_cancel a := by cases a <;> dsimp only [HMul.hMul, Nat.succ_eq_add_one] <;> rfl mul_comm a b := by cases a <;> cases b <;> rfl @[simp] lemma bosonic_mul_bosonic : bosonic * bosonic = bosonic := rfl @[simp] lemma bosonic_mul_fermionic : bosonic * fermionic = fermionic := rfl @[simp] lemma fermionic_mul_bosonic : fermionic * bosonic = fermionic := rfl @[simp] lemma fermionic_mul_fermionic : fermionic * fermionic = bosonic := rfl @[simp] lemma mul_self (a : FieldStatistic) : a * a = 1 := by cases a <;> rfl /-- Field statics form a finite type. -/ instance : Fintype FieldStatistic where elems := {bosonic, fermionic} complete := by intro c cases c Β· exact Finset.mem_insert_self bosonic {fermionic} Β· refine Finset.insert_eq_self.mp ?_ exact rfl @[simp] lemma fermionic_not_eq_bonsic : Β¬ fermionic = bosonic := by intro h exact FieldStatistic.noConfusion h lemma bonsic_eq_fermionic_false : bosonic = fermionic ↔ false := by simp only [reduceCtorEq, Bool.false_eq_true] @[simp] lemma neq_fermionic_iff_eq_bosonic (a : FieldStatistic) : Β¬ a = fermionic ↔ a = bosonic := by fin_cases a Β· simp Β· simp @[simp] lemma bosonic_neq_iff_fermionic_eq (a : FieldStatistic) : Β¬ bosonic = a ↔ fermionic = a := by fin_cases a Β· simp Β· simp @[simp] lemma fermionic_neq_iff_bosonic_eq (a : FieldStatistic) : Β¬ fermionic = a ↔ bosonic = a := by fin_cases a Β· simp Β· simp lemma eq_self_if_eq_bosonic {a : FieldStatistic} : (if a = bosonic then bosonic else fermionic) = a := by fin_cases a <;> rfl lemma eq_self_if_bosonic_eq {a : FieldStatistic} : (if bosonic = a then bosonic else fermionic) = a := by fin_cases a <;> rfl lemma mul_eq_one_iff (a b : FieldStatistic) : a * b = 1 ↔ a = b := by fin_cases a <;> fin_cases b <;> simp lemma one_eq_mul_iff (a b : FieldStatistic) : 1 = a * b ↔ a = b := by fin_cases a <;> fin_cases b <;> simp lemma mul_eq_iff_eq_mul (a b c : FieldStatistic) : a * b = c ↔ a = b * c := by fin_cases a <;> fin_cases b <;> fin_cases c <;> simp only [bosonic_mul_fermionic, fermionic_not_eq_bonsic, mul_self, reduceCtorEq, fermionic_mul_bosonic, true_iff, iff_true] all_goals rfl lemma mul_eq_iff_eq_mul' (a b c : FieldStatistic) : a * b = c ↔ b = a * c := by fin_cases a <;> fin_cases b <;> fin_cases c <;> simp only [bosonic_mul_fermionic, fermionic_not_eq_bonsic, mul_self, reduceCtorEq, fermionic_mul_bosonic, true_iff, iff_true] all_goals rfl /-- The field statistics of a list of fields is fermionic if ther is an odd number of fermions, otherwise it is bosonic. -/ def ofList (s : 𝓕 β†’ FieldStatistic) : (Ο†s : List 𝓕) β†’ FieldStatistic | [] => bosonic | Ο† :: Ο†s => if s Ο† = ofList s Ο†s then bosonic else fermionic lemma ofList_cons_eq_mul (s : 𝓕 β†’ FieldStatistic) (Ο† : 𝓕) (Ο†s : List 𝓕) : ofList s (Ο† :: Ο†s) = s Ο† * ofList s Ο†s := by have ha (a b : FieldStatistic) : (if a = b then bosonic else fermionic) = a * b := by fin_cases a <;> fin_cases b <;> rfl exact ha (s Ο†) (ofList s Ο†s) lemma ofList_eq_prod (s : 𝓕 β†’ FieldStatistic) : (Ο†s : List 𝓕) β†’ ofList s Ο†s = (List.map s Ο†s).prod | [] => rfl | Ο† :: Ο†s => by rw [ofList_cons_eq_mul, List.map_cons, List.prod_cons, ofList_eq_prod] @[simp] lemma ofList_singleton (s : 𝓕 β†’ FieldStatistic) (Ο† : 𝓕) : ofList s [Ο†] = s Ο† := by simp only [ofList, Fin.isValue] rw [eq_self_if_eq_bosonic] @[simp] lemma ofList_freeMonoid (s : 𝓕 β†’ FieldStatistic) (Ο† : 𝓕) : ofList s (FreeMonoid.of Ο†) = s Ο† := ofList_singleton s Ο† @[simp] lemma ofList_empty (s : 𝓕 β†’ FieldStatistic) : ofList s [] = bosonic := rfl @[simp] lemma ofList_append (s : 𝓕 β†’ FieldStatistic) (Ο†s Ο†s' : List 𝓕) : ofList s (Ο†s ++ Ο†s') = if ofList s Ο†s = ofList s Ο†s' then bosonic else fermionic := by induction Ο†s with | nil => simp only [List.nil_append, ofList_empty, Fin.isValue, eq_self_if_bosonic_eq] | cons a l ih => have hab (a b c : FieldStatistic) : (if a = (if b = c then bosonic else fermionic) then bosonic else fermionic) = if (if a = b then bosonic else fermionic) = c then bosonic else fermionic := by fin_cases a <;> fin_cases b <;> fin_cases c <;> rfl simp only [ofList, List.append_eq, Fin.isValue, ih, hab] lemma ofList_append_eq_mul (s : 𝓕 β†’ FieldStatistic) (Ο†s Ο†s' : List 𝓕) : ofList s (Ο†s ++ Ο†s') = ofList s Ο†s * ofList s Ο†s' := by rw [ofList_append] have ha (a b : FieldStatistic) : (if a = b then bosonic else fermionic) = a * b := by fin_cases a <;> fin_cases b <;> rfl exact ha _ _ lemma ofList_perm (s : 𝓕 β†’ FieldStatistic) {l l' : List 𝓕} (h : l.Perm l') : ofList s l = ofList s l' := by rw [ofList_eq_prod, ofList_eq_prod] exact List.Perm.prod_eq (List.Perm.map s h) lemma ofList_orderedInsert (s : 𝓕 β†’ FieldStatistic) (le1 : 𝓕 β†’ 𝓕 β†’ Prop) [DecidableRel le1] (Ο†s : List 𝓕) (Ο† : 𝓕) : ofList s (List.orderedInsert le1 Ο† Ο†s) = ofList s (Ο† :: Ο†s) := ofList_perm s (List.perm_orderedInsert le1 Ο† Ο†s) @[simp] lemma ofList_insertionSort (s : 𝓕 β†’ FieldStatistic) (le1 : 𝓕 β†’ 𝓕 β†’ Prop) [DecidableRel le1] (Ο†s : List 𝓕) : ofList s (List.insertionSort le1 Ο†s) = ofList s Ο†s := ofList_perm s (List.perm_insertionSort le1 Ο†s) lemma ofList_map_eq_finset_prod (s : 𝓕 β†’ FieldStatistic) : (Ο†s : List 𝓕) β†’ (l : List (Fin Ο†s.length)) β†’ (hl : l.Nodup) β†’ ofList s (l.map Ο†s.get) = ∏ (i : Fin Ο†s.length), if i ∈ l then s Ο†s[i] else 1 | [], [], _ => rfl | [], i :: l, hl => Fin.elim0 i | Ο† :: Ο†s, [], hl => by simp only [List.length_cons, List.map_nil, ofList_empty, List.not_mem_nil, ↓reduceIte, Finset.prod_const_one] rfl | Ο† :: Ο†s, i :: l, hl => by simp only [List.length_cons, List.map_cons, List.get_eq_getElem, List.mem_cons, instCommGroup, Fin.getElem_fin] rw [ofList_cons_eq_mul] rw [ofList_map_eq_finset_prod s (Ο† :: Ο†s) l] have h1 : s (Ο† :: Ο†s)[↑i] = ∏ (j : Fin (Ο† :: Ο†s).length), if j = i then s (Ο† :: Ο†s)[↑i] else 1 := by rw [Fintype.prod_ite_eq'] erw [h1] rw [← Finset.prod_mul_distrib] congr funext a simp only [instCommGroup, List.length_cons, mul_ite, ite_mul, one_mul, mul_one] by_cases ha : a = i Β· simp only [ha, ↓reduceIte, mul_self, true_or] rw [if_neg] rfl simp only [List.length_cons, List.nodup_cons] at hl exact hl.1 Β· simp only [ha, ↓reduceIte, false_or] rfl simp only [List.length_cons, List.nodup_cons] at hl exact hl.2 /-! ## ofList and take -/ section ofListTake open HepLean.List variable (q : 𝓕 β†’ FieldStatistic) lemma ofList_take_insert (n : β„•) (Ο† : 𝓕) (Ο†s : List 𝓕) : ofList q (List.take n Ο†s) = ofList q (List.take n (List.insertIdx n Ο† Ο†s)) := by congr 1 rw [take_insert_same] lemma ofList_take_eraseIdx (n : β„•) (Ο†s : List 𝓕) : ofList q (List.take n (Ο†s.eraseIdx n)) = ofList q (List.take n Ο†s) := by congr 1 rw [take_eraseIdx_same] lemma ofList_take_zero (Ο†s : List 𝓕) : ofList q (List.take 0 Ο†s) = 1 := by simp only [List.take_zero, ofList_empty] rfl lemma ofList_take_succ_cons (n : β„•) (Ο†1 : 𝓕) (Ο†s : List 𝓕) : ofList q ((Ο†1 :: Ο†s).take (n + 1)) = q Ο†1 * ofList q (Ο†s.take n) := by simp only [List.take_succ_cons, instCommGroup] rw [ofList_cons_eq_mul] lemma ofList_take_insertIdx_gt (n m : β„•) (Ο†1 : 𝓕) (Ο†s : List 𝓕) (hn : n < m) : ofList q ((List.insertIdx m Ο†1 Ο†s).take n) = ofList q (Ο†s.take n) := by rw [take_insert_gt Ο†1 n m hn Ο†s] lemma ofList_insert_lt_eq (n m : β„•) (Ο†1 : 𝓕) (Ο†s : List 𝓕) (hn : m ≀ n) (hm : m ≀ Ο†s.length) : ofList q ((List.insertIdx m Ο†1 Ο†s).take (n + 1)) = ofList q ((Ο†1 :: Ο†s).take (n + 1)) := by apply ofList_perm simp only [List.take_succ_cons] refine take_insert_let Ο†1 n m hn Ο†s hm lemma ofList_take_insertIdx_le (n m : β„•) (Ο†1 : 𝓕) (Ο†s : List 𝓕) (hn : m ≀ n) (hm : m ≀ Ο†s.length) : ofList q ((List.insertIdx m Ο†1 Ο†s).take (n + 1)) = q Ο†1 * ofList q (Ο†s.take n) := by rw [ofList_insert_lt_eq, ofList_take_succ_cons] Β· exact hn Β· exact hm end ofListTake end FieldStatistic