/- Copyright (c) 2025 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.FieldSpecification.CrAnSection import HepLean.PerturbationTheory.FieldSpecification.NormalOrder /-! # Time ordering of states -/ namespace FieldSpecification variable {๐“• : FieldSpecification} /-! ## Time ordering for states -/ /-- The time ordering relation on states. We have that `timeOrderRel ฯ†0 ฯ†1` is true if and only if `ฯ†1` has a time less-then or equal to `ฯ†0`, or `ฯ†1` is a negative asymptotic state, or `ฯ†0` is a positive asymptotic state. -/ def timeOrderRel : ๐“•.States โ†’ ๐“•.States โ†’ Prop | States.outAsymp _, _ => True | States.position ฯ†0, States.position ฯ†1 => ฯ†1.2 0 โ‰ค ฯ†0.2 0 | States.position _, States.inAsymp _ => True | States.position _, States.outAsymp _ => False | States.inAsymp _, States.outAsymp _ => False | States.inAsymp _, States.position _ => False | States.inAsymp _, States.inAsymp _ => True /-- The relation `timeOrderRel` is decidable, but not computablly so due to `Real.decidableLE`. -/ noncomputable instance : (ฯ† ฯ†' : ๐“•.States) โ†’ Decidable (timeOrderRel ฯ† ฯ†') | States.outAsymp _, _ => isTrue True.intro | States.position ฯ†0, States.position ฯ†1 => inferInstanceAs (Decidable (ฯ†1.2 0 โ‰ค ฯ†0.2 0)) | States.position _, States.inAsymp _ => isTrue True.intro | States.position _, States.outAsymp _ => isFalse (fun a => a) | States.inAsymp _, States.outAsymp _ => isFalse (fun a => a) | States.inAsymp _, States.position _ => isFalse (fun a => a) | States.inAsymp _, States.inAsymp _ => isTrue True.intro /-- Time ordering is total. -/ instance : IsTotal ๐“•.States ๐“•.timeOrderRel where total a b := by cases a <;> cases b <;> simp only [or_self, or_false, or_true, timeOrderRel, Fin.isValue, implies_true, imp_self, IsEmpty.forall_iff] exact LinearOrder.le_total _ _ /-- Time ordering is transitive. -/ instance : IsTrans ๐“•.States ๐“•.timeOrderRel where trans a b c := by cases a <;> cases b <;> cases c <;> simp only [timeOrderRel, Fin.isValue, implies_true, imp_self, IsEmpty.forall_iff] exact fun h1 h2 => Preorder.le_trans _ _ _ h2 h1 noncomputable section open FieldStatistic open HepLean.List /-- Given a list `ฯ† :: ฯ†s` of states, the (zero-based) position of the state which is of maximum time. For example - for the list `[ฯ†1(t = 4), ฯ†2(t = 5), ฯ†3(t = 3), ฯ†4(t = 5)]` this would return `1`. This is defined for a list `ฯ† :: ฯ†s` instead of `ฯ†s` to ensure that such a position exists. -/ def maxTimeFieldPos (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : โ„• := insertionSortMinPos timeOrderRel ฯ† ฯ†s lemma maxTimeFieldPos_lt_length (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : maxTimeFieldPos ฯ† ฯ†s < (ฯ† :: ฯ†s).length := by simp [maxTimeFieldPos] /-- Given a list `ฯ† :: ฯ†s` of states, the left-most state of maximum time, if there are more. As an example: - for the list `[ฯ†1(t = 4), ฯ†2(t = 5), ฯ†3(t = 3), ฯ†4(t = 5)]` this would return `ฯ†2(t = 5)`. It is the state at the position `maxTimeFieldPos ฯ† ฯ†s`. -/ def maxTimeField (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : ๐“•.States := insertionSortMin timeOrderRel ฯ† ฯ†s /-- Given a list `ฯ† :: ฯ†s` of states, the list with the left-most state of maximum time removed. As an example: - for the list `[ฯ†1(t = 4), ฯ†2(t = 5), ฯ†3(t = 3), ฯ†4(t = 5)]` this would return `[ฯ†1(t = 4), ฯ†3(t = 3), ฯ†4(t = 5)]`. -/ def eraseMaxTimeField (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : List ๐“•.States := insertionSortDropMinPos timeOrderRel ฯ† ฯ†s @[simp] lemma eraseMaxTimeField_length (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : (eraseMaxTimeField ฯ† ฯ†s).length = ฯ†s.length := by simp [eraseMaxTimeField, insertionSortDropMinPos, eraseIdx_length'] lemma maxTimeFieldPos_lt_eraseMaxTimeField_length_succ (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : maxTimeFieldPos ฯ† ฯ†s < (eraseMaxTimeField ฯ† ฯ†s).length.succ := by simp only [eraseMaxTimeField_length, Nat.succ_eq_add_one] exact maxTimeFieldPos_lt_length ฯ† ฯ†s /-- Given a list `ฯ† :: ฯ†s` of states, the position of the left-most state of maximum time as an eement of `Fin (eraseMaxTimeField ฯ† ฯ†s).length.succ`. As an example: - for the list `[ฯ†1(t = 4), ฯ†2(t = 5), ฯ†3(t = 3), ฯ†4(t = 5)]` this would return `โŸจ1,...โŸฉ`. -/ def maxTimeFieldPosFin (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : Fin (eraseMaxTimeField ฯ† ฯ†s).length.succ := insertionSortMinPosFin timeOrderRel ฯ† ฯ†s lemma lt_maxTimeFieldPosFin_not_timeOrder (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) (i : Fin (eraseMaxTimeField ฯ† ฯ†s).length) (hi : (maxTimeFieldPosFin ฯ† ฯ†s).succAbove i < maxTimeFieldPosFin ฯ† ฯ†s) : ยฌ timeOrderRel ((eraseMaxTimeField ฯ† ฯ†s)[i.val]) (maxTimeField ฯ† ฯ†s) := by exact insertionSortMin_lt_mem_insertionSortDropMinPos_of_lt timeOrderRel ฯ† ฯ†s i hi lemma timeOrder_maxTimeField (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) (i : Fin (eraseMaxTimeField ฯ† ฯ†s).length) : timeOrderRel (maxTimeField ฯ† ฯ†s) ((eraseMaxTimeField ฯ† ฯ†s)[i.val]) := by exact insertionSortMin_lt_mem_insertionSortDropMinPos timeOrderRel ฯ† ฯ†s _ /-- The sign associated with putting a list of states into time order (with the state of greatest time to the left). We pick up a minus sign for every fermion paired crossed. -/ def timeOrderSign (ฯ†s : List ๐“•.States) : โ„‚ := Wick.koszulSign ๐“•.statesStatistic ๐“•.timeOrderRel ฯ†s @[simp] lemma timeOrderSign_nil : timeOrderSign (๐“• := ๐“•) [] = 1 := by simp only [timeOrderSign] rfl lemma timeOrderSign_pair_ordered {ฯ† ฯˆ : ๐“•.States} (h : timeOrderRel ฯ† ฯˆ) : timeOrderSign [ฯ†, ฯˆ] = 1 := by simp only [timeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, ite_eq_left_iff, ite_eq_right_iff, and_imp] exact fun h' => False.elim (h' h) lemma timeOrderSign_pair_not_ordered {ฯ† ฯˆ : ๐“•.States} (h : ยฌ timeOrderRel ฯ† ฯˆ) : timeOrderSign [ฯ†, ฯˆ] = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯˆ) := by simp only [timeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, instCommGroup.eq_1] rw [if_neg h] simp [FieldStatistic.exchangeSign_eq_if] lemma timerOrderSign_of_eraseMaxTimeField (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : timeOrderSign (eraseMaxTimeField ฯ† ฯ†s) = timeOrderSign (ฯ† :: ฯ†s) * ๐“ข(๐“• |>โ‚› maxTimeField ฯ† ฯ†s, ๐“• |>โ‚› (ฯ† :: ฯ†s).take (maxTimeFieldPos ฯ† ฯ†s)) := by rw [eraseMaxTimeField, insertionSortDropMinPos, timeOrderSign, Wick.koszulSign_eraseIdx_insertionSortMinPos] rw [โ† timeOrderSign, โ† maxTimeField] rfl /-- The time ordering of a list of states. A schematic example is: - `normalOrderList [ฯ†1(t = 4), ฯ†2(t = 5), ฯ†3(t = 3), ฯ†4(t = 5)]` is equal to `[ฯ†2(t = 5), ฯ†4(t = 5), ฯ†1(t = 4), ฯ†3(t = 3)]` -/ def timeOrderList (ฯ†s : List ๐“•.States) : List ๐“•.States := List.insertionSort ๐“•.timeOrderRel ฯ†s lemma timeOrderList_pair_ordered {ฯ† ฯˆ : ๐“•.States} (h : timeOrderRel ฯ† ฯˆ) : timeOrderList [ฯ†, ฯˆ] = [ฯ†, ฯˆ] := by simp only [timeOrderList, List.insertionSort, List.orderedInsert, ite_eq_left_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h' h) lemma timeOrderList_pair_not_ordered {ฯ† ฯˆ : ๐“•.States} (h : ยฌ timeOrderRel ฯ† ฯˆ) : timeOrderList [ฯ†, ฯˆ] = [ฯˆ, ฯ†] := by simp only [timeOrderList, List.insertionSort, List.orderedInsert, ite_eq_right_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h h') @[simp] lemma timeOrderList_nil : timeOrderList (๐“• := ๐“•) [] = [] := by simp [timeOrderList] lemma timeOrderList_eq_maxTimeField_timeOrderList (ฯ† : ๐“•.States) (ฯ†s : List ๐“•.States) : timeOrderList (ฯ† :: ฯ†s) = maxTimeField ฯ† ฯ†s :: timeOrderList (eraseMaxTimeField ฯ† ฯ†s) := by exact insertionSort_eq_insertionSortMin_cons timeOrderRel ฯ† ฯ†s /-! ## Time ordering for CrAnStates -/ /-! ## timeOrderRel -/ /-- The time ordering relation on CrAnStates. -/ def crAnTimeOrderRel (a b : ๐“•.CrAnStates) : Prop := ๐“•.timeOrderRel a.1 b.1 /-- The relation `crAnTimeOrderRel` is decidable, but not computablly so due to `Real.decidableLE`. -/ noncomputable instance (ฯ† ฯ†' : ๐“•.CrAnStates) : Decidable (crAnTimeOrderRel ฯ† ฯ†') := inferInstanceAs (Decidable (๐“•.timeOrderRel ฯ†.1 ฯ†'.1)) /-- Time ordering of `CrAnStates` is total. -/ instance : IsTotal ๐“•.CrAnStates ๐“•.crAnTimeOrderRel where total a b := IsTotal.total (r := ๐“•.timeOrderRel) a.1 b.1 /-- Time ordering of `CrAnStates` is transitive. -/ instance : IsTrans ๐“•.CrAnStates ๐“•.crAnTimeOrderRel where trans a b c := IsTrans.trans (r := ๐“•.timeOrderRel) a.1 b.1 c.1 @[simp] lemma crAnTimeOrderRel_refl (ฯ† : ๐“•.CrAnStates) : crAnTimeOrderRel ฯ† ฯ† := by exact (IsTotal.to_isRefl (r := ๐“•.crAnTimeOrderRel)).refl ฯ† /-- The sign associated with putting a list of `CrAnStates` into time order (with the state of greatest time to the left). We pick up a minus sign for every fermion paired crossed. -/ def crAnTimeOrderSign (ฯ†s : List ๐“•.CrAnStates) : โ„‚ := Wick.koszulSign ๐“•.crAnStatistics ๐“•.crAnTimeOrderRel ฯ†s @[simp] lemma crAnTimeOrderSign_nil : crAnTimeOrderSign (๐“• := ๐“•) [] = 1 := by simp only [crAnTimeOrderSign] rfl lemma crAnTimeOrderSign_pair_ordered {ฯ† ฯˆ : ๐“•.CrAnStates} (h : crAnTimeOrderRel ฯ† ฯˆ) : crAnTimeOrderSign [ฯ†, ฯˆ] = 1 := by simp only [crAnTimeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, ite_eq_left_iff, ite_eq_right_iff, and_imp] exact fun h' => False.elim (h' h) lemma crAnTimeOrderSign_pair_not_ordered {ฯ† ฯˆ : ๐“•.CrAnStates} (h : ยฌ crAnTimeOrderRel ฯ† ฯˆ) : crAnTimeOrderSign [ฯ†, ฯˆ] = ๐“ข(๐“• |>โ‚› ฯ†, ๐“• |>โ‚› ฯˆ) := by simp only [crAnTimeOrderSign, Wick.koszulSign, Wick.koszulSignInsert, mul_one, instCommGroup.eq_1] rw [if_neg h] simp [FieldStatistic.exchangeSign_eq_if] lemma crAnTimeOrderSign_swap_eq_time {ฯ† ฯˆ : ๐“•.CrAnStates} (h1 : crAnTimeOrderRel ฯ† ฯˆ) (h2 : crAnTimeOrderRel ฯˆ ฯ†) (ฯ†s ฯ†s' : List ๐“•.CrAnStates) : crAnTimeOrderSign (ฯ†s ++ ฯ† :: ฯˆ :: ฯ†s') = crAnTimeOrderSign (ฯ†s ++ ฯˆ :: ฯ† :: ฯ†s') := by exact Wick.koszulSign_swap_eq_rel _ _ h1 h2 _ _ /-- Sort a list of `CrAnStates` based on `crAnTimeOrderRel`. -/ def crAnTimeOrderList (ฯ†s : List ๐“•.CrAnStates) : List ๐“•.CrAnStates := List.insertionSort ๐“•.crAnTimeOrderRel ฯ†s @[simp] lemma crAnTimeOrderList_nil : crAnTimeOrderList (๐“• := ๐“•) [] = [] := by simp [crAnTimeOrderList] lemma crAnTimeOrderList_pair_ordered {ฯ† ฯˆ : ๐“•.CrAnStates} (h : crAnTimeOrderRel ฯ† ฯˆ) : crAnTimeOrderList [ฯ†, ฯˆ] = [ฯ†, ฯˆ] := by simp only [crAnTimeOrderList, List.insertionSort, List.orderedInsert, ite_eq_left_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h' h) lemma crAnTimeOrderList_pair_not_ordered {ฯ† ฯˆ : ๐“•.CrAnStates} (h : ยฌ crAnTimeOrderRel ฯ† ฯˆ) : crAnTimeOrderList [ฯ†, ฯˆ] = [ฯˆ, ฯ†] := by simp only [crAnTimeOrderList, List.insertionSort, List.orderedInsert, ite_eq_right_iff, List.cons.injEq, and_true] exact fun h' => False.elim (h h') lemma orderedInsert_swap_eq_time {ฯ† ฯˆ : ๐“•.CrAnStates} (h1 : crAnTimeOrderRel ฯ† ฯˆ) (h2 : crAnTimeOrderRel ฯˆ ฯ†) (ฯ†s : List ๐“•.CrAnStates) : List.orderedInsert crAnTimeOrderRel ฯ† (List.orderedInsert crAnTimeOrderRel ฯˆ ฯ†s) = List.takeWhile (fun b => ยฌ crAnTimeOrderRel ฯˆ b) ฯ†s ++ ฯ† :: ฯˆ :: List.dropWhile (fun b => ยฌ crAnTimeOrderRel ฯˆ b) ฯ†s := by rw [List.orderedInsert_eq_take_drop crAnTimeOrderRel ฯˆ ฯ†s] simp only [decide_not] rw [List.orderedInsert_eq_take_drop] simp only [decide_not] have h1 (b : ๐“•.CrAnStates) : (crAnTimeOrderRel ฯ† b) โ†” (crAnTimeOrderRel ฯˆ b) := Iff.intro (fun h => IsTrans.trans _ _ _ h2 h) (fun h => IsTrans.trans _ _ _ h1 h) simp only [h1] rw [List.takeWhile_append] rw [List.takeWhile_takeWhile] simp only [Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, and_self, decide_not, โ†“reduceIte, crAnTimeOrderRel_refl, decide_true, Bool.false_eq_true, not_false_eq_true, List.takeWhile_cons_of_neg, List.append_nil, List.append_cancel_left_eq, List.cons.injEq, true_and] rw [List.dropWhile_append] simp only [List.isEmpty_eq_true, List.dropWhile_eq_nil_iff, Bool.not_eq_eq_eq_not, Bool.not_true, decide_eq_false_iff_not, crAnTimeOrderRel_refl, decide_true, Bool.false_eq_true, not_false_eq_true, List.dropWhile_cons_of_neg, ite_eq_left_iff, not_forall, Classical.not_imp, Decidable.not_not, List.append_left_eq_self, forall_exists_index, and_imp] intro x hx hxฯˆ intro y hy simpa using List.mem_takeWhile_imp hy lemma orderedInsert_in_swap_eq_time {ฯ† ฯˆ ฯ†': ๐“•.CrAnStates} (h1 : crAnTimeOrderRel ฯ† ฯˆ) (h2 : crAnTimeOrderRel ฯˆ ฯ†) : (ฯ†s ฯ†s' : List ๐“•.CrAnStates) โ†’ โˆƒ l1 l2, List.orderedInsert crAnTimeOrderRel ฯ†' (ฯ†s ++ ฯ† :: ฯˆ :: ฯ†s') = l1 ++ ฯ† :: ฯˆ :: l2 โˆง List.orderedInsert crAnTimeOrderRel ฯ†' (ฯ†s ++ ฯˆ :: ฯ† :: ฯ†s') = l1 ++ ฯˆ :: ฯ† :: l2 | [], ฯ†s' => by have h1 (b : ๐“•.CrAnStates) : (crAnTimeOrderRel b ฯ†) โ†” (crAnTimeOrderRel b ฯˆ) := Iff.intro (fun h => IsTrans.trans _ _ _ h h1) (fun h => IsTrans.trans _ _ _ h h2) by_cases h : crAnTimeOrderRel ฯ†' ฯ† ยท simp only [List.orderedInsert, h, โ†“reduceIte, โ† h1 ฯ†'] use [ฯ†'], ฯ†s' simp ยท simp only [List.orderedInsert, h, โ†“reduceIte, โ† h1 ฯ†'] use [], List.orderedInsert crAnTimeOrderRel ฯ†' ฯ†s' simp | ฯ†'' :: ฯ†s, ฯ†s' => by obtain โŸจl1, l2, hlโŸฉ := orderedInsert_in_swap_eq_time (ฯ†' := ฯ†') h1 h2 ฯ†s ฯ†s' simp only [List.orderedInsert, List.append_eq] rw [hl.1, hl.2] by_cases h : crAnTimeOrderRel ฯ†' ฯ†'' ยท simp only [h, โ†“reduceIte] use (ฯ†' :: ฯ†'' :: ฯ†s), ฯ†s' simp ยท simp only [h, โ†“reduceIte] use (ฯ†'' :: l1), l2 simp lemma crAnTimeOrderList_swap_eq_time {ฯ† ฯˆ : ๐“•.CrAnStates} (h1 : crAnTimeOrderRel ฯ† ฯˆ) (h2 : crAnTimeOrderRel ฯˆ ฯ†) : (ฯ†s ฯ†s' : List ๐“•.CrAnStates) โ†’ โˆƒ (l1 l2 : List ๐“•.CrAnStates), crAnTimeOrderList (ฯ†s ++ ฯ† :: ฯˆ :: ฯ†s') = l1 ++ ฯ† :: ฯˆ :: l2 โˆง crAnTimeOrderList (ฯ†s ++ ฯˆ :: ฯ† :: ฯ†s') = l1 ++ ฯˆ :: ฯ† :: l2 | [], ฯ†s' => by simp only [crAnTimeOrderList] simp only [List.insertionSort] use List.takeWhile (fun b => ยฌ crAnTimeOrderRel ฯˆ b) (List.insertionSort crAnTimeOrderRel ฯ†s'), List.dropWhile (fun b => ยฌ crAnTimeOrderRel ฯˆ b) (List.insertionSort crAnTimeOrderRel ฯ†s') apply And.intro ยท exact orderedInsert_swap_eq_time h1 h2 _ ยท have h1' (b : ๐“•.CrAnStates) : (crAnTimeOrderRel ฯ† b) โ†” (crAnTimeOrderRel ฯˆ b) := Iff.intro (fun h => IsTrans.trans _ _ _ h2 h) (fun h => IsTrans.trans _ _ _ h1 h) simp only [โ† h1', decide_not] simpa using orderedInsert_swap_eq_time h2 h1 _ | ฯ†'' :: ฯ†s, ฯ†s' => by rw [crAnTimeOrderList, crAnTimeOrderList] simp only [List.insertionSort, List.append_eq] obtain โŸจl1, l2, hlโŸฉ := crAnTimeOrderList_swap_eq_time h1 h2 ฯ†s ฯ†s' simp only [crAnTimeOrderList] at hl rw [hl.1, hl.2] obtain โŸจl1', l2', hl'โŸฉ := orderedInsert_in_swap_eq_time (ฯ†' := ฯ†'') h1 h2 l1 l2 rw [hl'.1, hl'.2] use l1', l2' /-! ## Relationship to sections -/ lemma koszulSignInsert_crAnTimeOrderRel_crAnSection {ฯ† : ๐“•.States} {ฯˆ : ๐“•.CrAnStates} (h : ฯˆ.1 = ฯ†) : {ฯ†s : List ๐“•.States} โ†’ (ฯˆs : CrAnSection ฯ†s) โ†’ Wick.koszulSignInsert ๐“•.crAnStatistics ๐“•.crAnTimeOrderRel ฯˆ ฯˆs.1 = Wick.koszulSignInsert ๐“•.statesStatistic ๐“•.timeOrderRel ฯ† ฯ†s | [], โŸจ[], hโŸฉ => by simp [Wick.koszulSignInsert] | ฯ†' :: ฯ†s, โŸจฯˆ' :: ฯˆs, h1โŸฉ => by simp only [Wick.koszulSignInsert, crAnTimeOrderRel, h] simp only [List.map_cons, List.cons.injEq] at h1 have hi := koszulSignInsert_crAnTimeOrderRel_crAnSection h (ฯ†s := ฯ†s) โŸจฯˆs, h1.2โŸฉ rw [hi] congr ยท exact h1.1 ยท simp only [crAnStatistics, crAnStatesToStates, Function.comp_apply] congr ยท simp only [crAnStatistics, crAnStatesToStates, Function.comp_apply] congr exact h1.1 @[simp] lemma crAnTimeOrderSign_crAnSection : {ฯ†s : List ๐“•.States} โ†’ (ฯˆs : CrAnSection ฯ†s) โ†’ crAnTimeOrderSign ฯˆs.1 = timeOrderSign ฯ†s | [], โŸจ[], hโŸฉ => by simp | ฯ† :: ฯ†s, โŸจฯˆ :: ฯˆs, hโŸฉ => by simp only [crAnTimeOrderSign, Wick.koszulSign, timeOrderSign] simp only [List.map_cons, List.cons.injEq] at h congr 1 ยท rw [koszulSignInsert_crAnTimeOrderRel_crAnSection h.1 โŸจฯˆs, h.2โŸฉ] ยท exact crAnTimeOrderSign_crAnSection โŸจฯˆs, h.2โŸฉ lemma orderedInsert_crAnTimeOrderRel_crAnSection {ฯ† : ๐“•.States} {ฯˆ : ๐“•.CrAnStates} (h : ฯˆ.1 = ฯ†) : {ฯ†s : List ๐“•.States} โ†’ (ฯˆs : CrAnSection ฯ†s) โ†’ (List.orderedInsert ๐“•.crAnTimeOrderRel ฯˆ ฯˆs.1).map ๐“•.crAnStatesToStates = List.orderedInsert ๐“•.timeOrderRel ฯ† ฯ†s | [], โŸจ[], _โŸฉ => by simp only [List.orderedInsert, List.map_cons, List.map_nil, List.cons.injEq, and_true] exact h | ฯ†' :: ฯ†s, โŸจฯˆ' :: ฯˆs, h1โŸฉ => by simp only [List.orderedInsert, crAnTimeOrderRel, h] simp only [List.map_cons, List.cons.injEq] at h1 by_cases hr : timeOrderRel ฯ† ฯ†' ยท simp only [hr, โ†“reduceIte] rw [โ† h1.1] at hr simp only [crAnStatesToStates] at hr simp only [hr, โ†“reduceIte, List.map_cons, List.cons.injEq] exact And.intro h (And.intro h1.1 h1.2) ยท simp only [hr, โ†“reduceIte] rw [โ† h1.1] at hr simp only [crAnStatesToStates] at hr simp only [hr, โ†“reduceIte, List.map_cons, List.cons.injEq] apply And.intro h1.1 exact orderedInsert_crAnTimeOrderRel_crAnSection h โŸจฯˆs, h1.2โŸฉ lemma crAnTimeOrderList_crAnSection_is_crAnSection : {ฯ†s : List ๐“•.States} โ†’ (ฯˆs : CrAnSection ฯ†s) โ†’ (crAnTimeOrderList ฯˆs.1).map ๐“•.crAnStatesToStates = timeOrderList ฯ†s | [], โŸจ[], hโŸฉ => by simp | ฯ† :: ฯ†s, โŸจฯˆ :: ฯˆs, hโŸฉ => by simp only [crAnTimeOrderList, List.insertionSort, timeOrderList] simp only [List.map_cons, List.cons.injEq] at h exact orderedInsert_crAnTimeOrderRel_crAnSection h.1 โŸจ(List.insertionSort crAnTimeOrderRel ฯˆs), crAnTimeOrderList_crAnSection_is_crAnSection โŸจฯˆs, h.2โŸฉโŸฉ /-- Time ordering of sections of a list of states. -/ def crAnSectionTimeOrder (ฯ†s : List ๐“•.States) (ฯˆs : CrAnSection ฯ†s) : CrAnSection (timeOrderList ฯ†s) := โŸจcrAnTimeOrderList ฯˆs.1, crAnTimeOrderList_crAnSection_is_crAnSection ฯˆsโŸฉ lemma orderedInsert_crAnTimeOrderRel_injective {ฯˆ ฯˆ' : ๐“•.CrAnStates} (h : ฯˆ.1 = ฯˆ'.1) : {ฯ†s : List ๐“•.States} โ†’ (ฯˆs ฯˆs' : ๐“•.CrAnSection ฯ†s) โ†’ (ho : List.orderedInsert crAnTimeOrderRel ฯˆ ฯˆs.1 = List.orderedInsert crAnTimeOrderRel ฯˆ' ฯˆs'.1) โ†’ ฯˆ = ฯˆ' โˆง ฯˆs = ฯˆs' | [], โŸจ[], _โŸฉ, โŸจ[], _โŸฉ, h => by simp only [List.orderedInsert, List.cons.injEq, and_true] at h simpa using h | ฯ† :: ฯ†s, โŸจฯˆ1 :: ฯˆs, h1โŸฉ, โŸจฯˆ1' :: ฯˆs', h1'โŸฉ, ho => by simp only [List.map_cons, List.cons.injEq] at h1 h1' have ih := orderedInsert_crAnTimeOrderRel_injective h โŸจฯˆs, h1.2โŸฉ โŸจฯˆs', h1'.2โŸฉ simp only [List.orderedInsert] at ho by_cases hr : crAnTimeOrderRel ฯˆ ฯˆ1 ยท simp_all only [ite_true] by_cases hr2 : crAnTimeOrderRel ฯˆ' ฯˆ1' ยท simp_all ยท simp only [crAnTimeOrderRel] at hr hr2 simp_all only rw [crAnStatesToStates] at h1 h1' rw [h1.1] at hr rw [h1'.1] at hr2 exact False.elim (hr2 hr) ยท simp_all only [ite_false] by_cases hr2 : crAnTimeOrderRel ฯˆ' ฯˆ1' ยท simp only [crAnTimeOrderRel] at hr hr2 simp_all only rw [crAnStatesToStates] at h1 h1' rw [h1.1] at hr rw [h1'.1] at hr2 exact False.elim (hr hr2) ยท simp only [hr2, โ†“reduceIte, List.cons.injEq] at ho have ih' := ih ho.2 simp_all only [and_self, implies_true, not_false_eq_true, true_and] apply Subtype.ext simp only [List.cons.injEq, true_and] rw [Subtype.eq_iff] at ih' exact ih'.2 lemma crAnSectionTimeOrder_injective : {ฯ†s : List ๐“•.States} โ†’ Function.Injective (๐“•.crAnSectionTimeOrder ฯ†s) | [], โŸจ[], _โŸฉ, โŸจ[], _โŸฉ => by simp | ฯ† :: ฯ†s, โŸจฯˆ :: ฯˆs, hโŸฉ, โŸจฯˆ' :: ฯˆs', h'โŸฉ => by intro h1 apply Subtype.ext simp only [List.cons.injEq] simp only [crAnSectionTimeOrder] at h1 rw [Subtype.eq_iff] at h1 simp only [crAnTimeOrderList, List.insertionSort] at h1 simp only [List.map_cons, List.cons.injEq] at h h' rw [crAnStatesToStates] at h h' have hin := orderedInsert_crAnTimeOrderRel_injective (by rw [h.1, h'.1]) (๐“•.crAnSectionTimeOrder ฯ†s โŸจฯˆs, h.2โŸฉ) (๐“•.crAnSectionTimeOrder ฯ†s โŸจฯˆs', h'.2โŸฉ) h1 apply And.intro hin.1 have hl := crAnSectionTimeOrder_injective hin.2 rw [Subtype.ext_iff] at hl simpa using hl lemma crAnSectionTimeOrder_bijective (ฯ†s : List ๐“•.States) : Function.Bijective (๐“•.crAnSectionTimeOrder ฯ†s) := by rw [Fintype.bijective_iff_injective_and_card] apply And.intro crAnSectionTimeOrder_injective apply CrAnSection.card_perm_eq simp only [timeOrderList] exact List.Perm.symm (List.perm_insertionSort timeOrderRel ฯ†s) lemma sum_crAnSections_timeOrder {ฯ†s : List ๐“•.States} [AddCommMonoid M] (f : CrAnSection (timeOrderList ฯ†s) โ†’ M) : โˆ‘ s, f s = โˆ‘ s, f (๐“•.crAnSectionTimeOrder ฯ†s s) := by erw [(Equiv.ofBijective _ (๐“•.crAnSectionTimeOrder_bijective ฯ†s)).sum_comp] /-! ## normTimeOrderRel -/ /-- The time ordering relation on `CrAnStates` such that if two CrAnStates have the same time, we normal order them. -/ def normTimeOrderRel (a b : ๐“•.CrAnStates) : Prop := crAnTimeOrderRel a b โˆง (crAnTimeOrderRel b a โ†’ normalOrderRel a b) /-- The relation `normTimeOrderRel` is decidable, but not computablly so due to `Real.decidableLE`. -/ noncomputable instance (ฯ† ฯ†' : ๐“•.CrAnStates) : Decidable (normTimeOrderRel ฯ† ฯ†') := instDecidableAnd /-- Norm-Time ordering of `CrAnStates` is total. -/ instance : IsTotal ๐“•.CrAnStates ๐“•.normTimeOrderRel where total a b := by simp only [normTimeOrderRel] match IsTotal.total (r := ๐“•.crAnTimeOrderRel) a b, IsTotal.total (r := ๐“•.normalOrderRel) a b with | Or.inl h1, Or.inl h2 => simp [h1, h2] | Or.inr h1, Or.inl h2 => simp only [h1, h2, imp_self, and_true, true_and] by_cases hn : crAnTimeOrderRel a b ยท simp [hn] ยท simp [hn] | Or.inl h1, Or.inr h2 => simp only [h1, true_and, h2, imp_self, and_true] by_cases hn : crAnTimeOrderRel b a ยท simp [hn] ยท simp [hn] | Or.inr h1, Or.inr h2 => simp [h1, h2] /-- Norm-Time ordering of `CrAnStates` is transitive. -/ instance : IsTrans ๐“•.CrAnStates ๐“•.normTimeOrderRel where trans a b c := by intro h1 h2 simp_all only [normTimeOrderRel] apply And.intro ยท exact IsTrans.trans _ _ _ h1.1 h2.1 ยท intro hc refine IsTrans.trans _ _ _ (h1.2 ?_) (h2.2 ?_) ยท exact IsTrans.trans _ _ _ h2.1 hc ยท exact IsTrans.trans _ _ _ hc h1.1 /-- The sign associated with putting a list of `CrAnStates` into normal-time order (with the state of greatest time to the left). We pick up a minus sign for every fermion paired crossed. -/ def normTimeOrderSign (ฯ†s : List ๐“•.CrAnStates) : โ„‚ := Wick.koszulSign ๐“•.crAnStatistics ๐“•.normTimeOrderRel ฯ†s /-- Sort a list of `CrAnStates` based on `normTimeOrderRel`. -/ def normTimeOrderList (ฯ†s : List ๐“•.CrAnStates) : List ๐“•.CrAnStates := List.insertionSort ๐“•.normTimeOrderRel ฯ†s end end FieldSpecification