Theory Semi_Mojmir

(*
    Author:      Salomon Sickert
    License:     BSD
*)

section ‹Mojmir Automata (Without Final States)›

theory Semi_Mojmir
  imports Main "Auxiliary/Preliminaries2" DTS
begin

subsection ‹Definitions›

locale semi_mojmir_def =
  fixes
    ― ‹Alphapet›
    Σ :: "'a set"
  fixes
    ― ‹Transition Function›
    δ :: "('b, 'a) DTS"
  fixes
    ― ‹Initial State›
    q0 :: "'b"
  fixes
    ― ‹$\omega$-Word›
    w :: "'a word"
begin

definition sink :: "'b  bool"
where
  "sink q  (q0  q)  (ν  Σ. δ q ν = q)"

declare sink_def [code]

fun token_run :: "nat  nat  'b"
where
  "token_run x n = run δ q0 (suffix x w) (n - x)"

fun configuration :: "'b  nat  nat set"
where
  "configuration q n = {x. x  n  token_run x n = q}"

fun oldest_token :: "'b  nat  nat option"
where
  "oldest_token q n = (if configuration q n  {} then Some (Min (configuration q n)) else None)"

fun senior :: "nat  nat  nat"
where
  "senior x n = the (oldest_token (token_run x n) n)"

fun older_seniors :: "nat  nat  nat set"
where
  "older_seniors x n = {s. y. s = senior y n  s < senior x n  ¬ sink (token_run s n)}"

fun rank :: "nat  nat  nat option"
where
  "rank x n =
    (if x  n  ¬sink (token_run x n) then Some (card (older_seniors x n)) else None)"

fun senior_states :: "'b  nat  'b set"
where
  "senior_states q n =
    {p. x y. oldest_token p n = Some y  oldest_token q n = Some x  y < x  ¬ sink p}"

fun state_rank :: "'b  nat  nat option"
where
  "state_rank q n = (if configuration q n  {}  ¬sink q then Some (card (senior_states q n)) else None)"

definition max_rank :: "nat"
where
  "max_rank = card (reach Σ δ q0 - {q. sink q})"

subsubsection ‹Iterative Computation of State-Ranks›

fun initial :: "'b  nat option"
where
  "initial q = (if q = q0 then Some 0 else None)"

fun pre_ranks :: "('b  nat option)  'a  'b  nat set"
where
  "pre_ranks r ν q = {i . q'. r q' = Some i  q = δ q' ν}  (if q = q0 then {max_rank} else {})"

fun step :: "('b  nat option)  'a  ('b  nat option)"
where
  "step r ν q = (
    if
      ¬sink q  pre_ranks r ν q  {}
    then
      Some (card {q'. ¬sink q'  pre_ranks r ν q'  {}  Min (pre_ranks r ν q') < Min (pre_ranks r ν q)})
    else
      None)"

subsubsection ‹Properties of Tokens›

definition token_squats :: "nat  bool"
where
  "token_squats x = (n. ¬sink (token_run x n))"

end

locale semi_mojmir = semi_mojmir_def +
  assumes
    ― ‹The alphabet is finite. Non-emptiness is derived from well-formed w›
    finite_Σ: "finite Σ"
  assumes
    ― ‹The set of reachable states is finite›
    finite_reach: "finite (reach Σ δ q0)"
  assumes
    ― ‹w only contains letters from the alphabet›
    bounded_w: "range w  Σ"
begin

lemma nonempty_Σ: "Σ  {}"
  using bounded_w by blast

lemma bounded_w': "w i  Σ"
  using bounded_w by blast

― ‹Naming Scheme:

This theory uses the following naming scheme to consistently name variables.

* Tokens: x, y, z
* Time: n, m
* Rank: i, j, k
* States: p, q›

lemma sink_rev_step:
  "¬sink q  q = δ q' ν  ν  Σ  ¬sink q'"
  "¬sink q  q = δ q' (w i)  ¬sink q'"
  using bounded_w' by (force simp only: sink_def)+

subsection ‹Token Run›

lemma token_stays_in_sink:
  assumes "sink q"
  assumes "token_run x n = q"
  shows "token_run x (n + m) = q"
proof (cases "x  n")
  case True
    show ?thesis
    proof (induction m)
      case 0
        show ?case
          using assms(2) by simp
    next
      case (Suc m)
        have "x  n + m"
          using True by simp
        moreover
        have "x. w x  Σ"
          using bounded_w by auto
        ultimately
        have "t. token_run x (n + m)  = q  token_run x (n + m + 1) = q"
          using sink q[unfolded sink_def] upt_add_eq_append[OF le0, of "n + m" 1]
          using Suc_diff_le by simp
        with Suc show ?case
          by simp
    qed
qed (insert assms, simp add: sink_def)

lemma token_is_not_in_sink:
  "token_run x n  A  token_run x (Suc n)  A  ¬sink (token_run x n)"
  by (metis Suc_eq_plus1 token_stays_in_sink)

lemma token_run_intial_state:
  "token_run x x = q0"
  by simp

lemma token_run_P:
  assumes "¬ P (token_run x n)"
  assumes "P (token_run x (Suc (n + m)))"
  shows "m'  m. ¬ P (token_run x (n + m'))  P (token_run x (Suc (n + m')))"
  using assms by (induction m) (simp_all, metis add_Suc_right le_Suc_eq)

lemma token_run_merge_Suc:
  assumes "x  n"
  assumes "y  n"
  assumes "token_run x n = token_run y n"
  shows "token_run x (Suc n) = token_run y (Suc n)"
proof -
  have "run δ q0 (suffix x w) (Suc (n - x)) = run δ q0 (suffix y w) (Suc (n - y))"
    using assms by fastforce
  thus ?thesis
    using Suc_diff_le assms(1,2) by force
qed

lemma token_run_merge:
  "x  n; y  n; token_run x n = token_run y n  token_run x (n + m) = token_run y (n + m)"
  using token_run_merge_Suc[of x _ y] by (induction m) auto

lemma token_run_mergepoint:
  assumes "x < y"
  assumes "token_run x (y + n) = token_run y (y + n)"
  obtains m where "x  (Suc m)" and "y  (Suc m)"
    and "y = Suc m  token_run x m  token_run y m"
    and "token_run x (Suc m) = token_run y (Suc m)"
  using assms by (induction n)
    ((metis add_0_iff le_Suc_eq le_add1 less_imp_Suc_add),
     (metis add_Suc_right le_add1 less_or_eq_imp_le order_trans))

subsubsection ‹Step Lemmas›

lemma token_run_step:
  assumes "x  n"
  assumes "token_run x n = q'"
  assumes "q = δ q' (w n)"
  shows "token_run x (Suc n) = q"
  using assms unfolding token_run.simps Suc_diff_le[OF x  n] by force

lemma token_run_step':
  "x  n  token_run x (Suc n) = δ (token_run x n) (w n)"
  using token_run_step by simp

subsection ‹Configuration›

subsubsection ‹Properties›

lemma configuration_distinct:
  "q  q'  configuration q n  configuration q' n = {}"
  by auto

lemma configuration_finite:
  "finite (configuration q n)"
  by simp

lemma configuration_non_empty:
  "x  n  configuration (token_run x n) n  {}"
  by fastforce

lemma configuration_token:
  "x  n  x  configuration (token_run x n) n"
  by fastforce

lemmas configuration_Max_in = Max_in[OF configuration_finite]
lemmas configuration_Min_in = Min_in[OF configuration_finite]

subsubsection ‹Monotonicity›

lemma configuration_monotonic_Suc:
  "x  n  configuration (token_run x n) n  configuration (token_run x (Suc n)) (Suc n)"
proof
  fix y
  assume "y  configuration (token_run x n) n"
  hence "y  n" and "token_run x n = token_run y n"
    by simp_all
  moreover
  assume "x  n"
  ultimately
  have "token_run x (Suc n) = token_run y (Suc n)"
    using token_run_merge_Suc by blast
  thus "y  configuration (token_run x (Suc n)) (Suc n)"
    using configuration_token y  n by simp
qed

subsubsection ‹Pull-Up and Push-Down›

lemma pull_up_token_run_tokens:
  "x  n; y  n; token_run x n = token_run y n  q. x  configuration q n  y  configuration q n"
  by force

lemma push_down_configuration_token_run:
  "x  configuration q n; y  configuration q n  x  n  y  n  token_run x n = token_run y n"
  by simp

subsubsection ‹Step Lemmas›

lemma configuration_step:
  "x  configuration q' n  q = δ q' (w n)  x  configuration q (Suc n)"
  using Suc_diff_le by simp

lemma configuration_step_non_empty:
  "configuration q' n  {}  q = δ q' (w n)  configuration q (Suc n)  {}"
  by (blast dest: configuration_step)

lemma configuration_rev_step':
  assumes "x  Suc n"
  assumes "x  configuration q (Suc n)"
  obtains q' where "q = δ q' (w n)" and "x  configuration q' n"
  using assms Suc_diff_le by force

lemma configuration_rev_step'':
  assumes "x  configuration q0 (Suc n)"
  shows "x = Suc n  (q'. q0 = δ q' (w n)  x  configuration q' n)"
  using assms configuration_rev_step' by metis

lemma configuration_step_eq_q0:
  "configuration q0 (Suc n) = {Suc n}  {configuration q' n | q'. q0 = δ q' (w n)}"
  apply rule using configuration_rev_step'' apply fast using configuration_step[of _ _ n q0] by fastforce

lemma configuration_rev_step:
  assumes "q  q0"
  assumes "x  configuration q (Suc n)"
  obtains q' where "q = δ q' (w n)" and "x  configuration q' n"
  using configuration_rev_step'[OF _ assms(2)] assms by fastforce

lemma configuration_step_eq:
  assumes "q  q0"
  shows "configuration q (Suc n) = {configuration q' n | q'. q = δ q' (w n)}"
  using configuration_rev_step[OF assms, of _ n] configuration_step by auto

