Theory DTS

(*  
    Author:      Salomon Sickert
    License:     BSD
*)

section ‹Deterministic Transition Systems›

theory DTS
  imports Main "HOL-Library.Omega_Words_Fun" "Auxiliary/Mapping2" KBPs.DFS
begin

― ‹DTS are realised by functions›

type_synonym ('a, 'b) DTS = "'a  'b  'a"
type_synonym ('a, 'b) transition = "('a × 'b × 'a)"

subsection ‹Infinite Runs›

fun run :: "('q,'a) DTS  'q  'a word  'q word"
where
  "run δ q0 w 0 = q0"
| "run δ q0 w (Suc i) = δ (run δ q0 w i) (w i)"

fun runt :: "('q,'a) DTS  'q  'a word  ('q * 'a * 'q) word"
where
  "runt δ q0 w i = (run δ q0 w i, w i, run δ q0 w (Suc i))"

lemma run_foldl:
  "run Δ q0 w i = foldl Δ q0 (map w [0..<i])"
  by (induction i; simp)

lemma runt_foldl:
  "runt Δ q0 w i = (foldl Δ q0 (map w [0..<i]), w i, foldl Δ q0 (map w [0..<Suc i]))"
  unfolding runt.simps run_foldl ..

subsection ‹Reachable States and Transitions›

definition reach :: "'a set  ('b, 'a) DTS  'b  'b set"
where
  "reach Σ δ q0 = {run δ q0 w n | w n. range w  Σ}"

definition reacht :: "'a set  ('b, 'a) DTS  'b  ('b, 'a) transition set"
where
  "reacht Σ δ q0 = {runt δ q0 w n | w n. range w  Σ}"

lemma reach_foldl_def:
  assumes "Σ  {}"
  shows "reach Σ δ q0 = {foldl δ q0 w | w. set w  Σ}"
proof -
  {
    fix w assume "set w  Σ"
    moreover
    obtain a where "a  Σ"
      using Σ  {} by blast
    ultimately
    have "foldl δ q0 w = foldl δ q0 (prefix (length w) (w  (iter [a])))" 
      and "range (w  (iter [a]))  Σ"
      by (unfold prefix_conc_length, auto simp add: iter_def conc_def) 
    hence "w' n. foldl δ q0 w = run δ q0 w' n  range w'  Σ"
      unfolding run_foldl subsequence_def by blast
  }
  thus ?thesis
    by (fastforce simp add: reach_def run_foldl)
qed

lemma reacht_foldl_def:
  "reacht Σ δ q0 = {(foldl δ q0 w, ν, foldl δ q0 (w@[ν])) | w ν. set w  Σ  ν  Σ}" (is "?lhs = ?rhs")
proof (cases "Σ  {}")
  case True
    show ?thesis
    proof
      {
        fix w ν assume "set w  Σ" "ν  Σ"
        moreover
        obtain a where "a  Σ"
          using Σ  {} by blast
        moreover
        have "w = map (λn. if n < length w then w ! n else if n - length w = 0 then [ν] ! (n - length w) else a) [0..<length w]"
          by (simp add: nth_equalityI)  
        ultimately
        have "foldl δ q0 w = foldl δ q0 (prefix (length w) ((w @ [ν])  (iter [a])))" 
          and"foldl δ q0 (w @ [ν]) = foldl δ q0 (prefix (length (w @ [ν])) ((w @ [ν])  (iter [a])))" 
          and "range ((w @ [ν])  (iter [a]))  Σ"
          by (simp_all only: prefix_conc_length conc_conc[symmetric] iter_def)
             (auto simp add: subsequence_def conc_def upt_Suc_append[OF le0])
        moreover
        have "((w @ [ν])  (iter [a])) (length w) = ν"
          by (simp add: conc_def)
        ultimately
        have "w' n. (foldl δ q0 w, ν, foldl δ q0 (w @ [ν])) = runt δ q0 w' n  range w'  Σ"
          by (metis runt_foldl length_append_singleton subsequence_def)
      }
      thus "?lhs  ?rhs"
        unfolding reacht_def runt.simps by blast
    qed (unfold reacht_def runt_foldl, fastforce simp add: upt_Suc_append)
qed (simp add: reacht_def)

lemma reach_card_0:
  assumes "Σ  {}"
  shows "infinite (reach Σ δ q0)  card (reach Σ δ q0) = 0"
proof -
  have "{run δ q0 w n | w n. range w  Σ}  {}"
    using assms by fast
  thus ?thesis
    unfolding reach_def card_eq_0_iff by auto
qed

lemma reacht_card_0:
  assumes "Σ  {}"
  shows "infinite (reacht Σ δ q0)  card (reacht Σ δ q0) = 0"