lemma configuration_step_eq_unified:
  shows "configuration q (Suc n) = {configuration q' n | q'. q = δ q' (w n)}  (if q = q0 then {Suc n} else {}) "
  using configuration_step_eq configuration_step_eq_q0 by force

subsection ‹Oldest Token›

subsubsection ‹Properties›

lemma oldest_token_always_def:
  "i. i  x  oldest_token (token_run x n) n = Some i"
proof (cases "x  n")
  case False
    let ?q = "token_run x n"
    from False have "n  configuration ?q n" and "configuration ?q n  {}"
      by auto
    then obtain i where "i  n" and "oldest_token ?q n = Some i"
      by (metis Min.coboundedI oldest_token.simps configuration_finite)
    moreover
    hence "i  x"
      using False by linarith
    ultimately
    show ?thesis
      by blast
qed fastforce

lemma oldest_token_bounded:
  "oldest_token q n = Some x  x  n"
  by (metis oldest_token.simps configuration_Min_in option.distinct(1) option.inject push_down_configuration_token_run)

lemma oldest_token_distinct:
  "q  q'  oldest_token q n = Some i  oldest_token q' n = Some j  i  j"
  by (metis configuration_Min_in configuration_distinct disjoint_iff_not_equal option.distinct(1) oldest_token.simps option.sel)

lemma oldest_token_equal:
  "oldest_token q n = Some i  oldest_token q' n = Some i  q = q'"
  using oldest_token_distinct by blast

subsubsection ‹Monotonicity›

lemma oldest_token_monotonic_Suc:
  assumes "x  n"
  assumes "oldest_token (token_run x n) n = Some i"
  assumes "oldest_token (token_run x (Suc n)) (Suc n) = Some j"
  shows "i  j"
proof -
  from assms have "i = Min (configuration (token_run x n) n)"
    and "j = Min (configuration (token_run x (Suc n)) (Suc n))"
    by (metis oldest_token.elims option.discI option.sel)+
  thus ?thesis
    using Min_antimono[OF configuration_monotonic_Suc[OF assms(1)] configuration_non_empty[OF assms(1)] configuration_finite] by blast
qed

subsubsection ‹Pull-Up and Push-Down›

lemma push_down_oldest_token_configuration:
  "oldest_token q n = Some x  x  configuration q n"
  by (metis configuration_Min_in oldest_token.simps option.distinct(2) option.inject)

lemma push_down_oldest_token_token_run:
  "oldest_token q n = Some x  token_run x n = q"
  using push_down_oldest_token_configuration configuration.simps by blast

subsection ‹Senior Token›

subsubsection ‹Properties›

lemma senior_le_token:
  "senior x n  x"
  using oldest_token_always_def[of x n] by fastforce

lemma senior_token_run:
  "senior x n = senior y n  token_run x n = token_run y n"
  by (metis oldest_token_always_def oldest_token_distinct option.sel senior.simps)

text ‹The senior of a token is always in the same state›

lemma senior_same_state:
  "token_run (senior x n) n = token_run x n"
proof -
  have X: "{t. t  n  token_run t n = token_run x n}  {}"
    by (cases "x  n") auto
  show ?thesis
    using Min_in[OF _ X] by force
qed

lemma senior_senior:
  "senior (senior x n) n = senior x n"
  using senior_same_state senior_token_run by blast

subsubsection ‹Monotonicity›

lemma senior_monotonic_Suc:
  "x  n  senior x n  senior x (Suc n)"
  by (metis oldest_token_always_def oldest_token_monotonic_Suc option.sel senior.simps)

subsubsection ‹Pull-Up and Push-Down›

lemma pull_up_configuration_senior:
  "x  configuration q n; y  configuration q n  senior x n = senior y n"
  by force

lemma push_down_senior_tokens:
  "x  n; y  n; senior x n = senior y n  q. x  configuration q n  y  configuration q n"
  using senior_token_run pull_up_token_run_tokens by blast

subsection ‹Set of Older Seniors›

subsubsection ‹Properties›

lemma older_seniors_cases_subseteq [case_names le ge]:
  assumes "older_seniors x n  older_seniors y n  P"
  assumes "older_seniors x n  older_seniors y n  P"
  shows "P" using assms by fastforce

lemma older_seniors_cases_subset [case_names less equal greater]:
  assumes "older_seniors x n  older_seniors y n  P"
  assumes "older_seniors x n = older_seniors y n  P"
  assumes "older_seniors x n  older_seniors y n  P"
  shows "P" using assms older_seniors_cases_subseteq by blast

lemma older_seniors_finite:
  "finite (older_seniors x n)"
  by fastforce

lemma older_seniors_older:
  "y  older_seniors x n  y < x"
  using less_le_trans[OF _ senior_le_token, of y x n] by force

lemma older_seniors_senior_simp:
  "older_seniors (senior x n) n = older_seniors x n"
  unfolding older_seniors.simps senior_senior ..

lemma older_seniors_not_self_referential:
  "senior x n  older_seniors x n"
  by simp

lemma older_seniors_not_self_referential_2:
  "x  older_seniors x n"
  using older_seniors_older older_seniors_not_self_referential less_not_refl by blast

lemma older_seniors_subset:
  "y  older_seniors x n  older_seniors y n  older_seniors x n"
  using older_seniors_not_self_referential_2 by (cases rule: older_seniors_cases_subset) blast+

lemma older_seniors_subset_2:
  assumes "¬ sink (token_run x n)"
  assumes "older_seniors x n  older_seniors y n"
  shows "senior x n  older_seniors y n"
proof -
  have "senior x n < senior y n"
    using assms(2) by fastforce
  thus ?thesis
    using assms(1)[unfolded senior_same_state[symmetric, of x n]]
    unfolding older_seniors.simps by blast
qed

lemmas older_seniors_Max_in = Max_in[OF older_seniors_finite]
lemmas older_seniors_Min_in = Min_in[OF older_seniors_finite]
lemmas older_seniors_Max_coboundedI = Max.coboundedI[OF older_seniors_finite]
lemmas older_seniors_Min_coboundedI = Min.coboundedI[OF older_seniors_finite]
lemmas older_seniors_card_mono = card_mono[OF older_seniors_finite]
lemmas older_seniors_psubset_card_mono = psubset_card_mono[OF older_seniors_finite]

lemma older_seniors_recursive:
  fixes x n
  defines "os  older_seniors x n"
  assumes "os  {}"
  shows "os = {Max os}  older_seniors (Max os) n"
  (is "?lhs = ?rhs")
proof
  show "?lhs  ?rhs"
  proof
    fix x
    assume "x  ?lhs"
    show "x  ?rhs"
    proof (cases "x = Max os")
      case False
        hence "x < Max os"
          by (metis older_seniors_Max_coboundedI os_def x  os dual_order.order_iff_strict)
        moreover
        obtain y' where "Max os = senior y' n"
          using older_seniors_Max_in assms(2)
          unfolding os_def older_seniors.simps by blast
        ultimately
        have "x < senior (Max os) n"
          using senior_senior by presburger
        moreover
        from x  ?lhs obtain y where "x = senior y n" and "¬ sink (token_run x n)"
          unfolding os_def older_seniors.simps by blast
        ultimately
        show ?thesis
          unfolding older_seniors.simps by blast
    qed blast
  qed
next
  show "?lhs  ?rhs"
    using older_seniors_subset older_seniors_Max_in assms(2)
    unfolding os_def by blast
qed

lemma older_seniors_recursive_card:
  fixes x n
  defines "os  older_seniors x n"
  assumes "os  {}"
  shows "card os = Suc (card (older_seniors (Max os) n))"
  by (metis older_seniors_recursive assms Un_empty_left Un_insert_left card_insert_disjoint older_seniors_finite older_seniors_not_self_referential_2)

lemma older_seniors_card:
  "card (older_seniors x n) = card (older_seniors y n)  older_seniors x n = older_seniors y n"
  by (metis less_not_refl older_seniors_cases_subset older_seniors_psubset_card_mono)

lemma older_seniors_card_le:
  "card (older_seniors x n) < card (older_seniors y n)  older_seniors x n  older_seniors y n"
  by (metis card_mono card_psubset not_le older_seniors_cases_subseteq older_seniors_finite psubset_card_mono)

lemma older_seniors_card_less:
  "card (older_seniors x n)  card (older_seniors y n)  older_seniors x n  older_seniors y n"
   by (metis not_le older_seniors_card_mono older_seniors_cases_subseteq older_seniors_psubset_card_mono subset_not_subset_eq)

subsubsection ‹Monotonicity›

lemma older_seniors_monotonic_Suc:
  assumes "x  n"
  shows "older_seniors x n  older_seniors x (Suc n)"
proof
  fix y
  assume "y  older_seniors x (Suc n)"
  then obtain ox where "y = senior ox (Suc n)"
    and "y < senior x (Suc n)"
    and "¬ sink (token_run y (Suc n))"
    unfolding older_seniors.simps by blast

  hence "y = senior y n"
    using senior_senior senior_le_token senior_monotonic_Suc assms
    by (metis add.commute add.left_commute dual_order.order_iff_strict linear not_add_less1 not_less le_iff_add)
  moreover
  have "y < senior x n"
    using assms less_le_trans[OF y < senior x (Suc n) senior_monotonic_Suc]
    by blast
  moreover
  have "¬ sink (token_run y n)"
    using ¬ sink (token_run y (Suc n)) token_stays_in_sink
    unfolding Suc_eq_plus1 by metis

  ultimately
  show "y  older_seniors x n"
    unfolding older_seniors.simps by blast
qed

lemma older_seniors_monotonic:
  "x  n  older_seniors x n  older_seniors x (n + m)"
  by (induction m) (simp, metis older_seniors_monotonic_Suc add_Suc_right dual_order.trans trans_le_add1)