proof -
  have "{runt δ q0 w n | w n. range w  Σ}  {}"
    using assms by fast
  thus ?thesis
    unfolding reacht_def card_eq_0_iff by blast
qed

subsubsection ‹Relation to runs›

lemma run_subseteq_reach:
  assumes "range w  Σ"
  shows "range (run δ q0 w)  reach Σ δ q0" 
    and "range (runt δ q0 w)  reacht Σ δ q0"
  using assms unfolding reach_def reacht_def by blast+

lemma limit_subseteq_reach:
  assumes "range w  Σ"
  shows "limit (run δ q0 w)  reach Σ δ q0"
    and "limit (runt δ q0 w)  reacht Σ δ q0"
  using run_subseteq_reach[OF assms] limit_in_range by fast+

lemma runt_finite:
  assumes "finite (reach Σ δ q0)"
  assumes "finite Σ"
  assumes "range w  Σ"
  defines "r  runt δ q0 w"
  shows "finite (range r)"
proof -
  let ?S = "(reach Σ δ q0) × Σ × (reach Σ δ q0)"
  have "i. w i  Σ" and "i. set (map w [0..<i])  Σ" and "Σ  {}"
    using range w  Σ by auto 
  hence "n. r n  ?S"
    unfolding runt.simps run_foldl reach_foldl_def[OF Σ  {}] r_def by blast
  hence "range r  ?S" and "finite ?S"
    using assms by blast+
  thus "finite (range r)"
    by (blast dest: finite_subset)
qed

subsubsection ‹Compute reach Using DFS›

definition QL :: "'a list  ('b, 'a) DTS  'b  'b set"
where
  "QL Σ δ q0 = (if Σ  [] then gen_dfs (λq. map (δ q) Σ) Set.insert (∈) {} [q0] else {})"

definition list_dfs :: "(('a, 'b) transition  ('a, 'b) transition list)  ('a, 'b) transition list => ('a, 'b) transition list => ('a, 'b) transition list"
where
  "list_dfs succ S start  gen_dfs succ List.insert (λx xs. x  set xs) S start"

definition δL :: "'a list  ('b, 'a) DTS  'b  ('b, 'a) transition set"
where
  "δL Σ δ q0 = set ( 
    let 
      start = map (λν. (q0, ν, δ q0 ν)) Σ; 
      succ = λ(_, _, q). (map (λν. (q, ν, δ q ν)) Σ)
    in 
      (list_dfs succ [] start))"

lemma QL_reach:
  assumes "finite (reach (set Σ) δ q0)"
  shows "QL Σ δ q0 = reach (set Σ) δ q0"
proof (cases "Σ  []")
  case True
    hence reach_redef: "reach (set Σ) δ q0 = {foldl δ q0 w | w. set w  set Σ}"
      using reach_foldl_def[of "set Σ"] unfolding set_empty[of Σ, symmetric] by force
  
    {
      fix x w y
      assume "set w  set Σ" "x = foldl δ q0 w" "y  set (map (δ x) Σ)"
      moreover
      then obtain ν where "y = δ x ν" and "ν  set Σ"
        by auto
      ultimately
      have "y = foldl δ q0 (w@[ν])" and "set (w@[ν])  set Σ"
        by simp+
      hence "w'. set w'  set Σ  y = foldl  δ q0 w'"
        by blast
    }
    note extend_run = this
    
    interpret DFS "λq. map (δ q) Σ" "λq. q  reach (set Σ) δ q0" "λS. S  reach (set Σ) δ q0" Set.insert "(∈)" "{}" id
      apply (unfold_locales; auto simp add: member_rec reach_redef list_all_iff elim: extend_run)
      apply (metis extend_run image_eqI set_map)
      apply (metis assms[unfolded reach_redef])
      done

    have Nil1: "set []  set Σ" and Nil2: "q0 = foldl δ q0 []"
      by simp+
    have list_all_init: "list_all (λq. q  reach (set Σ) δ q0) [q0]"
      unfolding list_all_iff list.set reach_redef using Nil1 Nil2 by blast

    have "reach (set Σ) δ q0  reachable {q0}"
    proof rule
      fix x 
      assume "x  reach (set Σ) δ q0"
      then obtain w where x_def: "x = foldl δ q0 w" and "set w  set Σ"
        unfolding reach_redef by blast
      hence "foldl δ q0 w  reachable {q0}"
      proof (induction w arbitrary: x rule: rev_induct)
        case (snoc ν w)
          hence "foldl δ q0 w  reachable {q0}" and "δ (foldl δ q0 w) ν  set (map (δ (foldl δ q0 w)) Σ)"
            by simp+
          thus ?case
            by (simp add: rtrancl.rtrancl_into_rtrancl reachable_def)
      qed (simp add: reachable_def)
      thus "x  reachable {q0}"
        by (simp add: x_def)
    qed
    moreover
    have "reachable {q0}  reach (set Σ) δ q0"
    proof rule
      fix x 
      assume "x  reachable {q0}"
      hence "(q0, x)  {(x, y). y  set (map (δ x) Σ)}*"
        unfolding reachable_def by blast
      thus "x  reach (set Σ) δ q0"
        apply (induction)
        apply (insert reach_redef Nil1 Nil2; blast)
        apply (metis r_into_rtrancl succsr_def succsr_isNode) 
        done
    qed
    ultimately
    have reachable_redef: "reachable {q0} = reach (set Σ) δ q0"
      by blast

    moreover

    have "reachable {q0}  QL Σ δ q0"
      using reachable_imp_dfs[OF _ list_all_init] unfolding list.set reachable_redef
      unfolding  reach_redef QL_def using Σ  [] by auto

    moreover

    have "QL Σ δ q0  reach (set Σ) δ q0"
      using dfs_invariant[of "{}", OF _ list_all_init] 
      by (auto simp add: reach_redef QL_def)

    ultimately
    show ?thesis
      using Σ  [] dfs_invariant[of "{}", OF _ list_all_init] by simp+
qed (simp add: reach_def QL_def)

lemma δL_reach: 
  assumes "finite (reacht (set Σ) δ q0)"
  shows "δL Σ δ q0 = reacht (set Σ) δ q0"
proof -
  {
    fix x w y
    assume "set w  set Σ" "x = foldl δ q0 w" "y  set (map (δ x) Σ)"
    moreover
    then obtain ν where "y = δ x ν" and "ν  set Σ"
      by auto
    ultimately
    have "y = foldl δ q0 (w@[ν])" and "set (w@[ν])  set Σ"
      by simp+
    hence "w'. set w'  set Σ  y = foldl  δ q0 w'"
      by blast
  }
  note extend_run = this

  let ?succs = "λ(_, _, q). (map (λν. (q, ν, δ q ν)) Σ)"

  interpret DFS "λ(_, _, q). (map (λν. (q, ν, δ q ν)) Σ)" "λt. t  reacht (set Σ) δ q0" "λS. set S  reacht (set Σ) δ q0" List.insert "λx xs. x  set xs" "[]" id
    apply (unfold_locales; auto simp add: member_rec reacht_foldl_def list_all_iff elim: extend_run)
    apply (metis extend_run image_eqI set_map)
    using  assms unfolding reacht_foldl_def by simp

  have Nil1: "set []  set Σ" and Nil2: "q0 = foldl δ q0 []"
    by simp+
  have list_all_init: "list_all (λq. q  reacht (set Σ) δ q0) (map (λν. (q0, ν, δ q0 ν)) Σ)"
    unfolding list_all_iff reacht_foldl_def set_map image_def using Nil2 by fastforce
  
  let ?q0 = "set (map (λν. (q0, ν, δ q0 ν)) Σ)"

  {
    fix q ν q'  
    assume "(q, ν, q')  reacht (set Σ) δ q0"
    then obtain w where q_def: "q = foldl δ q0 w" and q'_def: "q' = foldl δ q0 (w@[ν])" 
      and "set w  set Σ" and "ν  set Σ"
      unfolding reacht_foldl_def by blast
    hence "(foldl δ q0 w, ν, foldl δ q0 (w@[ν]))  reachable ?q0"
    proof (induction w arbitrary: q q' ν rule: rev_induct)
      case (snoc ν' w)
        hence "(foldl δ q0 w, ν', foldl δ q0 (w@[ν']))  reachable ?q0" (is "(?q, ν', ?q')  _")
          and "q. δ q ν  set (map (δ q) Σ)"
          and "ν  set Σ"
          by simp+
        then obtain x0 where 1: "(x0, (?q, ν', ?q'))  {(x, y). y  set (?succs x)}*" and 2: "x0  ?q0"
          unfolding reachable_def by auto
        moreover
        have 3: "((?q, ν', ?q'), (?q', ν, δ ?q' ν))  {(x, y). y  set (?succs x)}"
          using snoc q. δ q ν  set (map (δ q) Σ) by simp
        ultimately
        show ?case
          using rtrancl.rtrancl_into_rtrancl[OF 1 3] 2 unfolding reachable_def foldl_append foldl.simps  by auto
    qed (auto simp add: reachable_def)
    hence "(q, ν, q')  reachable ?q0"
      by (simp add: q_def q'_def)
  }
  hence "reacht (set Σ) δ q0  reachable ?q0"
    by auto
  moreover
  {
    fix y  
    assume "y  reachable ?q0" 
    then obtain x where "(x, y)  {(x, y). y  set (case x of (_, _, q)  map (λν. (q, ν, δ q ν)) Σ)}*"
      and "x  ?q0"
      unfolding reachable_def by auto
    hence "y  reacht (set Σ) δ q0"
    proof induction
      case base
        have "p ps. list_all p ps = (pa. pa  set ps  p pa)"
          by (meson list_all_iff)
        hence "x  {(foldl δ (foldl δ q0 []) bs, b, foldl δ (foldl δ q0 []) (bs @ [b])) | bs b. set bs  set Σ  b  set Σ}"
          using base by (metis (no_types) Nil2 list_all_init reacht_foldl_def)
        thus ?case
          unfolding reacht_foldl_def by auto
    next
      case (step x' y')
        thus ?case using succsr_def succsr_isNode by blast
    qed
  }
  hence "reachable ?q0  reacht (set Σ) δ q0"
     by blast
  ultimately
  have reachable_redef: "reachable ?q0 = reacht (set Σ) δ q0"
    by blast

  moreover

  have "reachable ?q0  (δL Σ δ q0)"
    using reachable_imp_dfs[OF _ list_all_init] unfolding δL_def reachable_redef list_dfs_def
    by (simp; blast)

  moreover

  have "δL Σ δ q0  reacht (set Σ) δ q0"
    using dfs_invariant[of "[]", OF _ list_all_init] 
    by (auto simp add: reacht_foldl_def δL_def list_dfs_def)

  ultimately
  show ?thesis 
    by simp
qed

lemma reach_reacht_fst:
  "reach Σ δ q0 = fst ` reacht Σ δ q0"
  unfolding reacht_def reach_def image_def by fastforce

lemma finite_reach:
  "finite (reacht Σ δ q0)  finite (reach Σ δ q0)"
  by (simp add: reach_reacht_fst)

lemma finite_reacht:
  assumes "finite (reach Σ δ q0)"
  assumes "finite Σ"
  shows "finite (reacht Σ δ q0)"
proof -
  have "reacht Σ δ q0  reach Σ δ q0 × Σ × reach Σ δ q0"
    unfolding reacht_def reach_def runt.simps by blast
  thus ?thesis
    using assms finite_subset by blast
qed

lemma QL_eq_δL:
  assumes "finite (reacht (set Σ) δ q0)"
  shows "QL Σ δ q0 = fst ` (δL Σ δ q0)"
  unfolding set_map δL_reach[OF assms] QL_reach[OF finite_reach[OF assms]] reach_reacht_fst ..

subsection ‹Product of DTS› 

fun simple_product :: "('a, 'c) DTS  ('b, 'c) DTS  ('a × 'b, 'c) DTS" ("_ × _")
where
  "δ1 × δ2 = (λ(q1, q2) ν. (δ1 q1 ν, δ2 q2 ν))"  

lemma simple_product_run:
  fixes δ1 δ2 w q1 q2
  defines "ρ  run (δ1 × δ2) (q1, q2) w"
  defines "ρ1  run δ1 q1 w"
  defines "ρ2  run δ2 q2 w"
  shows "ρ i = (ρ1 i, ρ2 i)"
  by (induction i) (insert assms, auto)

theorem finite_reach_simple_product:
  assumes "finite (reach Σ δ1 q1)"
  assumes "finite (reach Σ δ2 q2)"
  shows "finite (reach Σ (δ1 × δ2) (q1, q2))"
proof -
  have "reach Σ (δ1 × δ2) (q1, q2)  reach Σ δ1 q1 × reach Σ δ2 q2"
    unfolding reach_def simple_product_run by blast
  thus ?thesis
    using assms finite_subset by blast
qed

subsection ‹(Generalised) Product of DTS›

fun product :: "('a  ('b, 'c) DTS)  ('a  'b, 'c) DTS" ("Δ×")
where
  "Δ× δm = (λq ν. (λx. case q x of None  None | Some y  Some (δm x y ν)))"  

lemma product_run_None:
  fixes ιm δm w
  defines "ρ  run (Δ× δm) ιm w"
  assumes "ιm k = None"
  shows "ρ i k = None"
  by (induction i) (insert assms, auto)

lemma product_run_Some:
  fixes ιm δm w q0 k
  defines "ρ  run (Δ× δm) ιm w"
  defines "ρ'  run (δm k) q0 w"
  assumes "ιm k = Some q0"
  shows "ρ i k = Some (ρ' i)"
  by (induction i) (insert assms, auto)

theorem finite_reach_product:
  assumes "finite (dom ιm)"
  assumes "x. x  dom ιm  finite (reach Σ (δm x) (the (ιm x)))"
  shows "finite (reach Σ (Δ× δm) ιm)"
  using assms(1,2) 
proof (induction "dom ιm" arbitrary: ιm) 
  case empty
    hence "ιm = Map.empty"
      by auto
    hence "w i. run (Δ× δm) ιm w i = (λx. None)"
      using product_run_None by fast
    thus ?case
      unfolding reach_def by simp
next 
  case (insert k K)
    define f where "f = (λ(q :: 'b, m :: 'a  'b). m(k := Some q))"
    define Reach where "Reach = (reach Σ (δm k) (the (ιm k))) × ((reach Σ (Δ× δm) (ιm(k := None))))"

    have "(reach Σ (Δ× δm) ιm)  f ` Reach"
    proof
      fix x
      assume "x  reach Σ (Δ× δm) ιm"
      then obtain w n where x_def: "x = run (Δ× δm) ιm w n" and "range w  Σ"
        unfolding reach_def by blast
      {
        fix k'
        have "k'  dom ιm  x k' = run (Δ× δm) (ιm(k := None)) w n k'"
          unfolding x_def dom_def using product_run_None[of _ _ δm] by simp
        moreover
        have "k'  dom ιm - {k}  x k' = run (Δ× δm) (ιm(k := None)) w n k'"
          unfolding x_def dom_def using product_run_Some[of _ _ _ δm] by auto
        ultimately
        have "k'  k  x k' = run (Δ× δm) (ιm(k := None)) w n k'"
          by blast
      }
      hence "x(k := None) = run (Δ× δm) (ιm(k := None)) w n"
        using product_run_None[of _ _ δm] by auto
      moreover
      have "x k = Some (run (δm k) (the (ιm k)) w n)"
        unfolding x_def using product_run_Some[of ιm k _ δm] insert.hyps(4) by force 
      ultimately
      have "(the (x k), x(k := None))  Reach"
        unfolding Reach_def reach_def using range w  Σ by auto
      moreover
      have "x = f (the (x k), x(k := None))"
        unfolding f_def using x k = Some (run (δm k) (the (ιm k)) w n) by auto 
      ultimately
      show "x  f ` Reach"
         by simp
    qed
    moreover
    have "finite (reach Σ (Δ× δm) (ιm (k := None)))"
      using insert insert(3)[of "ιm (k:=None)"] by auto 
    hence "finite Reach"
      using insert Reach_def by blast
    hence "finite (f ` Reach)"
      .. 
    ultimately
    show ?case
      by (rule finite_subset)
qed

subsection ‹Simple Product Construction Helper Functions and Lemmas›

fun embed_transition_fst :: "('a, 'c) transition  ('a × 'b, 'c) transition set"
where
  "embed_transition_fst (q, ν, q') = {((q, x), ν, (q', y)) | x y. True}"

fun embed_transition_snd :: "('b, 'c) transition  ('a × 'b, 'c) transition set"
where
  "embed_transition_snd (q, ν, q') = {((x, q), ν, (y, q')) | x y. True}"

lemma embed_transition_snd_unfold:
  "embed_transition_snd t = {((x, fst t), fst (snd t), (y, snd (snd t))) | x y. True}"
  unfolding embed_transition_snd.simps[symmetric] by simp

fun project_transition_fst :: "('a × 'b, 'c) transition  ('a, 'c) transition" 
where
  "project_transition_fst (x, ν, y) = (fst x, ν, fst y)"

fun project_transition_snd :: "('a × 'b, 'c) transition  ('b, 'c) transition" 
where
  "project_transition_snd (x, ν, y) = (snd x, ν, snd y)"

lemma
  fixes δ1 δ2 w q1 q2
  defines "ρ  runt (δ1 × δ2) (q1, q2) w"
  defines "ρ1  runt δ1 q1 w"
  defines "ρ2  runt δ2 q2 w"
  shows product_run_project_fst: "project_transition_fst (ρ i) = ρ1 i" 
    and product_run_project_snd: "project_transition_snd (ρ i) = ρ2 i"
    and product_run_embed_fst: "ρ i  embed_transition_fst (ρ1 i)"
    and product_run_embed_snd: "ρ i  embed_transition_snd (ρ2 i)"
  unfolding assms runt.simps simple_product_run by simp_all

lemma 
  fixes δ1 δ2 w q1 q2
  defines "ρ  runt (δ1 × δ2) (q1, q2) w"
  defines "ρ1  runt δ1 q1 w"
  defines "ρ2  runt δ2 q2 w"
  assumes "finite (range ρ)"
  shows product_run_finite_fst: "finite (range ρ1)"
    and product_run_finite_snd: "finite (range ρ2)"
proof -
  have "k. project_transition_fst (ρ k) = ρ1 k"
    and "k. project_transition_snd (ρ k) = ρ2 k"
    unfolding assms product_run_project_fst product_run_project_snd by simp+
  hence "project_transition_fst ` range ρ = range ρ1"
    and "project_transition_snd ` range ρ = range ρ2"
    using range_composition[symmetric, of project_transition_fst ρ]
    using range_composition[symmetric, of project_transition_snd ρ] by presburger+
  thus "finite (range ρ1)" and "finite (range ρ2)"
    using assms finite_imageI by metis+
qed

lemma 
  fixes δ1 δ2 w q1 q2
  defines "ρ  runt (δ1 × δ2) (q1, q2) w"
  defines "ρ1  runt δ1 q1 w"
  assumes "finite (range ρ)"
  shows product_run_project_limit_fst: "project_transition_fst ` limit ρ = limit ρ1"
    and product_run_embed_limit_fst: "limit ρ   (embed_transition_fst ` (limit ρ1))"
proof -
  have "finite (range ρ1)"
    using assms product_run_finite_fst by metis

  then obtain i where "limit ρ = range (suffix i ρ)" and "limit ρ1 = range (suffix i ρ1)"
    using common_range_limit assms by metis
  moreover
  have "k. project_transition_fst (suffix i ρ k) = (suffix i ρ1 k)"
    by (simp only: assms runt.simps) (metis ρ1_def product_run_project_fst suffix_nth) 
  hence "project_transition_fst ` range (suffix i ρ) = (range (suffix i ρ1))"
    using range_composition[symmetric, of "project_transition_fst" "suffix i ρ"] by presburger
  moreover
  have "k. (suffix i ρ k)  embed_transition_fst (suffix i ρ1 k)"
    using assms product_run_embed_fst by simp
  ultimately
  show "project_transition_fst ` limit ρ = limit ρ1" 
    and "limit ρ   (embed_transition_fst ` (limit ρ1))"
    by auto
qed 

lemma 
  fixes δ1 δ2 w q1 q2
  defines "ρ  runt (δ1 × δ2) (q1, q2) w"
  defines "ρ2  runt δ2 q2 w"
  assumes "finite (range ρ)"
  shows product_run_project_limit_snd: "project_transition_snd ` limit ρ = limit ρ2"
    and product_run_embed_limit_snd: "limit ρ   (embed_transition_snd ` (limit ρ2))"
proof -
  have "finite (range ρ2)"
    using assms product_run_finite_snd by metis

  then obtain i where "limit ρ = range (suffix i ρ)" and "limit ρ2 = range (suffix i ρ2)"
    using common_range_limit assms by metis
  moreover
  have "k. project_transition_snd (suffix i ρ k) = (suffix i ρ2 k)"
    by (simp only: assms runt.simps) (metis ρ2_def product_run_project_snd suffix_nth) 
  hence "project_transition_snd ` range ((suffix i ρ)) = (range (suffix i ρ2))"
    using range_composition[symmetric, of "project_transition_snd" "(suffix i ρ)"] by presburger
  moreover
  have "k. (suffix i ρ k)  embed_transition_snd (suffix i ρ2 k)"
    using assms product_run_embed_snd by simp
  ultimately
  show "project_transition_snd ` limit ρ = limit ρ2" 
    and "limit ρ   (embed_transition_snd ` (limit ρ2))"
    by auto
qed 

lemma 
  fixes δ1 δ2 w q1 q2
  defines "ρ  runt (δ1 × δ2) (q1, q2) w"
  defines "ρ1  runt δ1 q1 w"
  defines "ρ2  runt δ2 q2 w"
  assumes "finite (range ρ)"
  shows product_run_embed_limit_finiteness_fst: "limit ρ  ( (embed_transition_fst ` S)) = {}  limit ρ1  S = {}" (is "?thesis1")
    and product_run_embed_limit_finiteness_snd: "limit ρ  ( (embed_transition_snd ` S')) = {}  limit ρ2  S' = {}" (is "?thesis2")
proof -
  show ?thesis1
    using assms product_run_project_limit_fst by fastforce
  show ?thesis2
    using assms product_run_project_limit_snd by fastforce
qed

subsection ‹Product Construction Helper Functions and Lemmas›

fun embed_transition :: "'a  ('b, 'c) transition  ('a  'b, 'c) transition set" ("↿⇩_")
where
  "↿⇩x (q, ν, q') = {(m, ν, m') | m m'. m x = Some q  m' x = Some q'}"

fun project_transition :: "'a  ('a  'b, 'c) transition  ('b, 'c) transition" ("⇃⇩_")
where
  "⇃⇩x (m, ν, m') = (the (m x), ν, the (m' x))"

fun embed_pair :: "'a  (('b, 'c) transition set × ('b, 'c) transition set)  (('a  'b, 'c) transition set × ('a  'b, 'c) transition set)" ("_")
where
  "x (S, S') = ((↿⇩x ` S), (↿⇩x ` S'))" 

fun project_pair :: "'a  (('a  'b, 'c) transition set × ('a  'b, 'c) transition set)  (('b, 'c) transition set × ('b, 'c) transition set)" ("_")
where
  "x (S, S') = (⇃⇩x ` S, ⇃⇩x ` S')" 

lemma embed_transition_unfold:
  "embed_transition x t = {(m, fst (snd t), m') | m m'. m x = Some (fst t)  m' x = Some (snd (snd t))}"
  unfolding embed_transition.simps[symmetric] by simp

lemma
  fixes ιm δm w q0 
  fixes x :: "'a"
  defines "ρ  runt (Δ× δm) ιm w"
  defines "ρ'  runt (δm x) q0 w"
  assumes "ιm x = Some q0"
  shows product_run_project: "⇃⇩x (ρ i) = ρ' i" 
    and product_run_embed: "ρ i  ↿⇩x (ρ' i)"
  using assms product_run_Some[of _ _ _ δm] by simp+

lemma 
  fixes ιm δm w q0 x
  defines "ρ  runt (Δ× δm) ιm w"
  defines "ρ'  runt (δm x) q0 w"
  assumes "ιm x = Some q0"
  assumes "finite (range ρ)"
  shows product_run_project_limit: "⇃⇩x ` limit ρ = limit ρ'" 
    and product_run_embed_limit: "limit ρ   (↿⇩x ` (limit ρ'))"
proof -
  have "k. ⇃⇩x (ρ k) = ρ' k"
    using assms product_run_embed[of _ _ _ δm] by simp
  hence "⇃⇩x ` range ρ = range ρ'"
    using range_composition[symmetric, of "⇃⇩x" ρ] by presburger
  hence "finite (range ρ')"
    using assms finite_imageI by metis

  then obtain i where "limit ρ = range (suffix i ρ)" and "limit ρ' = range (suffix i ρ')"
    using common_range_limit assms by metis
  moreover  
  have "k. ⇃⇩x (suffix i ρ k) = (suffix i ρ' k)"
    using assms product_run_embed[of _ _ _ δm] by simp
  hence "⇃⇩x ` range ((suffix i ρ)) = (range (suffix i ρ'))"
    using range_composition[symmetric, of "⇃⇩x" "(suffix i  ρ)"] by presburger
  moreover
  have "k. (suffix i ρ k)  ↿⇩x (suffix i ρ' k)"
    using assms product_run_embed[of _ _ _ δm] by simp
  ultimately
  show "⇃⇩x ` limit ρ = limit ρ'" and "limit ρ   (↿⇩x ` (limit ρ'))"
    by auto
qed

lemma product_run_embed_limit_finiteness:
  fixes ιm δm w q0 k
  defines "ρ  runt (Δ× δm) ιm w"
  defines "ρ'  runt (δm k) q0 w"
  assumes "ιm k = Some q0"
  assumes "finite (range ρ)"
  shows "limit ρ  ( (↿⇩k ` S)) = {}  limit ρ'  S = {}"
  (is "?lhs  ?rhs")
proof -
  have "⇃⇩k ` limit ρ  S  {}  limit ρ  ( (↿⇩k ` S))  {}"
  proof
    assume "⇃⇩k ` limit ρ  S  {}"
    then obtain q ν q' where "(q, ν, q')  ⇃⇩k ` limit ρ" and "(q, ν, q')  S"
      by auto
    moreover
    have "m ν m' i. (m, ν, m') = ρ i  p p'. m k = Some p  m' k = Some p'" 
      using assms product_run_Some[of ιm , OF assms(3)] by auto
    hence "m ν m'. (m, ν, m')  limit ρ  p p'. m k = Some p  m' k = Some p'" 
      using limit_in_range by fast
    ultimately
    obtain m m' where "m k = Some q" and "m' k = Some q'" and "(m, ν, m')  limit ρ"
      by auto
    moreover
    hence "(m, ν, m')   (↿⇩k ` S)"
      using (q, ν, q')  S by force
    ultimately
    show "limit ρ  ( (↿⇩k ` S))  {}"
      by blast
  qed
  hence "?lhs  ⇃⇩k ` limit ρ  S = {}"
    by auto
  also
  have "  ?rhs"
    using assms product_run_project_limit[of _ _ _ δm] by simp
  finally
  show ?thesis
    by simp
qed

subsection ‹Transfer Rules›

context includes lifting_syntax
begin

lemma product_parametric [transfer_rule]:
  "((A ===> B ===> C ===> B) ===> (A ===> rel_option B) ===> C ===> A ===> rel_option B) product product"
  by (auto simp add: rel_fun_def rel_option_iff split: option.split)

lemma run_parametric [transfer_rule]:
  "((A ===> B ===> A) ===> A ===> ((=) ===> B) ===> (=) ===> A) run run"
proof -
  {
    fix δ δ' q q' n w 
    fix w' :: "nat  'd" 
    assume "(A ===> B ===> A) δ δ'" "A q q'" "((=) ===> B) w w'" 
    hence "A (run δ q w n) (run δ' q' w' n)"
      by (induction n) (simp_all add: rel_fun_def)
  }
  thus ?thesis
    by blast
qed

lemma reach_parametric [transfer_rule]:
  assumes "bi_total B"
  assumes "bi_unique B"
  shows "(rel_set B ===> (A ===> B ===> A) ===> A ===> rel_set A) reach reach"
proof standard+
  fix Σ Σ' δ δ' q q'
  assume "rel_set B Σ Σ'" "(A ===> B ===> A) δ δ'" "A q q'"

  {
    fix z 
    assume "z  reach Σ δ q"
    then obtain w n where "z = run δ q w n" and "range w  Σ"
      unfolding reach_def by auto
    
    define w' where "w' n = (SOME x. B (w n) x)" for n
    
    have "n. w n  Σ"
      using range w  Σ by blast
    hence "n. w' n  Σ'"
      using assms rel_set B Σ Σ' by (simp add: w'_def bi_unique_def rel_set_def; metis someI) 
    hence "run δ' q' w' n  reach Σ' δ' q'"
      unfolding reach_def by auto
     
    moreover

    have "A z (run δ' q' w' n)"
      apply (unfold z = run δ q w n)
      apply (insert A q q' (A ===> B ===> A) δ δ' assms(1))  
      apply (induction n) 
      apply (simp_all add: rel_fun_def bi_total_def w'_def) 
      by (metis tfl_some)  

    ultimately

    have "z'  reach Σ' δ' q'. A z z'"
      by blast
  }

  moreover
  
  {
    fix z 
    assume "z  reach Σ' δ' q'"
    then obtain w n where "z = run δ' q' w n" and "range w  Σ'"
      unfolding reach_def by auto
    
    define w' where "w' n = (SOME x. B x (w n))" for n
    
    have "n. w n  Σ'"
      using range w  Σ' by blast
    hence "n. w' n  Σ"
      using assms rel_set B Σ Σ' by (simp add: w'_def bi_unique_def rel_set_def; metis someI)
    hence "run δ q w' n  reach Σ δ q"
      unfolding reach_def by auto
     
    moreover

    have "A (run δ q w' n) z"
      apply (unfold z = run δ' q' w n)
      apply (insert A q q' (A ===> B ===> A) δ δ' assms(1))  
      apply (induction n) 
      apply (simp_all add: rel_fun_def bi_total_def w'_def) 
      by (metis tfl_some)  

    ultimately

    have "z'  reach Σ δ q. A z' z"
      by blast
  }
  ultimately
  show "rel_set A (reach Σ δ q) (reach Σ' δ' q')"
    unfolding rel_set_def by blast
qed

end

subsection ‹Lift to Mapping›

lift_definition product_abs :: "('a  ('b, 'c) DTS)  (('a, 'b) mapping, 'c) DTS" ("↑Δ×") is product 
  parametric product_parametric .

lemma product_abs_run_None:
  "Mapping.lookup ιm k = None  Mapping.lookup (run (↑Δ× δm) ιm w i) k = None"
  by (transfer; insert product_run_None)

lemma product_abs_run_Some:
  "Mapping.lookup ιm k = Some q0  Mapping.lookup (run (↑Δ× δm) ιm w i) k = Some (run (δm k) q0 w i)"
  by (transfer; insert product_run_Some)

theorem finite_reach_product_abs:
  assumes "finite (Mapping.keys ιm)"
  assumes "x. x  (Mapping.keys ιm)  finite (reach Σ (δm x) (the (Mapping.lookup ιm x)))"
  shows "finite (reach Σ (↑Δ× δm) ιm)"
  using assms by (transfer; blast intro: finite_reach_product)

end