lemma older_seniors_stable:
  "x  n  older_seniors x n = older_seniors x (n + m + m')  older_seniors x n = older_seniors x (n + m)"
  by (induction m') (simp, unfold set_eq_subset, metis dual_order.trans le_add1 older_seniors_monotonic)

lemma card_older_seniors_monotonic:
  "x  n  card (older_seniors x n)  card (older_seniors x (n + m))"
  using older_seniors_monotonic older_seniors_card_mono by meson

subsubsection ‹Pull-Up and Push-Down›

lemma pull_up_senior_older_seniors:
  "senior x n = senior y n  older_seniors x n = older_seniors y n"
  unfolding older_seniors.simps senior.simps senior_token_run by presburger

lemma pull_up_senior_older_seniors_less:
  "senior x n < senior y n  older_seniors x n  older_seniors y n"
  by force

lemma pull_up_senior_older_seniors_less_2:
  assumes "¬ sink (token_run x n)"
  assumes "senior x n < senior y n"
  shows "older_seniors x n  older_seniors y n"
proof -
  from assms have "senior x n  older_seniors y n"
    unfolding senior_same_state[of x n, symmetric] older_seniors.simps by blast
  thus ?thesis
    using older_seniors_not_self_referential pull_up_senior_older_seniors_less[OF assms(2)] by blast
qed

lemma pull_up_senior_older_seniors_le:
  "senior x n  senior y n  older_seniors x n  older_seniors y n"
  using pull_up_senior_older_seniors pull_up_senior_older_seniors_less
  unfolding dual_order.order_iff_strict by blast

lemma push_down_older_seniors_senior:
  assumes "¬ sink (token_run x n)"
  assumes "¬ sink (token_run y n)"
  assumes "older_seniors x n = older_seniors y n"
  shows "senior x n = senior y n"
  using assms by (cases "senior x n" " senior y n" rule: linorder_cases) (fast dest: pull_up_senior_older_seniors_less_2)+

subsubsection ‹Tower Lemma›

lemma older_seniors_tower'':
  assumes "x  n"
  assumes "y  n"
  assumes "¬sink (token_run x n)"
  assumes "¬sink (token_run y n)"
  assumes "older_seniors x n = older_seniors x (Suc n)"
  assumes "older_seniors y n  older_seniors x n"
  shows "older_seniors y n = older_seniors y (Suc n)"
proof
  {
    fix s
    assume "s  older_seniors y n" and "older_seniors y n  older_seniors x n"
    hence "s  older_seniors x n"
      using assms by blast
    hence "¬sink (token_run s (Suc n))" and "z. s = senior z (Suc n)"
      unfolding assms by simp+
    moreover
    have "senior y n  senior y (Suc n)"
    proof (rule ccontr)
      assume "¬senior y n  senior y (Suc n)"
      moreover
      have "senior y n  n"
        by (metis assms(2) senior_le_token le_trans)
      ultimately
      have "z. senior y n  senior z (Suc n)"
        using token_run_merge_Suc[unfolded senior_token_run[symmetric], OF y  n]
        by (metis senior_senior le_refl)
      hence "senior y n  older_seniors x (Suc n)"
        using assms by simp
      moreover
      have "senior y n  older_seniors x n"
        using assms older_seniors y n  older_seniors x n older_seniors_subset_2 by meson
      ultimately
      show False
        unfolding assms ..
    qed
    hence "s < senior y (Suc n)"
      using s  older_seniors y n by fastforce
    ultimately
    have "s  older_seniors y (Suc n)"
      unfolding older_seniors.simps by blast
  }
  moreover
  {
    fix s
    assume "s  older_seniors y n" and "older_seniors y n = older_seniors x n"
    moreover
    hence "senior y n = senior x n"
      using assms(3-4) push_down_older_seniors_senior by blast
    hence "senior y (Suc n) = senior x (Suc n)"
      using token_run_merge_Suc[OF assms(2,1)] unfolding senior_token_run by blast
    ultimately
    have "s  older_seniors y (Suc n)"
      by (metis assms(5) older_seniors_senior_simp)
  }
  ultimately
  show "older_seniors y n  older_seniors y (Suc n)"
    using assms by blast
qed (metis older_seniors_monotonic_Suc assms(2))

lemma older_seniors_tower''2:
  assumes "x  n"
  assumes "y  n"
  assumes "¬sink (token_run x (n + m))"
  assumes "¬sink (token_run y (n + m))"
  assumes "older_seniors x n = older_seniors x (n + m)"
  assumes "older_seniors y n  older_seniors x n"
  shows "older_seniors y n = older_seniors y (n + m)"
  using assms
proof (induction m arbitrary: n)
  case (Suc m)
    have "¬sink (token_run x (n + m))" and "¬sink (token_run y (n + m))"
      using ¬sink (token_run x (n + Suc m)) ¬sink (token_run y (n + Suc m))
      using token_stays_in_sink[of _ _ "n + m" 1]
      unfolding Suc_eq_plus1 add.assoc[symmetric] by metis+
    moreover
    have "older_seniors x n = older_seniors x (n + m)"
      using Suc.prems(5) older_seniors_stable[OF x  n]
      unfolding Suc_eq_plus1 add.assoc by blast
    moreover
    hence "older_seniors x (n + m) = older_seniors x (Suc (n + m))"
      unfolding Suc.prems add_Suc_right ..
    ultimately
    have "older_seniors y n = older_seniors y (n + m)"
      using Suc by meson
    also
    have " = older_seniors y (Suc (n + m))"
      using older_seniors_tower''[OF _ _ ¬sink (token_run x (n + m)) ¬sink (token_run y (n + m)) older_seniors x (n + m) = older_seniors x (Suc (n + m))] Suc
      by (metis older_seniors x n = older_seniors x (n + m) add.commute add.left_commute calculation le_iff_add)
    finally
    show ?case
      unfolding add_Suc_right .
qed simp

lemma older_seniors_tower':
  assumes "y  older_seniors x n"
  assumes "older_seniors x n = older_seniors x (Suc n)"
  shows "older_seniors y n = older_seniors y (Suc n)"
  (is "?lhs = ?rhs")
  using assms
proof (induction "card (older_seniors x n)" arbitrary: x y)
  case 0
    hence "older_seniors x n = {}"
      using older_seniors_finite card_eq_0_iff by metis
    thus ?case
      using "0.prems" by blast
next
  case (Suc c)
    let ?os = "older_seniors x n"
    have "?os  {}"
      using Suc.prems(1) by blast

    hence "y = Max ?os  y  older_seniors (Max ?os) n"
      using Suc.prems(1) older_seniors_recursive by blast
    moreover
    have "older_seniors (Max ?os) n = older_seniors (Max ?os) (Suc n)"
      using Suc.prems(2) older_seniors_recursive ?os  {} older_seniors_not_self_referential_2
      by (metis Un_empty_left Un_insert_left insert_ident)
    moreover
    {
      fix s
      assume "s  older_seniors (Max ?os) n"
      moreover
      from Suc.hyps(2) have "card (older_seniors (Max ?os) n) = c"
        unfolding older_seniors_recursive_card[OF ?os  {}] by blast
      ultimately
      have "older_seniors s n = older_seniors s (Suc n)"
        by (metis Suc.hyps(1) older_seniors (Max ?os) n = older_seniors (Max ?os) (Suc n))
    }
    ultimately
    show ?case
      by blast
qed

lemma older_seniors_tower:
  "x  n; y  older_seniors x n; older_seniors x n = older_seniors x (n + m)  older_seniors y n = older_seniors y (n + m)"
proof (induction m)
  case (Suc m)
    hence "older_seniors x n = older_seniors x (n + m)"
      using older_seniors_monotonic older_seniors_monotonic_Suc subset_antisym
      by (metis Nat.add_0_right add.assoc add_Suc_shift trans_le_add1)
    hence "older_seniors y n = older_seniors y (n + m)"
      using Suc.IH[OF Suc.prems(1,2)] by blast
    also
    have " = older_seniors y (n + Suc m)"
      using older_seniors_tower'[of y x "n + m"] Suc.prems unfolding add_Suc_right
      by (metis older_seniors x n = older_seniors x (n + m))
    finally
    show ?case .
qed simp

subsection ‹Rank›

subsubsection ‹Properties›

lemma rank_None_before:
  "x > n  rank x n = None"
  by simp

lemma rank_None_Suc:
  assumes "x  n"
  assumes "rank x n = None"
  shows "rank x (Suc n) = None"
proof -
  have "sink (token_run x n)"
    using assms by (metis option.distinct(1) rank.simps)
  hence "sink (token_run x (Suc n))"
    using token_stays_in_sink by (metis (erased, opaque_lifting) Suc_leD le_Suc_ex not_less_eq_eq)
  thus ?thesis
    by simp
qed

lemma rank_Some_time:
  "rank x n = Some j  x  n"
  by (metis option.distinct(1) rank.simps)

lemma rank_Some_sink:
  "rank x n = Some j  ¬sink (token_run x n)"
  by fastforce

lemma rank_Some_card:
  "rank x n = Some j  card (older_seniors x n) = j"
  by (metis option.distinct(1) option.inject rank.simps)

lemma rank_initial:
  "i. rank x x = Some i"
  unfolding rank.simps sink_def by force

lemma rank_continuous:
  assumes "rank x n = Some i"
  assumes "rank x (n + m) = Some j"
  assumes "m'  m"
  shows "k. rank x (n + m') = Some k"
  using assms
proof (induction m arbitrary: j m')
  case (Suc m)
    thus ?case
    proof (cases "m' = Suc m")
      case False
        with Suc.prems have "m'  m"
          by linarith
        moreover
        obtain j' where "rank x (n + m) = Some j'"
          using Suc.prems(1,2) rank_Some_time rank_None_Suc
          by (metis add_Suc_right add_lessD1 not_less rank.simps)
        ultimately
        show ?thesis
          using Suc.IH[OF Suc.prems(1)] by blast
    qed simp
qed simp

lemma rank_token_squats:
  "token_squats x  x  n  i. rank x n = Some i"
  unfolding token_squats_def by simp

lemma rank_older_seniors_bounded:
  assumes "y  older_seniors x n"
  assumes "rank x n = Some j"
  shows "j' < j. rank y n = Some j'"
proof -
  from assms(1) have "¬sink (token_run y n)"
    by simp
  moreover
  from assms have "y  n"
    by (metis dual_order.trans linear not_less older_seniors_older option.distinct(1) rank.simps)
  moreover
  have "older_seniors y n  older_seniors x n"
    using older_seniors_subset assms(1) by presburger
  hence "card (older_seniors y n) < card (older_seniors x n)"
    by (rule older_seniors_psubset_card_mono)
  ultimately
  show ?thesis
    using rank_Some_card[OF assms(2)] rank.simps by meson
qed

subsubsection ‹Bounds›

lemma max_rank_lowerbound:
  "0 < max_rank"
proof -
  obtain a where "a  Σ"
    using nonempty_Σ by blast
  hence "range (λ_. a)  Σ" and "q0 = run δ q0 (λ_. a) 0"
    by auto
  hence "q0  reach Σ δ q0"
    unfolding reach_def by blast
  thus ?thesis
    using reach_card_0[OF nonempty_Σ] finite_reach max_rank_def sink_def by force
qed

lemma older_seniors_card_bounded:
  assumes "¬sink (token_run x n)" and "x  n"
  shows "card (older_seniors x n) < card (reach Σ δ q0 - {q. sink q})"
  (is "card ?S4 < card ?S0")
proof -
  let ?S1 = "{token_run x n | x n. True} - {q. sink q}"
  let ?S2 = "(λq. the (oldest_token q n)) ` ?S1"
  let ?S3 = "{s. x. s = senior x n  ¬(sink (token_run s n))}"

  have "?S1  ?S0"
    unfolding reach_def token_run.simps using bounded_w by fastforce
  hence "finite ?S1" and C1: "card ?S1  card ?S0"
    using finite_reach card_mono finite_subset
   apply (simp add: finite_subset) by (metis {token_run x n |x n. True} - Collect sink  reach Σ δ q0 - Collect sink card_mono finite_Diff local.finite_reach)
  hence "finite ?S2" and C2: "card ?S2  card ?S1"
    using finite_imageI card_image_le by blast+
  moreover
  have "?S3  ?S2"
  proof
    fix s
    assume "s  ?S3"
    hence "s = senior s n" and "¬sink (token_run s n)"
      using senior_senior by fastforce+
    thus "s  ?S2"
      by auto
  qed
  ultimately
  have "finite ?S3" and C3: "card ?S3  card ?S2"
    using card_mono finite_subset by blast+
  moreover
  have "senior x n  ?S3" and "senior x n  ?S4" and "?S4  ?S3"
    using assms older_seniors_not_self_referential senior_same_state by auto
  hence "?S4  ?S3"
    by blast
  ultimately
  have "finite ?S4" and C4: "card ?S4 < card ?S3"
    using psubset_card_mono finite_subset by blast+
  show ?thesis
    using C1 C2 C3 C4 by linarith
qed

lemma rank_upper_bound:
  "rank x n = Some i  i < max_rank"
  using older_seniors_card_bounded unfolding max_rank_def
  by (fast dest: rank_Some_card rank_Some_time rank_Some_sink )

lemma rank_range:
  "i. range (rank x)  {None}  Some ` {0..<i}"
proof
  {
    fix i_option
    assume "i_option  range (rank x)"
    hence "i_option  {None}  Some ` {0..<max_rank}"
    proof (cases i_option)
      case (Some i)
        hence "i  {0..<max_rank}"
          using i_option  range (rank x) rank_upper_bound by force
        thus ?thesis
          using Some by blast
    qed blast
  }
  thus "range (rank x)  ({None}  Some ` {0..<max_rank})" ..
qed

subsubsection ‹Monotonicity›

lemma rank_monotonic:
  "rank x n = Some i; rank x (n + m) = Some j  i  j"
  using card_older_seniors_monotonic rank_Some_card rank_Some_time by metis

subsubsection ‹Pull-Up and Push-Down›

lemma pull_up_senior_rank:
  "x  n; y  n; senior x n = senior y n  rank x n = rank y n"
  by (metis senior_token_run rank.simps pull_up_senior_older_seniors)

lemma pull_up_configuration_rank:
  "x  configuration q n; y  configuration q n  rank x n = rank y n"
  by force

lemma push_down_rank_older_seniors:
  "rank x n = rank y n; rank x n = Some i  older_seniors x n = older_seniors y n"
  by (metis older_seniors_card option.distinct(2) option.sel rank.simps)

lemma push_down_rank_senior:
  "rank x n = rank y n; rank x n = Some i  senior x n = senior y n"
  by (metis push_down_rank_older_seniors push_down_older_seniors_senior option.distinct(1) rank.elims)

lemma push_down_rank_tokens:
  "rank x n = rank y n; rank x n = Some i  (q. x  configuration q n  y  configuration q n)"
  by (metis push_down_senior_tokens rank_Some_time push_down_rank_senior)

subsubsection ‹Pulled-Up Lemmas›

lemma rank_senior_senior:
  "x  n  rank (senior x n) n = rank x n"
  by (metis le_iff_add add.commute add.left_commute pull_up_senior_rank senior_le_token senior_senior)

subsubsection ‹Stable Rank›

definition stable_rank :: "nat  nat  bool"
where
  "stable_rank x i = (n. rank x n = Some i)"

lemma stable_rank_unique:
  assumes "stable_rank x i"
  assumes "stable_rank x j"
  shows "i = j"
proof -
  from assms obtain n m where "n'. n'  n  rank x n' = Some i"
    and "m'. m'  m  rank x m' = Some j"
    unfolding stable_rank_def MOST_nat_le by blast
  hence "rank x (n + m) = Some i" and "rank x (n + m) = Some j"
    by (metis add.commute le_add1)+
  thus ?thesis
    by simp
qed

lemma stable_rank_equiv_token_squats:
  "token_squats x = (i. stable_rank x i)"
  (is "?lhs = ?rhs")
proof
  assume ?lhs
  define ranks where "ranks = {j | j n. rank x n = Some j}"
  hence "ranks  {0..<max_rank}" and "the (rank x x)  ranks"
    using rank_upper_bound rank_initial[of x] unfolding ranks_def by fastforce+ (* Takes roughly 10s *)
  hence "finite ranks" and "ranks  {}"
    using finite_reach finite_atLeastAtMost infinite_super by fast+

  define i where "i = Min ranks"
  obtain n where "rank x n = Some i"
    using Min_in[OF finite ranks ranks  {}]
    unfolding i_def ranks_def by blast

  have "j. j  ranks  j  i"
    using Min_in[OF finite ranks ranks  {}] unfolding i_def
    by (metis Min.coboundedI finite ranks)
  hence "m j. rank x (n + m) = Some j  j  i"
    unfolding ranks_def by blast
  moreover
  have "m j. rank x (n + m) = Some j  j  i"
    using rank_monotonic[OF rank x n = Some i] by blast
  moreover
  have "m. j. rank x (n + m) = Some j"
    using rank_token_squats[OF ?lhs] rank_Some_time[OF rank x n = Some i] by simp
  ultimately
  have "m. rank x (n + m) = Some i"
    by (metis le_antisym)
  thus ?rhs
    unfolding stable_rank_def MOST_nat_le by (metis le_iff_add)
next
  assume ?rhs
  thus ?lhs
    unfolding token_squats_def stable_rank_def MOST_nat_le
    by (metis le_add2 rank_Some_sink token_stays_in_sink)
qed

lemma stable_rank_same_tokens:
  assumes "stable_rank x i"
  assumes "stable_rank y j"
  assumes "x  configuration q n"
  assumes "y  configuration q n"
  shows "i = j"
proof -
  from assms(1) obtain n_i where "n_i  n" and "t  n_i. rank x t = Some i"
    unfolding stable_rank_def MOST_nat_le by (metis linear order_trans)
  moreover
  from assms(2) obtain n_j where "n_j  n" and "t  n_j. rank y t = Some j"
    unfolding stable_rank_def MOST_nat_le by (metis linear order_trans)
  moreover
  define m where "m = max n_i n_j"
  ultimately
  have "rank x m = Some i" and "rank y m = Some j"
    by (metis max.bounded_iff order_refl)+
  moreover
  have "m  n"
    by (metis n  n_j le_trans max.cobounded2 m_def)
  have "q'. x  configuration q' m  y  configuration q' m"
    using push_down_configuration_token_run[OF assms(3,4)]
    using token_run_merge[of x n y]
    using pull_up_token_run_tokens[of x m y]
    using m  n[unfolded le_iff_add] by force
  ultimately
  show ?thesis
    using pull_up_configuration_rank by (metis option.inject)
qed

subsubsection ‹Tower Lemma›

lemma rank_tower:
  assumes "i  j"
  assumes "rank x n = Some j"
  assumes "rank x (n + m) = Some j"
  assumes "rank y n = Some i"
  shows "rank y (n + m) = Some i"
proof (cases i j rule: linorder_cases)
  case less
    {
      hence "card (older_seniors (senior y n) n) < card (older_seniors x n)"
        using assms rank_Some_card senior_same_state by force
      hence "senior y n  older_seniors x n"
        by (metis older_seniors_card_le rank_Some_sink assms(4) older_seniors_senior_simp older_seniors_subset_2)
      moreover
      have "older_seniors x n = older_seniors x (n + m)"
        by (metis assms(2,3) rank_Some_card rank_Some_time card_subset_eq[OF older_seniors_finite] older_seniors_monotonic)
      ultimately
      have "older_seniors (senior y n) n = older_seniors (senior y n) (n + m)" and "senior y n  older_seniors x (n + m)"
        using older_seniors_tower rank_Some_time assms(2) by blast+
    }
    moreover
    have "rank (senior y n) n = Some i"
      by (metis assms(4) rank_Some_time rank_senior_senior)
    ultimately
    have "rank (senior y n) (n + m) = Some i"
      by (metis rank_older_seniors_bounded[OF _ assms(3)] rank_Some_card)
    moreover
    have "senior y n  n"
      by (metis rank (senior y n) n = Some i rank_Some_time)
    hence "senior y n  configuration (token_run y (n + m)) (n + m)"
      by (metis (full_types) token_run_merge[OF _ rank_Some_time[OF assms(4)] senior_same_state] configuration_token trans_le_add1)
    ultimately
    show ?thesis
      by (metis pull_up_configuration_rank le_iff_add add.assoc assms(4) configuration_token rank_Some_time)
next
  case equal
    hence "x  n" and "y  n" and "token_run x n = token_run y n"
      using assms(2-4) push_down_rank_tokens by force+
    moreover
    hence "token_run x (n + m) = token_run y (n + m)"
      using token_run_merge by blast
    ultimately
    show ?thesis
      by (metis assms(3) equal rank_senior_senior senior_token_run le_iff_add add.assoc)
qed (insert i  j, linarith)

lemma stable_rank_alt_def:
  "rank x n = Some j  stable_rank x j  (m  n. rank x m = Some j)"
  (is "?rhs  ?lhs")
proof
  assume ?rhs
  then obtain m' where "m  m'. rank x m = Some j"
    unfolding stable_rank_def MOST_nat_le by blast
  moreover
  hence "rank x n = Some j" and "rank x m' = Some j"
    using ?rhs by blast+
  {
    fix m
    assume "n  n + m" and "n + m < m'"
    then obtain j' where "rank x (n + m) = Some j'"
      by (metis ?rhs stable_rank_equiv_token_squats rank_Some_time rank_token_squats trans_le_add1)
    moreover
    hence "j'  j"
      using rank x n = Some j rank_monotonic by blast
    moreover
    have "j  j'"
      using rank x (n + m) = Some j' rank x m' = Some j  n + m < m' rank_monotonic
      by (metis add_Suc_right less_imp_Suc_add)
    ultimately
    have "rank x (n + m) = Some j"
      by simp
  }
  ultimately
  show ?lhs
    by (metis le_add_diff_inverse not_le)
qed (unfold stable_rank_def MOST_nat_le, blast)

lemma stable_rank_tower:
  assumes "j  i"
  assumes "rank x n = Some j"
  assumes "rank y n = Some i"
  assumes "stable_rank y i"
  shows "stable_rank x j"
  using assms rank_tower[OF j  i] stable_rank_alt_def[of y n i]
  unfolding stable_rank_def[of x j, unfolded MOST_nat_le] by (metis le_Suc_ex)

subsection ‹Senior States›

lemma senior_states_initial:
  "senior_states q 0 = {}"
  by simp

lemma senior_states_cases_subseteq [case_names le ge]:
  assumes "senior_states p n  senior_states q n  P"
  assumes "senior_states p n  senior_states q n  P"
  shows "P" using assms by force

lemma senior_states_cases_subset [case_names less equal greater]:
  assumes "senior_states p n  senior_states q n  P"
  assumes "senior_states p n = senior_states q n  P"
  assumes "senior_states p n  senior_states q n  P"
  shows "P" using assms senior_states_cases_subseteq by blast

lemma senior_states_finite:
  "finite (senior_states q n)"
  by fastforce

lemmas senior_states_card_mono = card_mono[OF senior_states_finite]
lemmas senior_states_psubset_card_mono = psubset_card_mono[OF senior_states_finite]

lemma senior_states_card:
  "card (senior_states p n) = card (senior_states q n)  senior_states p n = senior_states q n"
  by (metis less_not_refl senior_states_cases_subset senior_states_psubset_card_mono)

lemma senior_states_card_le:
  "card (senior_states p n) < card (senior_states q n)  senior_states p n  senior_states q n"
  by (metis card_mono not_less senior_states_cases_subseteq senior_states_finite senior_states_psubset_card_mono subset_not_subset_eq)

lemma senior_states_card_less:
  "card (senior_states p n)  card (senior_states q n)  senior_states p n  senior_states q n"
  by (metis card_mono card_seteq senior_states_cases_subseteq senior_states_finite)

lemma senior_states_older_seniors:
  "(λy. token_run y n) ` older_seniors x n = senior_states (token_run x n) n"
  (is "?lhs = ?rhs")
proof -
  have "?lhs = {q'. ost ot. q' = token_run ost n  ost = senior ot n  ost < senior x n  ¬ sink q'}"
    by auto
  also
  have " = {q'. t ot. oldest_token q' n = Some t  t = senior ot n  t < senior x n  ¬ sink q'}"
    unfolding senior.simps by (metis (erased, opaque_lifting) oldest_token_always_def push_down_oldest_token_token_run option.sel)
  also
  have " = {q'. t. oldest_token q' n = Some t  t < senior x n  ¬ sink q'}"
    by auto
  also
  have " = ?rhs"
    unfolding senior_states.simps senior.simps by (metis (erased, opaque_lifting) oldest_token_always_def option.sel)
  finally
  show "?lhs = ?rhs"
    .
qed

lemma card_older_senior_senior_states:
  assumes "x  configuration q n"
  shows "card (older_seniors x n) = card (senior_states q n)"
  (is "?lhs = ?rhs")
proof -
  have "inj_on (λt. token_run t n) (older_seniors x n)"
    unfolding inj_on_def using senior_same_state
    by (fastforce simp del: token_run.simps)
  moreover
  have "token_run x n = q"
    using assms by simp
  ultimately
  show "?lhs = ?rhs"
    using card_image[of "(λt. token_run t n)" "older_seniors x n"]
    unfolding senior_states_older_seniors by presburger
qed

subsection ‹Rank of States›

subsubsection ‹Alternative Definitions›

lemma state_rank_eq_rank:
  "state_rank q n = (case oldest_token q n of None  None | Some t  rank t n) "
  (is "?lhs = ?rhs")
proof (cases "oldest_token q n")
  case (None)
    thus ?thesis
      by (metis not_Some_eq oldest_token.elims option.simps(4) state_rank.elims)
next
  case (Some x)
    hence "?lhs = (if ¬sink q then Some (card (older_seniors x n)) else None)"
      by (metis emptyE push_down_oldest_token_configuration[OF Some] card_older_senior_senior_states state_rank.simps)
    also
    have " = rank x n"
      using oldest_token_bounded[OF Some] push_down_oldest_token_token_run[OF Some] by auto
    also
    have " = ?rhs"
      using Some by force
    finally
    show ?thesis .
qed

lemma state_rank_eq_rank_SOME:
  "state_rank q n = (if configuration q n  {} then rank (SOME x. x  configuration q n) n else None)"
proof (cases "oldest_token q n")
  case (Some x)
    thus ?thesis
      unfolding state_rank_eq_rank Some option.simps(5)
      by (metis Some ex_in_conv pull_up_configuration_rank push_down_oldest_token_configuration someI_ex)
qed (unfold state_rank_eq_rank; metis not_Some_eq oldest_token.elims option.simps(4))

lemma rank_eq_state_rank:
  "x  n  rank x n = state_rank (token_run x n) n"
  unfolding state_rank_eq_rank_SOME[of "token_run x n"]
  by (metis all_not_in_conv configuration_token pull_up_configuration_rank someI_ex)

subsubsection ‹Pull-Up and Push-Down›

lemma pull_up_configuration_state_rank:
  "configuration q n = {}  state_rank q n = None"
  by force

lemma push_down_state_rank_tokens:
  "state_rank q n = Some i  configuration q n  {}"
  by (metis not_Some_eq state_rank.elims)

lemma push_down_state_rank_configuration_None:
  "state_rank q n = None  ¬sink q  configuration q n = {}"
  unfolding state_rank.simps by (metis option.distinct(1))

lemma push_down_state_rank_oldest_token:
  "state_rank q n = Some i  x. oldest_token q n = Some x"
  by (metis oldest_token.elims state_rank.elims)

lemma push_down_state_rank_token_run:
  "state_rank q n = Some i  x. token_run x n = q  x  n"
  by (blast dest: push_down_state_rank_oldest_token push_down_oldest_token_token_run oldest_token_bounded)

subsubsection ‹Properties›

lemma state_rank_distinct:
  assumes distinct: "p  q"
  assumes ranked_1: "state_rank p n = Some i"
  assumes ranked_2: "state_rank q n = Some j"
  shows "i  j"
proof
  assume "i = j"
  obtain x y where "x  configuration p n" and "y  configuration q n"
    using assms push_down_state_rank_tokens by blast
  hence "rank x n = Some i" and "rank y n = Some j"
    using assms pull_up_configuration_rank unfolding state_rank_eq_rank_SOME
    by (metis all_not_in_conv someI_ex)+
  hence "x  configuration q n"
    using y  configuration q n push_down_rank_tokens
    unfolding i = j by auto
  hence "p = q"
    using x  configuration p n by fastforce
  thus "False"
    using distinct by blast
qed

lemma state_rank_initial_state:
  obtains i where "state_rank q0 n = Some i"
  unfolding state_rank.simps sink_def by fastforce

lemma state_rank_sink:
  "sink q  state_rank q n = None"
  by simp

lemma state_rank_upper_bound:
  "state_rank q n = Some i  i < max_rank"
  by (metis option.simps(5) rank_upper_bound push_down_state_rank_oldest_token state_rank_eq_rank)

lemma state_rank_range:
  "state_rank q n  {None}  Some ` {0..<max_rank}"
  by (cases "state_rank q n") (simp add: state_rank_upper_bound[of q n])+

lemma state_rank_None:
  "¬sink q  state_rank q n = None  oldest_token q n = None"
  by simp

lemma state_rank_Some:
  "¬sink q  (i. state_rank q n = Some i)  (j. oldest_token q n = Some j)"
  by simp

lemma state_rank_oldest_token:
  assumes "state_rank p n = Some i"
  assumes "state_rank q n = Some j"
  assumes "oldest_token p n = Some x"
  assumes "oldest_token q n = Some y"
  shows "i < j  x < y"
proof -
  have "configuration p n  {}" and "configuration q n  {}"
    using assms(3,4) by (metis oldest_token.simps option.distinct(1))+
  moreover
  have "¬sink p" and "¬sink q"
    using assms(1,2) state_rank_sink by auto
  ultimately
  have i_def: "i = card (senior_states p n)" and j_def: "j = card (senior_states q n)"
    using assms(1,2) option.sel by simp_all
  hence "i < j  senior_states p n  senior_states q n"
    using senior_states_card_le by presburger
  also
  with assms(3,4) have "  x < y"
  proof (cases rule: senior_states_cases_subset[of p n q])
    case equal
      thus ?thesis
        using assms state_rank_distinct i_def j_def
        by (metis less_irrefl option.sel)
  qed auto
  ultimately
  show ?thesis
    by meson
qed

lemma state_rank_oldest_token_le:
  assumes "state_rank p n = Some i"
  assumes "state_rank q n = Some j"
  assumes "oldest_token p n = Some x"
  assumes "oldest_token q n = Some y"
  shows "i  j  x  y"
  using state_rank_oldest_token[OF assms] assms state_rank_distinct oldest_token_equal
  by (cases "x = y") ((metis option.sel order_refl), (metis le_eq_less_or_eq option.inject))

lemma state_rank_in_function_set:
  shows "(λq. state_rank q t)  {f. (x. x  reach Σ δ q0  f x = None) 
      (x. x  reach Σ δ q0  f x  {None}  Some ` {0..<max_rank})}"
proof -
  {
    fix x
    assume "x  reach Σ δ q0"
    hence "token. x  token_run token t"
      unfolding reach_def token_run.simps using bounded_w by fastforce
    hence "state_rank x t = None"
      using pull_up_configuration_state_rank by auto
  }
  with state_rank_range show ?thesis
    by blast
qed

subsection ‹Step Function›

fun pre_oldest_tokens :: "'b  nat  nat set"
where
  "pre_oldest_tokens q n = {x. q'. oldest_token q' n = Some x  q = δ q' (w n)}  (if q = q0 then {Suc n} else {})"

lemma pre_oldest_configuration_range:
  "pre_oldest_tokens q n  {0..Suc n}"
proof -
  have "{x. q'. oldest_token q' n = Some x  q = δ q' (w n)}  {0..n}"
    (is "?lhs  ?rhs")
  proof
    fix x
    assume "x  ?lhs"
    then obtain q' where "oldest_token q' n = Some x"
      by blast
    thus "x  ?rhs"
      unfolding atLeastAtMost_iff using oldest_token_bounded[of q' n x] by blast
  qed
  thus ?thesis
    by (cases "q = q0") fastforce+
qed

lemma pre_oldest_configuration_finite:
  "finite (pre_oldest_tokens q n)"
  using pre_oldest_configuration_range finite_atLeastAtMost by (rule finite_subset)

lemmas pre_oldest_configuration_Min_in = Min_in[OF pre_oldest_configuration_finite]

lemma pre_oldest_configuration_obtain:
  assumes "x  pre_oldest_tokens q n - {Suc n}"
  obtains q' where "oldest_token q' n = Some x" and "q = δ q' (w n)"
  using assms by (cases "q = q0", auto)

lemma pre_oldest_configuration_element:
  assumes "oldest_token q' n = Some ot"
  assumes "q = δ q' (w n)"
  shows "ot  pre_oldest_tokens q n"
proof
  show "ot  {ot. q'. oldest_token q' n = Some ot  q = δ q' (w n)}"
    (is "_  ?A")
    using assms by blast
  show "?A  pre_oldest_tokens q n"
    by simp
qed

lemma pre_oldest_configuration_initial_state:
  "Suc n  pre_oldest_tokens q n  q = q0"
  using oldest_token_bounded[of _ n "Suc n"]
  by (cases "q = q0") auto

lemma pre_oldest_configuration_initial_state_2:
  "q = q0  Suc n  pre_oldest_tokens q n"
  by fastforce

lemma pre_oldest_configuration_tokens:
  "pre_oldest_tokens q n  {}  configuration q (Suc n)  {}"
  (is "?lhs  ?rhs")
proof
  assume ?lhs
  then obtain ot where ot_def: "ot  pre_oldest_tokens q n"
    by blast
  thus ?rhs
  proof (cases "ot = Suc n")
    case True
      thus ?thesis
    using pre_oldest_configuration_initial_state configuration_non_empty[of "Suc n" "Suc n"] ot  pre_oldest_tokens q n unfolding token_run_intial_state  by blast
  next
    case False
      then obtain q' where "oldest_token q' n = Some ot" and "q = δ q' (w n)"
        using ot_def pre_oldest_configuration_obtain by blast
      moreover
      hence "configuration q' n  {}"
        by (metis oldest_token.simps option.distinct(2))
      ultimately
      show ?rhs
        by (elim configuration_step_non_empty)
  qed
next
  assume ?rhs
  then obtain token where "token  configuration q (Suc n)" and "token  Suc n" and "token_run token (Suc n) = q"
    by auto
  moreover
  {
    assume "token  n"
    then obtain q' where "token_run token n = q'" and "q = δ q' (w n)"
      using token_run token (Suc n) = q unfolding token_run.simps Suc_diff_le[OF token  n] by fastforce
    then obtain ot where "oldest_token q' n = Some ot"
      using oldest_token_always_def by blast
    with q = δ q' (w n) have ?lhs
      using pre_oldest_configuration_element by blast
  }
  ultimately
  show ?lhs
    using pre_oldest_configuration_initial_state_2 by fastforce
qed

lemma oldest_token_rec:
  "oldest_token q (Suc n) = (if pre_oldest_tokens q n  {} then Some (Min (pre_oldest_tokens q n)) else None)"
proof (cases "oldest_token q (Suc n)")
  case (Some ot)
    moreover
    hence "ot  configuration q (Suc n)"
      by (rule push_down_oldest_token_configuration)
    hence "configuration q (Suc n)  {}"
      by blast
    hence "pre_oldest_tokens q n  {}"
      unfolding pre_oldest_configuration_tokens .
    let ?ot = "Min (pre_oldest_tokens q n)"
    {
      {
        {
          assume "ot < Suc n"
          hence "ot  Suc n"
            by blast
          then obtain q' where "ot  configuration q' n" and "q = δ q' (w n)"
            using configuration_rev_step' ot  configuration q (Suc n) by metis
          {
            fix token
            assume "token  configuration q' n"
            hence "token  configuration q (Suc n)"
              using q = δ q' (w n) by (rule configuration_step)
            hence "ot  token"
              using Some by (metis Min.coboundedI configuration q (Suc n)  {} configuration_finite oldest_token.simps option.inject)
          }
          hence "Min (configuration q' n) = ot"
            by (metis Min_eqI ot  configuration q' n configuration_finite)
          hence "oldest_token q' n = Some ot"
            using ot  configuration q' n unfolding oldest_token.simps by auto
          hence "ot  pre_oldest_tokens q n"
            using q = δ q' (w n) by (rule pre_oldest_configuration_element)
        }
        moreover
        {
          assume "ot = Suc n"
          moreover
          hence "q = q0"
            using Some by (metis push_down_oldest_token_token_run token_run_intial_state)
          ultimately
          have "ot  pre_oldest_tokens q n"
            by simp
        }
        ultimately
        have "ot  pre_oldest_tokens q n"
          using Some[THEN oldest_token_bounded] by linarith
      }
      moreover
      {
        fix ot' q'
        assume "oldest_token q' n = Some ot'" and "q = δ q' (w n)"
        moreover
        hence "ot'  configuration q (Suc n)"
          using push_down_oldest_token_configuration configuration_step by blast
        hence "ot  ot'"
          using Some by (metis Min.coboundedI configuration q (Suc n)  {} configuration_finite oldest_token.simps option.inject)
      }
      hence "y. y  pre_oldest_tokens q n - {Suc n}  ot  y"
        using pre_oldest_configuration_obtain by metis
      hence "y. y  pre_oldest_tokens q n  ot  y"
        using Some[THEN oldest_token_bounded] by force
      ultimately
      have "?ot = ot"
        using Min_eqI[OF pre_oldest_configuration_finite, of q n ot] by fast
    }
    ultimately
    show ?thesis
      unfolding pre_oldest_configuration_tokens oldest_token.simps
      by (metis configuration q (Suc n)  {})
qed (unfold pre_oldest_configuration_tokens oldest_token.simps, metis option.distinct(2))

lemma pre_ranks_range:
  "pre_ranks (λq. state_rank q n) ν q   {0..max_rank}"
proof -
  have "{i | q' i. state_rank q' n = Some i  q = δ q' ν}  {0..max_rank}"
    using state_rank_upper_bound by fastforce
  thus ?thesis
    by auto
qed

lemma pre_ranks_finite:
  "finite (pre_ranks (λq. state_rank q n) ν q)"
  using pre_ranks_range finite_atLeastAtMost by (rule finite_subset)

lemmas pre_ranks_Min_in = Min_in[OF pre_ranks_finite]

lemma pre_ranks_state_obtain:
  assumes "rq  pre_ranks r ν q - {max_rank}"
  obtains q' where "r q' = Some rq" and "q = δ q' ν"
  using assms by (cases "q = q0", auto)

lemma pre_ranks_element:
  assumes "state_rank q' n = Some r"
  assumes "q = δ q' (w n)"
  shows "r  pre_ranks (λq. state_rank q n) (w n) q"
proof
  show "r  {i. q'. (λq. state_rank q n) q' = Some i  q = δ q' (w n)}"
    (is "_  ?A")
    using assms by blast
  show "?A  pre_ranks (λq. state_rank q n) (w n) q"
    by simp
qed

lemma pre_ranks_initial_state:
  "max_rank  pre_ranks (λq. state_rank q n) ν q  q = q0"
  using state_rank_upper_bound by (cases "q = q0") auto

lemma pre_ranks_initial_state_2:
  "q = q0  max_rank  pre_ranks r ν q"
  by fastforce

lemma pre_ranks_tokens:
  assumes "¬sink q"
  shows "pre_ranks (λq. state_rank q n) (w n) q  {}  configuration q (Suc n)  {}"
  (is "?lhs = ?rhs")
proof
  assume ?lhs
  thus ?rhs
  proof (cases "q  q0")
    case True
      hence "{i. q'. state_rank q' n = Some i  q = δ q' (w n)}  {}"
        using ?lhs by simp
      then obtain q' where "state_rank q' n  None" and "q = δ q' (w n)"
        by blast
      moreover
      hence "configuration q' n  {}"
        unfolding state_rank.simps by meson
      ultimately
      show ?rhs
        by (elim configuration_step_non_empty)
  qed auto
next
  assume ?rhs
  then obtain token where "token  configuration q (Suc n)" and "token  Suc n" and "token_run token (Suc n) = q"
    by auto
  moreover
  {
    assume "token  n"
    then obtain q' where "token_run token n = q'" and "q = δ q' (w n)"
      using token_run token (Suc n) = q unfolding token_run.simps Suc_diff_le[OF token  n] by fastforce
    hence "¬sink q'"
      using ¬sink q sink_rev_step bounded_w by blast
    then obtain r where "state_rank q' n = Some r"
      using ¬sink q configuration_non_empty[OF token  n] unfolding token_run token n = q' by simp
    with q = δ q' (w n) have ?lhs
      using pre_ranks_element by blast
  }
  ultimately
  show ?lhs
    by fastforce
qed

lemma pre_ranks_pre_oldest_token_Min_state_special:
  assumes "¬sink q"
  assumes "configuration q (Suc n)  {}"
  shows "Min (pre_ranks (λq. state_rank q n) (w n) q) = max_rank  Min (pre_oldest_tokens q n) = Suc n"
  (is "?lhs  ?rhs")
proof
  from assms have "pre_oldest_tokens q n  {}"
    and "pre_ranks  (λq. state_rank q n) (w n) q  {}"
    using pre_ranks_tokens pre_oldest_configuration_tokens by simp_all

  {
    assume ?lhs
    have "q = q0"
      apply (rule ccontr)
      using state_rank_upper_bound pre_ranks_Min_in[OF pre_ranks (λq. state_rank q n) (w n) q  {}] ?lhs
      by auto
    moreover
    {
      fix q'
      assume "q = δ q' (w n)"
      hence "¬sink q'"
        using ¬sink q bounded_w unfolding sink_def
        using calculation by blast
      {
        fix i
        assume "state_rank q' n = Some i"
        hence "False"
          using q = δ q' (w n)
        using Min.coboundedI[OF pre_ranks_finite, of _ n "(w n)" q]
        unfolding ?lhs using state_rank_upper_bound[of q' n] by fastforce
      }
      hence "state_rank q' n = None"
        by fastforce
      hence "oldest_token q' n = None"
        using ¬sink q' by (metis state_rank_None)
    }
    hence "{ot. q'. oldest_token q' n = Some ot  q = δ q' (w n)} = {}"
      by fastforce
    ultimately
    show "?rhs"
      by auto
  }

  {
    assume ?rhs
    {
      fix q'
      assume "q = δ q' (w n)"
      have "state_rank q' n = None"
      proof (cases "oldest_token q' n")
        case (Some t)
          hence "t  n"
            using oldest_token_bounded[of q' n] by blast
          moreover
          have "Suc n  t"
            using q = δ q' (w n)
            using Min.coboundedI[OF pre_oldest_configuration_finite, of _ q n]
            unfolding ?rhs using oldest_token q' n = Some t by auto
          ultimately
          have "False"
            by linarith
          thus ?thesis
            ..
      qed (unfold state_rank_eq_rank, auto)
    }
    hence X: "{i. q'.  (λq. state_rank q n) q' = Some i  q = δ q' (w n)} = {}"
      by fastforce

    have "q = q0"
      apply (rule ccontr)
      using pre_ranks (λq. state_rank q n) (w n) q  {}
      unfolding pre_ranks.simps X by simp
    hence "pre_ranks (λq. state_rank q n) (w n) q = {max_rank}"
      unfolding pre_ranks.simps X by force
    thus ?lhs
      by fastforce
  }
qed

lemma pre_ranks_pre_oldest_token_Min_state:
  assumes "¬sink q"
  assumes "q = δ q' (w n)"
  assumes "configuration q (Suc n)  {}"
  defines "min_r  Min (pre_ranks (λq. state_rank q n) (w n) q)"
  defines "min_ot  Min (pre_oldest_tokens q n)"
  shows "state_rank q' n = Some min_r  oldest_token q' n = Some min_ot"
  (is "?lhs  ?rhs")
proof
  from assms have "pre_oldest_tokens q n  {}" and "¬sink q'"
    and "pre_ranks (λq. state_rank q n) (w n) q  {}"
    using pre_ranks_tokens pre_oldest_configuration_tokens bounded_w unfolding sink_def
    by (simp_all, metis rangeI subset_iff)

  {
    assume ?lhs
    thus ?rhs
    proof (cases min_r max_rank rule: linorder_cases)
      case less
        then obtain ot where "oldest_token q' n = Some ot"
           by (metis push_down_state_rank_oldest_token ?lhs)
        moreover
        {
          {
            fix q'' ot''
            assume "q = δ q'' (w n)"
            assume "oldest_token q'' n = Some ot''"
            moreover
            have "¬sink q''"
              using q = δ q'' (w n) assms unfolding sink_def
              by (metis rangeI subset_eq bounded_w)
            then obtain r'' where "state_rank q'' n = Some r''"
              using oldest_token q'' n = Some ot'' by (metis state_rank_Some)
            moreover
            hence "r''  pre_ranks (λq. state_rank q n) (w n) q" (* Move to special lemma *)
              using q = δ q'' (w n) unfolding pre_ranks.simps by blast
            then have "min_r  r''"
              unfolding min_r_def by (metis Min.coboundedI pre_ranks_finite)
            ultimately
            have "ot  ot''"
              using state_rank_oldest_token_le[OF ?lhs _ oldest_token q' n = Some ot] by blast
          }
          hence "x. x  {ot. q'. oldest_token q' n = Some ot  q = δ q' (w n)}  ot  x"
            by blast
          moreover
          have "ot  Suc n"
            using oldest_token_bounded[OF oldest_token q' n = Some ot] by simp
          ultimately
          have "x. x  pre_oldest_tokens q n  ot  x"
            unfolding pre_oldest_tokens.simps apply (cases "q0 = q") apply auto done
          hence "ot  min_ot"
            unfolding min_ot_def
            unfolding Min_ge_iff[OF pre_oldest_configuration_finite pre_oldest_tokens q n  {}, of ot]
             by simp
        }
        moreover
        have "ot  min_ot"
          using Min.coboundedI[OF pre_oldest_configuration_finite] pre_oldest_configuration_element
          unfolding min_ot_def by (metis assms(2) calculation(1))
        ultimately
        show ?thesis
          by simp
    qed (insert not_less, blast intro: state_rank_upper_bound less_imp_le_nat)+
  }

  {
    assume ?rhs
    thus ?lhs
    proof (cases min_ot "Suc n" rule: linorder_cases)
      case less
        then obtain r where "state_rank q' n = Some r"
          using ?rhs ¬sink q' by (metis state_rank_Some)
        moreover
        {
          {
            fix r''
            assume "r''  pre_ranks (λq. state_rank q n) (w n) q - {max_rank}"
            then obtain q'' where "state_rank q'' n = Some r''"
              and "q = δ q'' (w n)"
              using pre_ranks_state_obtain by blast
            moreover
            then obtain ot'' where "oldest_token q'' n = Some ot''"
               using push_down_state_rank_oldest_token by fastforce
            moreover
            hence "min_ot  ot''"
              using q = δ q'' (w n) pre_oldest_configuration_element Min.coboundedI pre_oldest_configuration_finite
              unfolding min_ot_def by metis
            ultimately
            have "r  r''"
              using state_rank_oldest_token_le[OF state_rank q' n = Some r _ ?rhs] by blast
          }
          moreover
          have "r  max_rank"
            using state_rank_upper_bound[OF state_rank q' n = Some r] by linarith
          ultimately
          have "x. x   pre_ranks (λq. state_rank q n) (w n) q  r  x"
            unfolding pre_ranks.simps apply (cases "q0 = q") apply auto done
          hence "r  min_r"
            unfolding min_r_def Min_ge_iff[OF pre_ranks_finite pre_ranks (λq. state_rank q n) (w n) q  {}]
            by simp
        }
        moreover
        have "r  min_r"
          using Min.coboundedI[OF pre_ranks_finite] pre_ranks_element
          unfolding min_r_def by (metis assms(2) calculation(1))
        ultimately
        show ?thesis
          by simp
    qed (insert not_less, blast intro: oldest_token_bounded Suc_lessD)+
  }
qed

lemma Min_pre_ranks_pre_oldest_tokens:
  fixes n
  defines "r  (λq. state_rank q n)"
  assumes "configuration p (Suc n)  {}"
      and "configuration q (Suc n)  {}"
  assumes "¬sink q"
      and "¬sink p"
  shows "Min (pre_ranks r (w n) p) < Min (pre_ranks r (w n) q)  Min (pre_oldest_tokens p n) < Min (pre_oldest_tokens q n)"
  (is "?lhs  ?rhs")
proof
  have pre_ranks_Min: "x ν. (x < Min (pre_ranks r (w n) q)) = (a  pre_ranks r (w n) q. x < a)"
    using assms pre_ranks_finite Min.bounded_iff pre_ranks_tokens by simp
  have pre_oldest_configuration_Min: "x. (x < Min (pre_oldest_tokens q n)) = (apre_oldest_tokens q n. x < a)"
    using assms pre_oldest_configuration_finite Min.bounded_iff pre_oldest_configuration_tokens by simp
  have "x. w x  Σ"
    using bounded_w by auto

  {
    let ?min_i = "Min (pre_ranks r (w n) p)"
    let ?min_j = "Min (pre_ranks r (w n) q)"

    assume ?lhs

    have "?min_i  pre_ranks r (w n) p" and "?min_j  pre_ranks r (w n) q"
      using Min_in[OF pre_ranks_finite] assms pre_ranks_tokens by presburger+
    hence "?min_i  max_rank" and "?min_j  max_rank"
      using pre_ranks_range atLeastAtMost_iff unfolding r_def by blast+
    with ?lhs have "?min_i  max_rank"
      by linarith
    then obtain p' i' where "i' = ?min_i" and "r p' = Some i'" and "p = δ p' (w n)"
      using ?min_i  pre_ranks r (w n) p apply (cases "p = q0") apply auto[1] by fastforce
    then obtain ot' where "oldest_token p' n = Some ot'"
      unfolding assms by (metis push_down_state_rank_oldest_token)
    have "state_rank p' n = Some ?min_i"
      using i' = ?min_i r p' = Some i' unfolding assms by simp
    hence "ot' = Min (pre_oldest_tokens p n)"
      using pre_ranks_pre_oldest_token_Min_state[OF ¬sink p p = δ p' (w n) configuration p (Suc n)  {}] oldest_token p' n = Some ot'
      unfolding r_def by (metis option.inject)
    moreover
    have "ot' < Suc n"
    proof (cases ot' "Suc n" rule: linorder_cases)
      case equal
        hence "?min_i = max_rank"
          using pre_ranks_pre_oldest_token_Min_state_special[of p n, OF ¬sink p configuration p (Suc n)  {}] assms
          unfolding ot' = Min (pre_oldest_tokens p n) by simp
        thus ?thesis
          using ?min_i  max_rank by simp
    next
      case greater
        moreover
        have "ot'  {0..Suc n}"
          using oldest_token p' n = Some ot'[THEN oldest_token_bounded] by fastforce
        ultimately
        show ?thesis
          by simp
    qed simp
    moreover
    {
      fix otq
      assume "otq  pre_oldest_tokens q n - {Suc n}"
      then obtain q' where "oldest_token q' n = Some otq" and "q = δ q' (w n)"
        using pre_oldest_configuration_obtain by blast
      moreover
      hence "¬sink q'"
        using ¬sink q x. w x  Σ unfolding sink_def by auto
      then obtain rq where "state_rank q' n = Some rq"
        unfolding assms state_rank.simps using oldest_token q' n = Some otq
        by (metis oldest_token.simps option.distinct(2))
      moreover
      hence "rq  pre_ranks r (w n) q"
        using q = δ q' (w n)
        unfolding pre_ranks.simps assms by blast
      hence "?min_j  rq"
        using Min.coboundedI[OF pre_ranks_finite] unfolding assms by blast
      hence "?min_i < rq"
        using ?lhs by linarith
      hence "ot' < otq"
        using state_rank_oldest_token[OF state_rank p' n = Some ?min_i state_rank q' n = Some rq oldest_token p' n = Some ot' oldest_token q' n = Some otq]
        unfolding assms by simp
    }
    ultimately
    show ?rhs
      using pre_oldest_configuration_Min by blast
  }

  {
    define ot_p where "ot_p = Min (pre_oldest_tokens p n)"
    define ot_q where "ot_q = Min (pre_oldest_tokens q n)"
    assume ?rhs
    hence "ot_p < ot_q"
      unfolding ot_p_def ot_q_def .

    have "oldest_token p (Suc n) = Some ot_p" and "oldest_token q (Suc n) = Some ot_q"
      unfolding ot_p_def ot_q_def oldest_token_rec pre_oldest_configuration_tokens by (metis assms)+

   (* Min oldest ⟷ Min rank *)
    define min_rp where "min_rp = Min (pre_ranks r (w n) p)"
    hence "min_rp  pre_ranks r (w n) p"
      using pre_ranks_Min_in assms pre_ranks_tokens by simp
    hence *: "min_rp < max_rank"
    proof (cases min_rp max_rank rule: linorder_cases)
      case equal
        hence "ot_p = Suc n"
          using pre_ranks_pre_oldest_token_Min_state_special[of p n, OF _ configuration p (Suc n)  {}] assms
          unfolding ot_p_def min_rp_def by simp
        moreover
        have "Min (pre_oldest_tokens q n)  pre_oldest_tokens q n"
          using Min_in[OF pre_oldest_configuration_finite ] assms pre_oldest_configuration_tokens by presburger
        hence "ot_q  {0..Suc n}"
          using pre_oldest_configuration_range[of q n]
          unfolding ot_q_def by blast
        hence "ot_q  Suc n"
          by simp
        ultimately
        show ?thesis
          using ot_p < ot_q by simp
    next
      case greater
        moreover
        have "min_rp  {0..max_rank}"
          using pre_ranks_range min_rp  pre_ranks r (w n) p
          unfolding r_def ..
        ultimately
        show ?thesis
          by simp
    qed simp
    moreover
    from * have "min_rp  pre_ranks r (w n) p - {max_rank}"
      using min_rp  pre_ranks r (w n) p by simp
    then obtain p' where "r p' = Some min_rp" and "p = δ p' (w n)"
      using pre_ranks_state_obtain by blast
    hence "oldest_token p' n = Some ot_p"
      using pre_ranks_pre_oldest_token_Min_state[OF ¬sink p p = δ p' (w n) configuration p (Suc n)  {}]
      unfolding r_def[symmetric] min_rp_def[symmetric] ot_p_def[symmetric] by (metis r_def)
    {
      fix rq
      assume "rq  pre_ranks r (w n) q - {max_rank}"
      then obtain q' where q': "r q' = Some rq" "q = δ q' (w n)"
        using pre_ranks_state_obtain by blast
      moreover
      from q' obtain ot_q' where ot_q': "oldest_token q' n = Some ot_q'"
        unfolding assms by (metis push_down_state_rank_oldest_token)
      moreover
      from ot_q' have "ot_q'  pre_oldest_tokens q n"
        using q = δ q' (w n)
        unfolding pre_oldest_tokens.simps by blast
      hence "ot_q  ot_q'"
        unfolding ot_q_def
        by (rule Min.coboundedI[OF pre_oldest_configuration_finite])
      hence "ot_p < ot_q'"
        using ot_p < ot_q by linarith
      ultimately
      have "min_rp < rq"
        using state_rank_oldest_token r p' = Some min_rp oldest_token p' n = Some ot_p
        unfolding assms by blast
    }
    ultimately
    show ?lhs
      using pre_ranks_Min unfolding min_rp_def by blast
  }
qed

subsubsection ‹Definition of initial and step›

lemma state_rank_initial:
  "state_rank q 0 = initial q"
  using state_rank_initial_state by force

lemma state_rank_step:
  "state_rank q (Suc n) = step (λq. state_rank q n) (w n) q"
  (is "?lhs = ?rhs")
proof (cases "sink q")
  case False
    {
      assume "configuration q (Suc n) = {}"
      hence ?thesis
        using False pull_up_configuration_state_rank pre_ranks_tokens
        unfolding step.simps by presburger
    }
    moreover
    {
      assume "configuration q (Suc n)  {}"
      hence "?lhs = Some (card (senior_states q (Suc n)))"
        using False unfolding state_rank.simps by presburger
      also
      have " = ?rhs"
      proof -
        let ?r = "λq. state_rank q n"
        have "{q'. ¬sink q'  pre_ranks ?r (w n) q'  {}  Min (pre_ranks ?r (w n) q') < Min (pre_ranks ?r (w n) q)} = senior_states q (Suc n)"
          (is "?S = ?S'")
        proof (rule set_eqI)
          fix q'
          have "q'  ?S  ¬sink q'  configuration q' (Suc n)  {}  Min (pre_ranks  ?r (w n) q') < Min (pre_ranks ?r (w n) q)"
            using pre_ranks_tokens by blast
          also
          have "  ¬sink q'  configuration q' (Suc n)  {}  Min (pre_oldest_tokens q' n) < Min (pre_oldest_tokens q n)"
            by (metis configuration q (Suc n)  {} ¬sink q Min_pre_ranks_pre_oldest_tokens)
          also
          have "  ¬sink q'  (x y. oldest_token q' (Suc n) = Some y  oldest_token q (Suc n) = Some x  y < x)"
            unfolding oldest_token_rec by (metis pre_oldest_configuration_tokens configuration q (Suc n)  {} option.distinct(2) option.sel)
          finally
          show "q'  ?S  q'  ?S'"
            unfolding senior_states.simps by blast
        qed
        thus ?thesis
          using ¬sink q configuration q (Suc n)  {}
          unfolding step.simps pre_ranks_tokens[OF ¬sink q] by presburger
      qed
      finally
      have ?thesis .
    }
    ultimately
    show ?thesis
      by blast
qed auto

lemma state_rank_step_foldl:
  "(λq. state_rank q n) = foldl step initial (map w [0..<n])"
  by (induction n) (unfold state_rank_initial state_rank_step, simp_all)

end

end