diff --git a/thys/Buchi_Complementation/Complementation.thy b/thys/Buchi_Complementation/Complementation.thy --- a/thys/Buchi_Complementation/Complementation.thy +++ b/thys/Buchi_Complementation/Complementation.thy @@ -1,608 +1,608 @@ section \Complementation\ theory Complementation imports "Transition_Systems_and_Automata.Maps" "Ranking" begin subsection \Level Rankings and Complementation States\ type_synonym 'state lr = "'state \ nat" definition lr_succ :: "('label, 'state) nba \ 'label \ 'state lr \ 'state lr set" where "lr_succ A a f \ {g. - dom g = \((succ A a) ` (dom f)) \ - (\ p \ dom f. \ q \ succ A a p. the (g q) \ the (f p)) \ + dom g = \ (transition A a ` dom f) \ + (\ p \ dom f. \ q \ transition A a p. the (g q) \ the (f p)) \ (\ q \ dom g. accepting A q \ even (the (g q)))}" type_synonym 'state st = "'state set" definition st_succ :: "('label, 'state) nba \ 'label \ 'state lr \ 'state st \ 'state st" where - "st_succ A a g P \ {q \ if P = {} then dom g else \((succ A a) ` P). even (the (g q))}" + "st_succ A a g P \ {q \ if P = {} then dom g else \ (transition A a ` P). even (the (g q))}" type_synonym 'state cs = "'state lr \ 'state st" definition complement_succ :: "('label, 'state) nba \ 'label \ 'state cs \ 'state cs set" where "complement_succ A a \ \ (f, P). {(g, st_succ A a g P) |g. g \ lr_succ A a f}" definition complement :: "('label, 'state) nba \ ('label, 'state cs) nba" where "complement A \ nba (alphabet A) ({const (Some (2 * card (nodes A))) |` initial A} \ {{}}) (complement_succ A) (\ (f, P). P = {})" lemma dom_nodes: assumes "fP \ nodes (complement A)" shows "dom (fst fP) \ nodes A" using assms unfolding complement_def complement_succ_def lr_succ_def by (induct) (auto, blast) lemma ran_nodes: assumes "fP \ nodes (complement A)" shows "ran (fst fP) \ {0 .. 2 * card (nodes A)}" using assms proof induct case (initial fP) show ?case using initial unfolding complement_def by (auto) (metis eq_refl option.inject ran_restrictD) next case (execute fP agQ) obtain f P where 1: "fP = (f, P)" by force have 2: "ran f \ {0 .. 2 * card (nodes A)}" using execute(2) unfolding 1 by auto obtain a g Q where 3: "agQ = (a, (g, Q))" using prod_cases3 by this - have 4: "p \ dom f \ q \ succ A a p \ the (g q) \ the (f p)" for p q + have 4: "p \ dom f \ q \ transition A a p \ the (g q) \ the (f p)" for p q using execute(3) unfolding 1 3 complement_def nba.simps complement_succ_def lr_succ_def by simp - have 8: "dom g = \((succ A a) ` (dom f))" + have 8: "dom g = \((transition A a) ` (dom f))" using execute(3) unfolding 1 3 complement_def nba.simps complement_succ_def lr_succ_def by simp show ?case unfolding 1 3 ran_def proof safe fix q k assume 5: "fst (snd (a, (g, Q))) q = Some k" have 6: "q \ dom g" using 5 by auto - obtain p where 7: "p \ dom f" "q \ succ A a p" using 6 unfolding 8 by auto + obtain p where 7: "p \ dom f" "q \ transition A a p" using 6 unfolding 8 by auto have "k = the (g q)" using 5 by auto also have "\ \ the (f p)" using 4 7 by this also have "\ \ 2 * card (nodes A)" using 2 7(1) by (simp add: domD ranI subset_eq) finally show "k \ {0 .. 2 * card (nodes A)}" by auto qed qed lemma states_nodes: assumes "fP \ nodes (complement A)" shows "snd fP \ nodes A" using assms proof induct case (initial fP) show ?case using initial unfolding complement_def by auto next case (execute fP agQ) obtain f P where 1: "fP = (f, P)" by force have 2: "P \ nodes A" using execute(2) unfolding 1 by auto obtain a g Q where 3: "agQ = (a, (g, Q))" using prod_cases3 by this have 11: "a \ alphabet A" using execute(3) unfolding 3 complement_def by auto have 10: "(g, Q) \ nodes (complement A)" using execute(1, 3) unfolding 1 3 by auto have 4: "dom g \ nodes A" using dom_nodes[OF 10] by simp - have 5: "\((succ A a) ` P) \ nodes A" using 2 11 by auto + have 5: "\ (transition A a ` P) \ nodes A" using 2 11 by auto have 6: "Q \ nodes A" using execute(3) unfolding 1 3 complement_def nba.simps complement_succ_def st_succ_def using 4 5 by (auto split: if_splits) show ?case using 6 unfolding 3 by auto qed theorem complement_finite: assumes "finite (nodes A)" shows "finite (nodes (complement A))" proof - let ?lrs = "{f. dom f \ nodes A \ ran f \ {0 .. 2 * card (nodes A)}}" have 1: "finite ?lrs" using finite_set_of_finite_maps' assms by auto let ?states = "Pow (nodes A)" have 2: "finite ?states" using assms by simp have "nodes (complement A) \ ?lrs \ ?states" by (force dest: dom_nodes ran_nodes states_nodes) also have "finite \" using 1 2 by simp finally show ?thesis by this qed lemma complement_trace_snth: assumes "run (complement A) (w ||| r) p" defines "m \ p ## trace (w ||| r) p" obtains "fst (m !! Suc k) \ lr_succ A (w !! k) (fst (m !! k))" "snd (m !! Suc k) = st_succ A (w !! k) (fst (m !! Suc k)) (snd (m !! k))" proof - have 1: "r !! k \ succ (complement A) (w !! k) (m !! k)" using nba.run_snth assms by force + have 1: "r !! k \ transition (complement A) (w !! k) (m !! k)" using nba.run_snth assms by force show "fst (m !! Suc k) \ lr_succ A (w !! k) (fst (m !! k))" using assms(2) 1 unfolding complement_def complement_succ_def trace_alt_def by auto show "snd (m !! Suc k) = st_succ A (w !! k) (fst (m !! Suc k)) (snd (m !! k))" using assms(2) 1 unfolding complement_def complement_succ_def trace_alt_def by auto qed subsection \Word in Complement Language Implies Ranking\ lemma complement_ranking: assumes "w \ language (complement A)" obtains f where "ranking A w f" proof - obtain r p where 1: "run (complement A) (w ||| r) p" "p \ initial (complement A)" "infs (accepting (complement A)) (trace (w ||| r) p)" using assms by rule let ?m = "p ## trace (w ||| r) p" obtain 100: "fst (?m !! Suc k) \ lr_succ A (w !! k) (fst (?m !! k))" "snd (?m !! Suc k) = st_succ A (w !! k) (fst (?m !! Suc k)) (snd (?m !! k))" for k using complement_trace_snth 1(1) by metis define f where "f \ \ (k, q). the (fst (?m !! k) q)" define P where "P k \ snd (?m !! k)" for k have 2: "snd v \ dom (fst (?m !! fst v))" if "v \ gunodes A w" for v using that proof induct case (initial v) then show ?case using 1(2) unfolding complement_def by auto next case (execute v u) - have "snd u \ \ ((succ A (w !! fst v)) ` (dom (fst (?m !! fst v))))" + have "snd u \ \ ((transition A (w !! fst v)) ` (dom (fst (?m !! fst v))))" using execute(2, 3) by auto also have "\ = dom (fst (?m !! Suc (fst v)))" using 100 unfolding lr_succ_def by simp also have "Suc (fst v) = fst u" using execute(3) by auto finally show ?case by this qed have 3: "f u \ f v" if 10: "v \ gunodes A w" and 11: "u \ gusuccessors A w v" for u v proof - - have 15: "snd u \ succ A (w !! fst v) (snd v)" using 11 by auto + have 15: "snd u \ transition A (w !! fst v) (snd v)" using 11 by auto have 16: "snd v \ dom (fst (?m !! fst v))" using 2 10 by this have "f u = the (fst (?m !! fst u) (snd u))" unfolding f_def by (simp add: case_prod_beta) also have "fst u = Suc (fst v)" using 11 by auto also have "the (fst (?m !! \) (snd u)) \ the (fst (?m !! fst v) (snd v))" using 100 15 16 unfolding lr_succ_def by auto also have "\ = f v" unfolding f_def by (simp add: case_prod_beta) finally show "f u \ f v" by this qed have 4: "\ l \ k. P l = {}" for k proof - have 15: "infs (\ (k, P). P = {}) ?m" using 1(3) unfolding complement_def by auto obtain l where 17: "l \ k" "snd (?m !! l) = {}" using 15 unfolding infs_snth by force have 19: "P l = {}" unfolding P_def using 17 by auto show ?thesis using 19 17(1) by auto qed show ?thesis proof (rule that, unfold ranking_def, intro conjI ballI impI allI) fix v assume "v \ gunodes A w" then show "f v \ 2 * card (nodes A)" proof induct case (initial v) then show ?case using 1(2) unfolding complement_def f_def by auto next case (execute v u) have "f u \ f v" using 3[OF execute(1)] execute(3) by simp also have "\ \ 2 * card (nodes A)" using execute(2) by this finally show ?case by this qed next fix v u assume 10: "v \ gunodes A w" assume 11: "u \ gusuccessors A w v" show "f u \ f v" using 3 10 11 by this next fix v assume 10: "v \ gunodes A w" assume 11: "gaccepting A v" show "even (f v)" using 10 proof cases case (initial) then show ?thesis using 1(2) unfolding complement_def f_def by auto next case (execute u) have 12: "snd v \ dom (fst (?m !! fst v))" using execute graph.nodes.execute 2 by blast have 12: "snd v \ dom (fst (?m !! Suc (fst u)))" using 12 execute(2) by auto have 13: "accepting A (snd v)" using 11 by auto have "f v = the (fst (?m !! fst v) (snd v))" unfolding f_def by (simp add: case_prod_beta) also have "fst v = Suc (fst u)" using execute(2) by auto also have "even (the (fst (?m !! Suc (fst u)) (snd v)))" using 100 12 13 unfolding lr_succ_def by simp finally show ?thesis by this qed next fix v s k assume 10: "v \ gunodes A w" assume 11: "gurun A w s v" assume 12: "smap f (gtrace s v) = sconst k" show "odd k" proof assume 13: "even k" obtain t u where 14: "u \ ginitial A" "gupath A w t u" "v = gtarget t u" using 10 by auto obtain l where 15: "l \ length t" "P l = {}" using 4 by auto obtain l' where 16: "l' \ Suc l" "P l' = {}" using 4 by auto have 30: "gurun A w (t @- s) u" using 11 14 by auto have 21: "fst (gtarget (stake (Suc l) (t @- s)) u) = Suc l" for l unfolding sscan_snth[symmetric] using 30 14(1) by (auto elim!: grun_elim) have 17: "snd (gtarget (stake (Suc l + i) (t @- s)) u) \ P (Suc l + i)" for i proof (induct i) case (0) have 20: "gtarget (stake (Suc l) (t @- s)) u \ gunodes A w" using 14 11 by (force simp add: 15(1) le_SucI graph.run_stake stake_shift) have "snd (gtarget (stake (Suc l) (t @- s)) u) \ dom (fst (?m !! fst (gtarget (stake (Suc l) (t @- s)) u)))" using 2[OF 20] by this also have "fst (gtarget (stake (Suc l) (t @- s)) u) = Suc l" using 21 by this finally have 22: "snd (gtarget (stake (Suc l) (t @- s)) u) \ dom (fst (?m !! Suc l))" by this have "gtarget (stake (Suc l) (t @- s)) u = gtrace (t @- s) u !! l" unfolding sscan_snth by rule also have "\ = gtrace s v !! (l - length t)" using 15(1) by simp also have "f \ = smap f (gtrace s v) !! (l - length t)" by simp also have "smap f (gtrace s v) = sconst k" unfolding 12 by rule also have "sconst k !! (l - length t) = k" by simp finally have 23: "even (f (gtarget (stake (Suc l) (t @- s)) u))" using 13 by simp have 24: "snd (gtarget (stake (Suc l) (t @- s)) u) \ {p \ dom (fst (?m !! Suc l)). even (f (Suc l, p))}" using 21 22 23 by (metis (no_types, lifting) mem_Collect_eq prod.collapse) also have "\ = st_succ A (w !! l) (fst (?m !! Suc l)) (P l)" unfolding 15(2) st_succ_def f_def by simp also have "\ = P (Suc l)" using 100(2) unfolding P_def by rule also have "P (Suc l) = P (l + Suc 0)" by simp finally show ?case by auto next case (Suc i) have 20: "P (Suc l + i) \ {}" using Suc by auto have 21: "fst (gtarget (stake (Suc l + Suc i) (t @- s)) u) = Suc l + Suc i" using 21 by (simp add: stake_shift) have "gtarget (stake (Suc l + Suc i) (t @- s)) u = gtrace (t @- s) u !! (l + Suc i)" unfolding sscan_snth by simp also have "\ \ gusuccessors A w (gtarget (stake (Suc (l + i)) (t @- s)) u)" using graph.run_snth[OF 30, of "l + Suc i"] by simp finally have 220: "snd (gtarget (stake (Suc (Suc l + i)) (t @- s)) u) \ - succ A (w !! (Suc l + i)) (snd (gtarget (stake (Suc (l + i)) (t @- s)) u))" + transition A (w !! (Suc l + i)) (snd (gtarget (stake (Suc (l + i)) (t @- s)) u))" using 21 by auto have 22: "snd (gtarget (stake (Suc l + Suc i) (t @- s)) u) \ - \ (succ A (w !! (Suc l + i)) ` P (Suc l + i))" using 220 Suc by auto + \ (transition A (w !! (Suc l + i)) ` P (Suc l + i))" using 220 Suc by auto have "gtarget (stake (Suc l + Suc i) (t @- s)) u = gtrace (t @- s) u !! (l + Suc i)" unfolding sscan_snth by simp also have "\ = gtrace s v !! (l + Suc i - length t)" using 15(1) by (metis add.commute shift_snth_ge sscan_const trans_le_add2) also have "f \ = smap f (gtrace s v) !! (l + Suc i - length t)" by simp also have "smap f (gtrace s v) = sconst k" unfolding 12 by rule also have "sconst k !! (l + Suc i - length t) = k" by simp finally have 23: "even (f (gtarget (stake (Suc l + Suc i) (t @- s)) u))" using 13 by auto have "snd (gtarget (stake (Suc l + Suc i) (t @- s)) u) \ - {p \ \ (succ A (w !! (Suc l + i)) ` P (Suc l + i)). even (f (Suc (Suc l + i), p))}" + {p \ \ (transition A (w !! (Suc l + i)) ` P (Suc l + i)). even (f (Suc (Suc l + i), p))}" using 21 22 23 by (metis (mono_tags) add_Suc_right mem_Collect_eq prod.collapse) also have "\ = st_succ A (w !! (Suc l + i)) (fst (?m !! Suc (Suc l + i))) (P (Suc l + i))" unfolding st_succ_def f_def using 20 by simp also have "\ = P (Suc (Suc l + i))" unfolding 100(2)[folded P_def] by rule also have "\ = P (Suc l + Suc i)" by simp finally show ?case by this qed show "False" using 16 17 using nat_le_iff_add by auto qed qed qed subsection \Ranking Implies Word in Complement Language\ definition reach where "reach A w i \ {target r p |r p. path A r p \ p \ initial A \ map fst r = stake i w}" lemma reach_0[simp]: "reach A w 0 = initial A" unfolding reach_def by auto lemma reach_Suc_empty: assumes "w !! n \ alphabet A" shows "reach A w (Suc n) = {}" proof safe fix q assume 1: "q \ reach A w (Suc n)" obtain r p where 2: "q = target r p" "path A r p" "p \ initial A" "map fst r = stake (Suc n) w" using 1 unfolding reach_def by blast have 3: "path A (take n r @ drop n r) p" using 2(2) by simp have 4: "map fst r = stake n w @ [w !! n]" using 2(4) stake_Suc by auto have 5: "map snd r = take n (map snd r) @ [q]" using 2(1, 4) 4 by (metis One_nat_def Suc_inject Suc_neq_Zero Suc_pred append.right_neutral append_eq_conv_conj drop_map id_take_nth_drop last_ConsR last_conv_nth length_0_conv length_map length_stake lessI nba.target_alt_def states_alt_def zero_less_Suc) have 6: "drop n r = [(w !! n, q)]" using 4 5 by (metis append_eq_conv_conj append_is_Nil_conv append_take_drop_id drop_map length_greater_0_conv length_stake stake_cycle_le stake_invert_Nil take_map zip_Cons_Cons zip_map_fst_snd) show "q \ {}" using assms 3 unfolding 6 by auto qed lemma reach_Suc_succ: assumes "w !! n \ alphabet A" - shows "reach A w (Suc n) = \((succ A (w !! n) ` (reach A w n)))" + shows "reach A w (Suc n) = \((transition A (w !! n) ` (reach A w n)))" proof safe fix q assume 1: "q \ reach A w (Suc n)" obtain r p where 2: "q = target r p" "path A r p" "p \ initial A" "map fst r = stake (Suc n) w" using 1 unfolding reach_def by blast have 3: "path A (take n r @ drop n r) p" using 2(2) by simp have 4: "map fst r = stake n w @ [w !! n]" using 2(4) stake_Suc by auto have 5: "map snd r = take n (map snd r) @ [q]" using 2(1, 4) 4 by (metis One_nat_def Suc_inject Suc_neq_Zero Suc_pred append.right_neutral append_eq_conv_conj drop_map id_take_nth_drop last_ConsR last_conv_nth length_0_conv length_map length_stake lessI nba.target_alt_def states_alt_def zero_less_Suc) have 6: "drop n r = [(w !! n, q)]" using 4 5 by (metis append_eq_conv_conj append_is_Nil_conv append_take_drop_id drop_map length_greater_0_conv length_stake stake_cycle_le stake_invert_Nil take_map zip_Cons_Cons zip_map_fst_snd) - show "q \ \((succ A (w !! n) ` (reach A w n)))" + show "q \ \((transition A (w !! n) ` (reach A w n)))" unfolding reach_def proof (intro UN_I CollectI exI conjI) show "target (take n r) p = target (take n r) p" by rule show "path A (take n r) p" using 3 by blast show "p \ initial A" using 2(3) by this show "map fst (take n r) = stake n w" using 2 by (metis length_stake lessI nat.distinct(1) stake_cycle_le stake_invert_Nil take_map take_stake) - show "q \ succ A (w !! n) (target (take n r) p)" using 3 unfolding 6 by auto + show "q \ transition A (w !! n) (target (take n r) p)" using 3 unfolding 6 by auto qed next fix p q - assume 1: "p \ reach A w n" "q \ succ A (w !! n) p" + assume 1: "p \ reach A w n" "q \ transition A (w !! n) p" obtain r x where 2: "p = target r x" "path A r x" "x \ initial A" "map fst r = stake n w" using 1(1) unfolding reach_def by blast show "q \ reach A w (Suc n)" unfolding reach_def proof (intro CollectI exI conjI) show "q = target (r @ [(w !! n, q)]) x" using 1 2 by auto show "path A (r @ [(w !! n, q)]) x" using assms 1(2) 2(1, 2) by auto show "x \ initial A" using 2(3) by this show "map fst (r @ [(w !! n, q)]) = stake (Suc n) w" using 1 2 by (metis eq_fst_iff list.simps(8) list.simps(9) map_append stake_Suc) qed qed lemma reach_Suc[simp]: "reach A w (Suc n) = (if w !! n \ alphabet A - then \((succ A (w !! n) ` (reach A w n))) else {})" + then \((transition A (w !! n) ` (reach A w n))) else {})" using reach_Suc_empty reach_Suc_succ by metis lemma reach_nodes: "reach A w i \ nodes A" by (induct i) (auto) lemma reach_gunodes: "{i} \ reach A w i \ gunodes A w" by (induct i) (auto intro: graph.nodes.execute) lemma ranking_complement: assumes "finite (nodes A)" "w \ streams (alphabet A)" "ranking A w f" shows "w \ language (complement A)" proof - define f' where "f' \ \ (k, p). if k = 0 then 2 * card (nodes A) else f (k, p)" have 0: "ranking A w f'" unfolding ranking_def proof (intro conjI ballI impI allI) show "\ v. v \ gunodes A w \ f' v \ 2 * card (nodes A)" using assms(3) unfolding ranking_def f'_def by auto show "\ v u. v \ gunodes A w \ u \ gusuccessors A w v \ f' u \ f' v" using assms(3) unfolding ranking_def f'_def by fastforce show "\ v. v \ gunodes A w \ gaccepting A v \ even (f' v)" using assms(3) unfolding ranking_def f'_def by auto next have 1: "v \ gunodes A w \ gurun A w r v \ smap f (gtrace r v) = sconst k \ odd k" for v r k using assms(3) unfolding ranking_def by meson fix v r k assume 2: "v \ gunodes A w" "gurun A w r v" "smap f' (gtrace r v) = sconst k" have 20: "shd r \ gureachable A w v" using 2 by (auto) (metis graph.reachable.reflexive graph.reachable_trace gtrace_alt_def subsetD shd_sset) obtain 3: "shd r \ gunodes A w" "gurun A w (stl r) (shd r)" "smap f' (gtrace (stl r) (shd r)) = sconst k" using 2 20 by (metis (no_types, lifting) eq_id_iff graph.nodes_trans graph.run_scons_elim siterate.simps(2) sscan.simps(2) stream.collapse stream.map_sel(2)) have 4: "k \ 0" if "(k, p) \ sset r" for k p proof - obtain ra ka pa where 1: "r = fromN (Suc ka) ||| ra" "v = (ka, pa)" using grun_elim[OF 2(2)] by this have 2: "k \ sset (fromN (Suc ka))" using 1(1) that by (metis image_eqI prod.sel(1) szip_smap_fst stream.set_map) show ?thesis using 2 by simp qed have 5: "smap f' (gtrace (stl r) (shd r)) = smap f (gtrace (stl r) (shd r))" proof (rule stream.map_cong) show "gtrace (stl r) (shd r) = gtrace (stl r) (shd r)" by rule next fix z assume 1: "z \ sset (gtrace (stl r) (shd r))" have 2: "fst z \ 0" using 4 1 by (metis gtrace_alt_def prod.collapse stl_sset) show "f' z = f z" using 2 unfolding f'_def by (auto simp: case_prod_beta) qed show "odd k" using 1 3 5 by simp qed define g where "g i p \ if p \ reach A w i then Some (f' (i, p)) else None" for i p have g_dom[simp]: "dom (g i) = reach A w i" for i unfolding g_def by (auto) (metis option.simps(3)) have g_0[simp]: "g 0 = const (Some (2 * card (nodes A))) |` initial A" unfolding g_def f'_def by auto have g_Suc[simp]: "g (Suc n) \ lr_succ A (w !! n) (g n)" for n unfolding lr_succ_def proof (intro CollectI conjI ballI impI) - show "dom (g (Suc n)) = \ (succ A (w !! n) ` dom (g n))" using snth_in assms(2) by auto + show "dom (g (Suc n)) = \ (transition A (w !! n) ` dom (g n))" using snth_in assms(2) by auto next fix p q - assume 100: "p \ dom (g n)" "q \ succ A (w !! n) p" + assume 100: "p \ dom (g n)" "q \ transition A (w !! n) p" have 101: "q \ reach A w (Suc n)" using snth_in assms(2) 100 by auto have 102: "(n, p) \ gunodes A w" using 100(1) reach_gunodes g_dom by blast have 103: "(Suc n, q) \ gusuccessors A w (n, p)" using snth_in assms(2) 102 100(2) by auto have 104: "p \ reach A w n" using 100(1) by simp have "g (Suc n) q = Some (f' (Suc n, q))" using 101 unfolding g_def by simp also have "the \ = f' (Suc n, q)" by simp also have "\ \ f' (n, p)" using 0 unfolding ranking_def using 102 103 by simp also have "\ = the (Some (f' (n, p)))" by simp also have "Some (f' (n, p)) = g n p" using 104 unfolding g_def by simp finally show "the (g (Suc n) q) \ the (g n p)" by this next fix p assume 100: "p \ dom (g (Suc n))" "accepting A p" have 101: "p \ reach A w (Suc n)" using 100(1) by simp have 102: "(Suc n, p) \ gunodes A w" using 101 reach_gunodes by blast have 103: "gaccepting A (Suc n, p)" using 100(2) by simp have "the (g (Suc n) p) = f' (Suc n, p)" using 101 unfolding g_def by simp also have "even \" using 0 unfolding ranking_def using 102 103 by auto finally show "even (the (g (Suc n) p))" by this qed define P where "P \ rec_nat {} (\ n. st_succ A (w !! n) (g (Suc n)))" have P_0[simp]: "P 0 = {}" unfolding P_def by simp have P_Suc[simp]: "P (Suc n) = st_succ A (w !! n) (g (Suc n)) (P n)" for n unfolding P_def by simp have P_reach: "P n \ reach A w n" for n using snth_in assms(2) by (induct n) (auto simp add: st_succ_def) have "P n \ reach A w n" for n using P_reach by auto also have "\ n \ nodes A" for n using reach_nodes by this also have "finite (nodes A)" using assms(1) by this finally have P_finite: "finite (P n)" for n by this define s where "s \ smap g nats ||| smap P nats" show ?thesis proof show "run (complement A) (w ||| stl s) (shd s)" proof (intro nba.snth_run conjI, simp_all del: stake.simps stake_szip) fix k show "w !! k \ alphabet (complement A)" using snth_in assms(2) unfolding complement_def by auto have "stl s !! k = s !! Suc k" by simp also have "\ \ complement_succ A (w !! k) (s !! k)" unfolding complement_succ_def s_def using P_Suc by simp also have "\ = complement_succ A (w !! k) (target (stake k (w ||| stl s)) (shd s))" unfolding sscan_scons_snth[symmetric] trace_alt_def by simp - also have "\ = succ (complement A) (w !! k) (target (stake k (w ||| stl s)) (shd s))" + also have "\ = transition (complement A) (w !! k) (target (stake k (w ||| stl s)) (shd s))" unfolding complement_def nba.sel by rule finally show "stl s !! k \ - succ (complement A) (w !! k) (target (stake k (w ||| stl s)) (shd s))" by this + transition (complement A) (w !! k) (target (stake k (w ||| stl s)) (shd s))" by this qed show "shd s \ initial (complement A)" unfolding complement_def s_def using P_0 by simp show "infs (accepting (complement A)) (trace (w ||| stl s) (shd s))" proof - have 10: "\ n. \ k \ n. P k = {}" proof (rule ccontr) assume 20: "\ (\ n. \ k \ n. P k = {})" obtain k where 22: "P (k + n) \ {}" for n using 20 using le_add1 by blast - define m where "m n S \ {p \ \((succ A (w !! n) ` S)). even (the (g (Suc n) p))}" for n S + define m where "m n S \ {p \ \((transition A (w !! n) ` S)). even (the (g (Suc n) p))}" for n S define R where "R i n S \ rec_nat S (\ i. m (n + i)) i" for i n S have R_0[simp]: "R 0 n = id" for n unfolding R_def by auto have R_Suc[simp]: "R (Suc i) n = m (n + i) \ R i n" for i n unfolding R_def by auto have R_Suc': "R (Suc i) n = R i (Suc n) \ m n" for i n unfolding R_Suc by (induct i) (auto) have R_reach: "R i n S \ reach A w (n + i)" if "S \ reach A w n" for i n S using snth_in assms(2) that m_def by (induct i) (auto) have P_R: "P (k + i) = R i k (P k)" for i using 22 by (induct i) (auto simp add: case_prod_beta' m_def st_succ_def) have 50: "R i n S = (\ p \ S. R i n {p})" for i n S by (induct i) (auto simp add: m_def prod.case_eq_if) have 51: "R (i + j) n S = {}" if "R i n S = {}" for i j n S using that by (induct j) (auto simp add: m_def prod.case_eq_if) have 52: "R j n S = {}" if "i \ j" "R i n S = {}" for i j n S using 51 by (metis le_add_diff_inverse that(1) that(2)) have 1: "\ p \ S. \ i. R i n {p} \ {}" if assms: "finite S" "\ i. R i n S \ {}" for n S proof (rule ccontr) assume 1: "\ (\ p \ S. \ i. R i n {p} \ {})" obtain f where 3: "\ p. p \ S \ R (f p) n {p} = {}" using 1 by metis have 4: "R (Sup (f ` S)) n {p} = {}" if "p \ S" for p proof (rule 52) show "f p \ Sup (f ` S)" using assms(1) that by (auto intro: le_cSup_finite) show "R (f p) n {p} = {}" using 3 that by this qed have "R (Sup (f ` S)) n S = (\ p \ S. R (Sup (f ` S)) n {p})" using 50 by this also have "\ = {}" using 4 by simp finally have 5: "R (Sup (f ` S)) n S = {}" by this show "False" using that(2) 5 by auto qed have 2: "\ i. R i (k + 0) (P k) \ {}" using 22 P_R by simp obtain p where 3: "p \ P k" "\ i. R i k {p} \ {}" using 1[OF P_finite 2] by auto define Q where "Q n p \ (\ i. R i (k + n) {p} \ {}) \ p \ P (k + n)" for n p - have 5: "\ q \ succ A (w !! (k + n)) p. Q (Suc n) q" if "Q n p" for n p + have 5: "\ q \ transition A (w !! (k + n)) p. Q (Suc n) q" if "Q n p" for n p proof - have 11: "p \ P (k + n)" "\ i. R i (k + n) {p} \ {}" using that unfolding Q_def by auto have 12: "R (Suc i) (k + n) {p} \ {}" for i using 11(2) by this have 13: "R i (k + Suc n) (m (k + n) {p}) \ {}" for i using 12 unfolding R_Suc' by simp have "{p} \ P (k + n)" using 11(1) by auto also have "\ \ reach A w (k + n)" using P_reach by this finally have "R 1 (k + n) {p} \ reach A w (k + n + 1)" using R_reach by blast also have "\ \ nodes A" using reach_nodes by this also have "finite (nodes A)" using assms(1) by this finally have 14: "finite (m (k + n) {p})" by simp obtain q where 14: "q \ m (k + n) {p}" "\ i. R i (k + Suc n) {q} \ {}" using 1[OF 14 13] by auto show ?thesis unfolding Q_def prod.case proof (intro bexI conjI allI) show "\ i. R i (k + Suc n) {q} \ {}" using 14(2) by this show "q \ P (k + Suc n)" using 14(1) 11(1) 22 unfolding m_def by (auto simp add: st_succ_def) - show "q \ succ A (w !! (k + n)) p" using 14(1) unfolding m_def by simp + show "q \ transition A (w !! (k + n)) p" using 14(1) unfolding m_def by simp qed qed obtain r where 23: "run A r p" "\ i. Q i ((p ## trace r p) !! i)" "\ i. fst (r !! i) = w !! (k + i)" proof (rule nba.invariant_run_index[of Q 0 p A "\ n p a. fst a = w !! (k + n)"]) show "Q 0 p" unfolding Q_def using 3 by auto - show "\ a. (fst a \ alphabet A \ snd a \ succ A (fst a) p) \ + show "\ a. (fst a \ alphabet A \ snd a \ transition A (fst a) p) \ Q (Suc n) (snd a) \ fst a = w !! (k + n)" if "Q n p" for n p using snth_in assms(2) 5 that by fastforce qed auto have 20: "smap fst r = sdrop k w" using 23(3) by (intro eqI_snth) (simp add: case_prod_beta) have 21: "(p ## smap snd r) !! i \ P (k + i)" for i using 23(2) unfolding Q_def unfolding trace_alt_def by simp obtain r where 23: "run A (sdrop k w ||| stl r) (shd r)" "\ i. r !! i \ P (k + i)" using 20 21 23(1) by (metis stream.sel(1) stream.sel(2) szip_smap) let ?v = "(k, shd r)" let ?r = "fromN (Suc k) ||| stl r" have "shd r = r !! 0" by simp also have "\ \ P k" using 23(2)[of 0] by simp also have "\ \ reach A w k" using P_reach by this finally have 24: "?v \ gunodes A w" using reach_gunodes by blast have 25: "gurun A w ?r ?v" using run_grun 23(1) by this obtain l where 26: "sset (smap f' (gtrace (sdrop l ?r) (gtarget (stake l ?r) ?v))) \ Collect odd" using ranking_stuck_odd 0 24 25 by this have 27: "f' (Suc (k + l), r !! Suc l) = shd (smap f' (gtrace (sdrop l ?r) (gtarget (stake l ?r) ?v)))" by (simp add: algebra_simps) also have "\ \ sset (smap f' (gtrace (sdrop l ?r) (gtarget (stake l ?r) ?v)))" using shd_sset by this also have "\ \ Collect odd" using 26 by this finally have 28: "odd (f' (Suc (k + l), r !! Suc l))" by simp have "r !! Suc l \ P (Suc (k + l))" using 23(2) by (metis add_Suc_right) - also have "\ = {p \ \ (succ A (w !! (k + l)) ` P (k + l)). + also have "\ = {p \ \ (transition A (w !! (k + l)) ` P (k + l)). even (the (g (Suc (k + l)) p))}" using 23(2) by (auto simp: st_succ_def) also have "\ \ {p. even (the (g (Suc (k + l)) p))}" by auto finally have 29: "even (the (g (Suc (k + l)) (r !! Suc l)))" by auto have 30: "r !! Suc l \ reach A w (Suc (k + l))" using 23(2) P_reach by (metis add_Suc_right subsetCE) have 31: "even (f' (Suc (k + l), r !! Suc l))" using 29 30 unfolding g_def by simp show "False" using 28 31 by simp qed have 11: "infs (\ k. P k = {}) nats" using 10 unfolding infs_snth by simp have "infs (\ S. S = {}) (smap snd (smap g nats ||| smap P nats))" using 11 by (simp add: comp_def) then have "infs (\ x. snd x = {}) (smap g nats ||| smap P nats)" by (simp add: comp_def del: szip_smap_snd) then have "infs (\ (f, P). P = {}) (smap g nats ||| smap P nats)" by (simp add: case_prod_beta') then have "infs (\ (f, P). P = {}) (stl (smap g nats ||| smap P nats))" by blast then have "infs (\ (f, P). P = {}) (smap snd (w ||| stl (smap g nats ||| smap P nats)))" by simp then have "infs (\ (f, P). P = {}) (trace (w ||| stl s) (shd s))" unfolding trace_alt_def s_def by this then show ?thesis unfolding complement_def by auto qed qed qed subsection \Correctness Theorem\ theorem complement_language: assumes "finite (nodes A)" shows "language (complement A) = streams (alphabet A) - language A" proof (safe del: notI) have 1: "alphabet (complement A) = alphabet A" unfolding complement_def nba.sel by rule show "w \ streams (alphabet A)" if "w \ language (complement A)" for w using language_alphabet that 1 by force show "w \ language A" if "w \ language (complement A)" for w using complement_ranking ranking_language that by metis show "w \ language (complement A)" if "w \ streams (alphabet A)" "w \ language A" for w using language_ranking ranking_complement assms that by blast qed end diff --git a/thys/Buchi_Complementation/Complementation_Final.thy b/thys/Buchi_Complementation/Complementation_Final.thy --- a/thys/Buchi_Complementation/Complementation_Final.thy +++ b/thys/Buchi_Complementation/Complementation_Final.thy @@ -1,89 +1,89 @@ section \Complementation to Explicit Büchi Automaton\ theory Complementation_Final imports "Complementation_Implement" "Transition_Systems_and_Automata.NBA_Translate" "HOL-Library.Permutation" begin definition "hci k \ uint32_of_nat k * 1103515245 + 12345" definition "hc \ \ (p, q, b). hci p + hci q * 31 + (if b then 1 else 0)" definition "list_hash xs \ fold (bitXOR \ hc) xs 0" lemma list_hash_eq: assumes "distinct xs" "distinct ys" "set xs = set ys" shows "list_hash xs = list_hash ys" proof - have "remdups xs <~~> remdups ys" using eq_set_perm_remdups assms(3) by this then have "xs <~~> ys" using assms(1, 2) by (simp add: distinct_remdups_id) then have "fold (bitXOR \ hc) xs a = fold (bitXOR \ hc) ys a" for a proof (induct arbitrary: a) case (swap y x l) have "x XOR y XOR a = y XOR x XOR a" for x y by (transfer) (simp add: word_bw_lcs(3)) then show ?case by simp qed simp+ then show ?thesis unfolding list_hash_def by this qed definition state_hash :: "nat \ Complementation_Implement.state \ nat" where "state_hash n p \ nat_of_hashcode (list_hash p) mod n" lemma state_hash_bounded_hashcode[autoref_ga_rules]: "is_bounded_hashcode state_rel (gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\))) state_hash" proof show [param]: "(gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\)), (=)) \ state_rel \ state_rel \ bool_rel" by autoref show "state_hash n xs = state_hash n ys" if "xs \ Domain state_rel" "ys \ Domain state_rel" "gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (=)) xs ys" for xs ys n proof - have 1: "distinct (map fst xs)" "distinct (map fst ys)" using that(1, 2) unfolding list_map_rel_def list_map_invar_def by (auto simp: in_br_conv) have 2: "distinct xs" "distinct ys" using 1 by (auto intro: distinct_mapI) have 3: "(xs, map_of xs) \ state_rel" "(ys, map_of ys) \ state_rel" using 1 unfolding list_map_rel_def list_map_invar_def by (auto simp: in_br_conv) have 4: "(gen_equals (Gen_Map.gen_ball (foldli \ list_map_to_list)) (list_map_lookup (=)) (prod_eq (=) (\)) xs ys, map_of xs = map_of ys) \ bool_rel" using 3 by parametricity have 5: "map_to_set (map_of xs) = map_to_set (map_of ys)" using that(3) 4 by simp have 6: "set xs = set ys" using map_to_set_map_of 1 5 by blast show "state_hash n xs = state_hash n ys" unfolding state_hash_def using list_hash_eq 2 6 by metis qed show "state_hash n x < n" if "1 < n" for n x using that unfolding state_hash_def by simp qed schematic_goal complement_impl: assumes [simp]: "finite (nodes A)" assumes [autoref_rules]: "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" shows "(?f :: ?'c, RETURN (to_nbaei (complement_4 A))) \ ?R" by (autoref_monadic (plain)) concrete_definition complement_impl uses complement_impl[unfolded autoref_tag_defs] theorem complement_impl_correct: assumes "finite (nodes A)" assumes "(Ai, A) \ \Id, nat_rel\ nbai_nba_rel" shows "language (nbae_nba (nbaei_nbae (complement_impl Ai))) = streams (alphabet A) - language A" proof - have "(language ((nbae_nba \ nbaei_nbae) (complement_impl Ai)), language (id (complement_4 A))) \ \\Id_on (alphabet (complement_4 A))\ stream_rel\ set_rel" using complement_impl.refine[OF assms, unfolded to_nbaei_def id_apply, THEN RETURN_nres_relD] by parametricity also have "language (id (complement_4 A)) = streams (alphabet A) - language A" using assms(1) complement_4_correct by simp finally show ?thesis by simp qed definition nbaei_nbai :: "(String.literal, nat) nbaei \ (String.literal, nat) nbai" where "nbaei_nbai \ nbae_nba_impl" (* TODO: possible optimizations: - introduce op_map_map operation for maps instead of manually iterating via FOREACH - consolidate various binds and maps in expand_map_get_7 *) export_code nat_of_integer integer_of_nat - nbaei alphabetei initialei transei acceptingei + nbaei alphabetei initialei transitionei acceptingei nbaei_nbai complement_impl in SML module_name Complementation file_prefix Complementation_Export end \ No newline at end of file diff --git a/thys/Buchi_Complementation/Complementation_Implement.thy b/thys/Buchi_Complementation/Complementation_Implement.thy --- a/thys/Buchi_Complementation/Complementation_Implement.thy +++ b/thys/Buchi_Complementation/Complementation_Implement.thy @@ -1,974 +1,974 @@ section \Complementation Implementation\ theory Complementation_Implement imports "HOL-Library.Lattice_Syntax" "Transition_Systems_and_Automata.NBA_Implement" "Complementation" begin type_synonym item = "nat \ bool" type_synonym 'state items = "'state \ item" type_synonym state = "(nat \ item) list" abbreviation "item_rel \ nat_rel \\<^sub>r bool_rel" abbreviation "state_rel \ \nat_rel, item_rel\ list_map_rel" - abbreviation "pred A a q \ {p. q \ succ A a p}" + abbreviation "pred A a q \ {p. q \ transition A a p}" subsection \Phase 1\ definition cs_lr :: "'state items \ 'state lr" where "cs_lr f \ map_option fst \ f" definition cs_st :: "'state items \ 'state st" where "cs_st f \ f -` Some ` snd -` {True}" abbreviation cs_abs :: "'state items \ 'state cs" where "cs_abs f \ (cs_lr f, cs_st f)" definition cs_rep :: "'state cs \ 'state items" where "cs_rep \ \ (g, P) p. map_option (\ k. (k, p \ P)) (g p)" lemma cs_abs_rep[simp]: "cs_rep (cs_abs f) = f" proof show "cs_rep (cs_abs f) x = f x" for x unfolding cs_lr_def cs_st_def cs_rep_def by (cases "f x") (force+) qed lemma cs_rep_lr[simp]: "cs_lr (cs_rep (g, P)) = g" proof show "cs_lr (cs_rep (g, P)) x = g x" for x unfolding cs_rep_def cs_lr_def by (cases "g x") (auto) qed lemma cs_rep_st[simp]: "cs_st (cs_rep (g, P)) = P \ dom g" unfolding cs_rep_def cs_st_def by force lemma cs_lr_dom[simp]: "dom (cs_lr f) = dom f" unfolding cs_lr_def by simp lemma cs_lr_apply[simp]: assumes "p \ dom f" shows "the (cs_lr f p) = fst (the (f p))" using assms unfolding cs_lr_def by auto lemma cs_rep_dom[simp]: "dom (cs_rep (g, P)) = dom g" unfolding cs_rep_def by auto lemma cs_rep_apply[simp]: assumes "p \ dom f" shows "fst (the (cs_rep (f, P) p)) = the (f p)" using assms unfolding cs_rep_def by auto abbreviation cs_rel :: "('state items \ 'state cs) set" where "cs_rel \ br cs_abs top" lemma cs_rel_inv_single_valued: "single_valued (cs_rel\)" by (auto intro!: inj_onI) (metis cs_abs_rep) definition refresh_1 :: "'state items \ 'state items" where "refresh_1 f \ if True \ snd ` ran f then f else map_option (apsnd top) \ f" definition ranks_1 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "ranks_1 A a f \ {g. - dom g = \((succ A a) ` (dom f)) \ - (\ p \ dom f. \ q \ succ A a p. fst (the (g q)) \ fst (the (f p))) \ + dom g = \((transition A a) ` (dom f)) \ + (\ p \ dom f. \ q \ transition A a p. fst (the (g q)) \ fst (the (f p))) \ (\ q \ dom g. accepting A q \ even (fst (the (g q)))) \ - cs_st g = {q \ \((succ A a) ` (cs_st f)). even (fst (the (g q)))}}" + cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}}" definition complement_succ_1 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_1 A a = ranks_1 A a \ refresh_1" definition complement_1 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_1 A \ nba (alphabet A) ({const (Some (2 * card (nodes A), False)) |` initial A}) (complement_succ_1 A) (\ f. cs_st f = {})" lemma refresh_1_dom[simp]: "dom (refresh_1 f) = dom f" unfolding refresh_1_def by simp lemma refresh_1_apply[simp]: "fst (the (refresh_1 f p)) = fst (the (f p))" unfolding refresh_1_def by (cases "f p") (auto) lemma refresh_1_cs_st[simp]: "cs_st (refresh_1 f) = (if cs_st f = {} then dom f else cs_st f)" unfolding refresh_1_def cs_st_def ran_def image_def vimage_def by auto lemma complement_succ_1_abs: assumes "g \ complement_succ_1 A a f" shows "cs_abs g \ complement_succ A a (cs_abs f)" unfolding complement_succ_def proof (simp, rule) have 1: - "dom g = \((succ A a) ` (dom f))" - "\ p \ dom f. \ q \ succ A a p. fst (the (g q)) \ fst (the (f p))" + "dom g = \((transition A a) ` (dom f))" + "\ p \ dom f. \ q \ transition A a p. fst (the (g q)) \ fst (the (f p))" "\ p \ dom g. accepting A p \ even (fst (the (g p)))" using assms unfolding complement_succ_1_def ranks_1_def by simp_all show "cs_lr g \ lr_succ A a (cs_lr f)" unfolding lr_succ_def proof (intro CollectI conjI ballI impI) - show "dom (cs_lr g) = \ (succ A a ` dom (cs_lr f))" using 1 by simp + show "dom (cs_lr g) = \ (transition A a ` dom (cs_lr f))" using 1 by simp next fix p q - assume 2: "p \ dom (cs_lr f)" "q \ succ A a p" + assume 2: "p \ dom (cs_lr f)" "q \ transition A a p" have 3: "q \ dom (cs_lr g)" using 1 2 by auto show "the (cs_lr g q) \ the (cs_lr f p)" using 1 2 3 by simp next fix p assume 2: "p \ dom (cs_lr g)" "accepting A p" show "even (the (cs_lr g p))" using 1 2 by auto qed - have 2: "cs_st g = {q \ \ (succ A a ` cs_st (refresh_1 f)). even (fst (the (g q)))}" + have 2: "cs_st g = {q \ \ (transition A a ` cs_st (refresh_1 f)). even (fst (the (g q)))}" using assms unfolding complement_succ_1_def ranks_1_def by simp show "cs_st g = st_succ A a (cs_lr g) (cs_st f)" proof (cases "cs_st f = {}") case True - have 3: "the (cs_lr g q) = fst (the (g q))" if "q \ \((succ A a) ` (dom f))" for q + have 3: "the (cs_lr g q) = fst (the (g q))" if "q \ \((transition A a) ` (dom f))" for q using that 1(1) by simp show ?thesis using 2 3 unfolding st_succ_def refresh_1_cs_st True cs_lr_dom 1(1) by force next case False - have 3: "the (cs_lr g q) = fst (the (g q))" if "q \ \((succ A a) ` (cs_st f))" for q + have 3: "the (cs_lr g q) = fst (the (g q))" if "q \ \((transition A a) ` (cs_st f))" for q using that 1(1) by (auto intro!: cs_lr_apply) (metis IntE UN_iff cs_abs_rep cs_lr_dom cs_rep_st domD prod.collapse) - have "cs_st g = {q \ \ (succ A a ` cs_st (refresh_1 f)). even (fst (the (g q)))}" + have "cs_st g = {q \ \ (transition A a ` cs_st (refresh_1 f)). even (fst (the (g q)))}" using 2 by this also have "cs_st (refresh_1 f) = cs_st f" using False by simp - also have "{q \ \((succ A a) ` (cs_st f)). even (fst (the (g q)))} = - {q \ \((succ A a) ` (cs_st f)). even (the (cs_lr g q))}" using 3 by metis + also have "{q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))} = + {q \ \((transition A a) ` (cs_st f)). even (the (cs_lr g q))}" using 3 by metis also have "\ = st_succ A a (cs_lr g) (cs_st f)" unfolding st_succ_def using False by simp finally show ?thesis by this qed qed lemma complement_succ_1_rep: assumes "P \ dom f" "(g, Q) \ complement_succ A a (f, P)" shows "cs_rep (g, Q) \ complement_succ_1 A a (cs_rep (f, P))" unfolding complement_succ_1_def ranks_1_def comp_apply proof (intro CollectI conjI ballI impI) have 1: - "dom g = \((succ A a) ` (dom f))" - "\ p \ dom f. \ q \ succ A a p. the (g q) \ the (f p)" + "dom g = \((transition A a) ` (dom f))" + "\ p \ dom f. \ q \ transition A a p. the (g q) \ the (f p)" "\ p \ dom g. accepting A p \ even (the (g p))" using assms(2) unfolding complement_succ_def lr_succ_def by simp_all - have 2: "Q = {q \ if P = {} then dom g else \((succ A a) ` P). even (the (g q))}" + have 2: "Q = {q \ if P = {} then dom g else \((transition A a) ` P). even (the (g q))}" using assms(2) unfolding complement_succ_def st_succ_def by simp have 3: "Q \ dom g" unfolding 2 1(1) using assms(1) by auto - show "dom (cs_rep (g, Q)) = \ (succ A a ` dom (refresh_1 (cs_rep (f, P))))" using 1 by simp - show "\ p q. p \ dom (refresh_1 (cs_rep (f, P))) \ q \ succ A a p \ + show "dom (cs_rep (g, Q)) = \ (transition A a ` dom (refresh_1 (cs_rep (f, P))))" using 1 by simp + show "\ p q. p \ dom (refresh_1 (cs_rep (f, P))) \ q \ transition A a p \ fst (the (cs_rep (g, Q) q)) \ fst (the (refresh_1 (cs_rep (f, P)) p))" using 1(1, 2) by (auto) (metis UN_I cs_rep_apply domI option.sel) show "\ p. p \ dom (cs_rep (g, Q)) \ accepting A p \ even (fst (the (cs_rep (g, Q) p)))" using 1(1, 3) by auto - show "cs_st (cs_rep (g, Q)) = {q \ \ (succ A a ` cs_st (refresh_1 (cs_rep (f, P)))). + show "cs_st (cs_rep (g, Q)) = {q \ \ (transition A a ` cs_st (refresh_1 (cs_rep (f, P)))). even (fst (the (cs_rep (g, Q) q)))}" proof (cases "P = {}") case True have "cs_st (cs_rep (g, Q)) = Q" using 3 by auto also have "\ = {q \ dom g. even (the (g q))}" unfolding 2 using True by auto also have "\ = {q \ dom g. even (fst (the (cs_rep (g, Q) q)))}" using cs_rep_apply by metis - also have "dom g = \((succ A a) ` (dom f))" using 1(1) by this + also have "dom g = \((transition A a) ` (dom f))" using 1(1) by this also have "dom f = cs_st (refresh_1 (cs_rep (f, P)))" using True by simp finally show ?thesis by this next case False - have 4: "fst (the (cs_rep (g, Q) q)) = the (g q)" if "q \ \((succ A a) ` P)" for q + have 4: "fst (the (cs_rep (g, Q) q)) = the (g q)" if "q \ \((transition A a) ` P)" for q using 1(1) that assms(1) by (fast intro: cs_rep_apply) have "cs_st (cs_rep (g, Q)) = Q" using 3 by auto - also have "\ = {q \ \((succ A a) ` P). even (the (g q))}" unfolding 2 using False by auto - also have "\ = {q \ \((succ A a) ` P). even (fst (the (cs_rep (g, Q) q)))}" using 4 by force + also have "\ = {q \ \((transition A a) ` P). even (the (g q))}" unfolding 2 using False by auto + also have "\ = {q \ \((transition A a) ` P). even (fst (the (cs_rep (g, Q) q)))}" using 4 by force also have "P = (cs_st (refresh_1 (cs_rep (f, P))))" using assms(1) False by auto finally show ?thesis by simp qed qed lemma complement_succ_1_refine: "(complement_succ_1, complement_succ) \ Id \ Id \ cs_rel \ \cs_rel\ set_rel" proof (clarsimp simp: br_set_rel_alt in_br_conv) fix A :: "('a, 'b) nba" fix a f show "complement_succ A a (cs_abs f) = cs_abs ` complement_succ_1 A a f" proof safe fix g Q assume 1: "(g, Q) \ complement_succ A a (cs_abs f)" have 2: "Q \ dom g" using 1 unfolding complement_succ_def lr_succ_def st_succ_def by (auto) (metis IntE cs_abs_rep cs_lr_dom cs_rep_st) have 3: "cs_st f \ dom (cs_lr f)" unfolding cs_st_def by auto show "(g, Q) \ cs_abs ` complement_succ_1 A a f" proof show "(g, Q) = cs_abs (cs_rep (g, Q))" using 2 by auto have "cs_rep (g, Q) \ complement_succ_1 A a (cs_rep (cs_abs f))" using complement_succ_1_rep 3 1 by this also have "cs_rep (cs_abs f) = f" by simp finally show "cs_rep (g, Q) \ complement_succ_1 A a f" by this qed next fix g assume 1: "g \ complement_succ_1 A a f" show "cs_abs g \ complement_succ A a (cs_abs f)" using complement_succ_1_abs 1 by this qed qed lemma complement_1_refine: "(complement_1, complement) \ \Id, Id\ nba_rel \ \Id, cs_rel\ nba_rel" unfolding complement_1_def complement_def proof parametricity fix A B :: "('a, 'b) nba" assume 1: "(A, B) \ \Id, Id\ nba_rel" have 2: "(const (Some (2 * card (nodes B), False)) |` initial B, const (Some (2 * card (nodes B))) |` initial B, {}) \ cs_rel" unfolding cs_lr_def cs_st_def in_br_conv by (force simp: restrict_map_def) show "(complement_succ_1 A, complement_succ B) \ Id \ cs_rel \ \cs_rel\ set_rel" using complement_succ_1_refine 1 by parametricity auto show "({const (Some (2 * card (nodes A), False)) |` initial A}, {const (Some (2 * card (nodes B))) |` initial B} \ {{}}) \ \cs_rel\ set_rel" using 1 2 by simp parametricity show "(\ f. cs_st f = {}, \ (f, P). P = {}) \ cs_rel \ bool_rel" by (auto simp: in_br_conv) qed subsection \Phase 2\ definition ranks_2 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "ranks_2 A a f \ {g. - dom g = \((succ A a) ` (dom f)) \ + dom g = \((transition A a) ` (dom f)) \ (\ q l d. g q = Some (l, d) \ l \ \ (fst ` Some -` f ` pred A a q) \ (d \ \ (snd ` Some -` f ` pred A a q) \ even l) \ (accepting A q \ even l))}" definition complement_succ_2 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_2 A a \ ranks_2 A a \ refresh_1" definition complement_2 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_2 A \ nba (alphabet A) ({const (Some (2 * card (nodes A), False)) |` initial A}) (complement_succ_2 A) (\ f. True \ snd ` ran f)" lemma ranks_2_refine: "ranks_2 = ranks_1" proof (intro ext) fix A :: "('a, 'b) nba" and a f show "ranks_2 A a f = ranks_1 A a f" proof safe fix g assume 1: "g \ ranks_2 A a f" - have 2: "dom g = \((succ A a) ` (dom f))" using 1 unfolding ranks_2_def by auto + have 2: "dom g = \((transition A a) ` (dom f))" using 1 unfolding ranks_2_def by auto have 3: "g q = Some (l, d) \ l \ \ (fst ` Some -` f ` pred A a q)" for q l d using 1 unfolding ranks_2_def by auto have 4: "g q = Some (l, d) \ d \ \ (snd ` Some -` f ` pred A a q) \ even l" for q l d using 1 unfolding ranks_2_def by auto have 5: "g q = Some (l, d) \ accepting A q \ even l" for q l d using 1 unfolding ranks_2_def by auto show "g \ ranks_1 A a f" unfolding ranks_1_def proof (intro CollectI conjI ballI impI) - show "dom g = \((succ A a) ` (dom f))" using 2 by this + show "dom g = \((transition A a) ` (dom f))" using 2 by this next fix p q - assume 10: "p \ dom f" "q \ succ A a p" + assume 10: "p \ dom f" "q \ transition A a p" obtain k c where 11: "f p = Some (k, c)" using 10(1) by auto have 12: "q \ dom g" using 10 2 by auto obtain l d where 13: "g q = Some (l, d)" using 12 by auto have "fst (the (g q)) = l" unfolding 13 by simp also have "\ \ \ (fst ` Some -` f ` pred A a q)" using 3 13 by this also have "\ \ k" proof (rule cInf_lower) show "k \ fst ` Some -` f ` pred A a q" using 11 10(2) by force show "bdd_below (fst ` Some -` f ` pred A a q)" by simp qed also have "\ = fst (the (f p))" unfolding 11 by simp finally show "fst (the (g q)) \ fst (the (f p))" by this next fix q assume 10: "q \ dom g" "accepting A q" show "even (fst (the (g q)))" using 10 5 by auto next - show "cs_st g = {q \ \((succ A a) ` (cs_st f)). even (fst (the (g q)))}" + show "cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" proof - show "cs_st g \ {q \ \((succ A a) ` (cs_st f)). even (fst (the (g q)))}" + show "cs_st g \ {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" using 4 unfolding cs_st_def image_def vimage_def by auto metis+ - show "{q \ \((succ A a) ` (cs_st f)). even (fst (the (g q)))} \ cs_st g" + show "{q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))} \ cs_st g" proof safe fix p q - assume 10: "even (fst (the (g q)))" "p \ cs_st f" "q \ succ A a p" + assume 10: "even (fst (the (g q)))" "p \ cs_st f" "q \ transition A a p" have 12: "q \ dom g" using 10 2 unfolding cs_st_def by auto show "q \ cs_st g" using 10 4 12 unfolding cs_st_def image_def by force qed qed qed next fix g assume 1: "g \ ranks_1 A a f" - have 2: "dom g = \((succ A a) ` (dom f))" using 1 unfolding ranks_1_def by auto - have 3: "\ p q. p \ dom f \ q \ succ A a p \ fst (the (g q)) \ fst (the (f p))" + have 2: "dom g = \((transition A a) ` (dom f))" using 1 unfolding ranks_1_def by auto + have 3: "\ p q. p \ dom f \ q \ transition A a p \ fst (the (g q)) \ fst (the (f p))" using 1 unfolding ranks_1_def by auto have 4: "\ q. q \ dom g \ accepting A q \ even (fst (the (g q)))" using 1 unfolding ranks_1_def by auto - have 5: "cs_st g = {q \ \((succ A a) ` (cs_st f)). even (fst (the (g q)))}" + have 5: "cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" using 1 unfolding ranks_1_def by auto show "g \ ranks_2 A a f" unfolding ranks_2_def proof (intro CollectI conjI allI impI) - show "dom g = \((succ A a) ` (dom f))" using 2 by this + show "dom g = \((transition A a) ` (dom f))" using 2 by this next fix q l d assume 10: "g q = Some (l, d)" have 11: "q \ dom g" using 10 by auto show "l \ \ (fst ` Some -` f ` pred A a q)" proof (rule cInf_greatest) show "fst ` Some -` f ` pred A a q \ {}" using 11 unfolding 2 image_def vimage_def by force show "\ x. x \ fst ` Some -` f ` pred A a q \ l \ x" using 3 10 by (auto) (metis domI fst_conv option.sel) qed have "d \ q \ cs_st g" unfolding cs_st_def by (force simp: 10) - also have "cs_st g = {q \ \((succ A a) ` (cs_st f)). even (fst (the (g q)))}" using 5 by this - also have "q \ \ \ (\ x \ cs_st f. q \ succ A a x) \ even l" + also have "cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" using 5 by this + also have "q \ \ \ (\ x \ cs_st f. q \ transition A a x) \ even l" unfolding mem_Collect_eq 10 by simp also have "\ \ \ (snd ` Some -` f ` pred A a q) \ even l" unfolding cs_st_def image_def vimage_def by auto metis+ finally show "d \ \ (snd ` Some -` f ` pred A a q) \ even l" by this show "accepting A q \ even l" using 4 10 11 by force qed qed qed lemma complement_2_refine: "(complement_2, complement_1) \ \Id, Id\ nba_rel \ \Id, Id\ nba_rel" unfolding complement_2_def complement_1_def complement_succ_2_def complement_succ_1_def unfolding ranks_2_refine cs_st_def image_def vimage_def ran_def by auto subsection \Phase 3\ definition bounds_3 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items" where "bounds_3 A a f \ \ q. let S = Some -` f ` pred A a q in if S = {} then None else Some (\(fst ` S), \(snd ` S))" definition items_3 :: "('label, 'state) nba \ 'state \ item \ item set" where "items_3 A p \ \ (k, c). {(l, c \ even l) |l. l \ k \ (accepting A p \ even l)}" definition get_3 :: "('label, 'state) nba \ 'state items \ ('state \ item set)" where "get_3 A f \ \ p. map_option (items_3 A p) (f p)" definition complement_succ_3 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_3 A a \ expand_map \ get_3 A \ bounds_3 A a \ refresh_1" definition complement_3 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_3 A \ nba (alphabet A) ({(Some \ (const (2 * card (nodes A), False))) |` initial A}) (complement_succ_3 A) (\ f. \ (p, k, c) \ map_to_set f. \ c)" - lemma bounds_3_dom[simp]: "dom (bounds_3 A a f) = \((succ A a) ` (dom f))" + lemma bounds_3_dom[simp]: "dom (bounds_3 A a f) = \((transition A a) ` (dom f))" unfolding bounds_3_def Let_def dom_def by (force split: if_splits) lemma items_3_nonempty[intro!, simp]: "items_3 A p s \ {}" unfolding items_3_def by auto lemma items_3_finite[intro!, simp]: "finite (items_3 A p s)" unfolding items_3_def by (auto split: prod.splits) lemma get_3_dom[simp]: "dom (get_3 A f) = dom f" unfolding get_3_def by (auto split: bind_splits) lemma get_3_finite[intro, simp]: "S \ ran (get_3 A f) \ finite S" unfolding get_3_def ran_def by auto lemma get_3_update[simp]: "get_3 A (f (p \ s)) = (get_3 A f) (p \ items_3 A p s)" unfolding get_3_def by auto lemma expand_map_get_bounds_3: "expand_map \ get_3 A \ bounds_3 A a = ranks_2 A a" proof (intro ext set_eqI, unfold comp_apply) fix f g have 1: "(\ x S y. get_3 A (bounds_3 A a f) x = Some S \ g x = Some y \ y \ S) \ (\ q S l d. get_3 A (bounds_3 A a f) q = Some S \ g q = Some (l, d) \ (l, d) \ S)" by auto have 2: "(\ S. get_3 A (bounds_3 A a f) q = Some S \ g q = Some (l, d) \ (l, d) \ S) \ (g q = Some (l, d) \ l \ \(fst ` (Some -` f ` pred A a q)) \ (d \ \(snd ` (Some -` f ` pred A a q)) \ even l) \ (accepting A q \ even l))" - if 3: "dom g = \((succ A a) ` (dom f))" for q l d + if 3: "dom g = \((transition A a) ` (dom f))" for q l d proof - have 4: "q \ dom g" if "Some -` f ` pred A a q = {}" unfolding 3 using that by force show ?thesis unfolding get_3_def items_3_def bounds_3_def Let_def using 4 by auto qed show "g \ expand_map (get_3 A (bounds_3 A a f)) \ g \ ranks_2 A a f" unfolding expand_map_alt_def ranks_2_def mem_Collect_eq unfolding get_3_dom bounds_3_dom 1 using 2 by blast qed lemma complement_succ_3_refine: "complement_succ_3 = complement_succ_2" unfolding complement_succ_3_def complement_succ_2_def expand_map_get_bounds_3 by rule lemma complement_initial_3_refine: "{const (Some (2 * card (nodes A), False)) |` initial A} = {(Some \ (const (2 * card (nodes A), False))) |` initial A}" unfolding comp_apply by rule lemma complement_accepting_3_refine: "True \ snd ` ran f \ (\ (p, k, c) \ map_to_set f. \ c)" unfolding map_to_set_def ran_def by auto lemma complement_3_refine: "(complement_3, complement_2) \ \Id, Id\ nba_rel \ \Id, Id\ nba_rel" unfolding complement_3_def complement_2_def unfolding complement_succ_3_refine complement_initial_3_refine complement_accepting_3_refine by auto subsection \Phase 4\ definition items_4 :: "('label, 'state) nba \ 'state \ item \ item set" where "items_4 A p \ \ (k, c). {(l, c \ even l) |l. k \ Suc l \ l \ k \ (accepting A p \ even l)}" definition get_4 :: "('label, 'state) nba \ 'state items \ ('state \ item set)" where "get_4 A f \ \ p. map_option (items_4 A p) (f p)" definition complement_succ_4 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_4 A a \ expand_map \ get_4 A \ bounds_3 A a \ refresh_1" definition complement_4 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_4 A \ nba (alphabet A) ({(Some \ (const (2 * card (nodes A), False))) |` initial A}) (complement_succ_4 A) (\ f. \ (p, k, c) \ map_to_set f. \ c)" lemma get_4_dom[simp]: "dom (get_4 A f) = dom f" unfolding get_4_def by (auto split: bind_splits) definition R :: "'state items rel" where "R \ {(f, g). dom f = dom g \ (\ p \ dom f. fst (the (f p)) \ fst (the (g p))) \ (\ p \ dom f. snd (the (f p)) \ snd (the (g p)))}" lemma bounds_R: assumes "(f, g) \ R" assumes "bounds_3 A a (refresh_1 f) p = Some (n, e)" assumes "bounds_3 A a (refresh_1 g) p = Some (k, c)" shows "n \ k" "e \ c" proof - have 1: "dom f = dom g" "\ p \ dom f. fst (the (f p)) \ fst (the (g p))" "\ p \ dom f. snd (the (f p)) \ snd (the (g p))" using assms(1) unfolding R_def by auto have "n = \(fst ` (Some -` refresh_1 f ` pred A a p))" using assms(2) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) also have "fst ` Some -` refresh_1 f ` pred A a p = fst ` Some -` f ` pred A a p" proof show " fst ` Some -` refresh_1 f ` pred A a p \ fst ` Some -` f ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (force) show "fst ` Some -` f ` pred A a p \ fst ` Some -` refresh_1 f ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (metis fst_conv option.sel) qed also have "\ = fst ` Some -` f ` (pred A a p \ dom f)" unfolding dom_def image_def Int_def by auto metis also have "\ = fst ` the ` f ` (pred A a p \ dom f)" unfolding dom_def by force also have "\ = (fst \ the \ f) ` (pred A a p \ dom f)" by force also have "\((fst \ the \ f) ` (pred A a p \ dom f)) \ \((fst \ the \ g) ` (pred A a p \ dom g))" proof (rule cINF_mono) show "pred A a p \ dom g \ {}" using assms(2) 1(1) unfolding bounds_3_def refresh_1_def by (auto simp: Let_def split: if_splits) (force+) show "bdd_below ((fst \ the \ f) ` (pred A a p \ dom f))" by rule show "\ n \ pred A a p \ dom f. (fst \ the \ f) n \ (fst \ the \ g) m" if "m \ pred A a p \ dom g" for m using 1 that by auto qed also have "(fst \ the \ g) ` (pred A a p \ dom g) = fst ` the ` g ` (pred A a p \ dom g)" by force also have "\ = fst ` Some -` g ` (pred A a p \ dom g)" unfolding dom_def by force also have "\ = fst ` Some -` g ` pred A a p" unfolding dom_def image_def Int_def by auto metis also have "\ = fst ` Some -` refresh_1 g ` pred A a p" proof show "fst ` Some -` g ` pred A a p \ fst ` Some -` refresh_1 g ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (metis fst_conv option.sel) show "fst ` Some -` refresh_1 g ` pred A a p \ fst ` Some -` g ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (force) qed also have "\(fst ` (Some -` refresh_1 g ` pred A a p)) = k" using assms(3) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) finally show "n \ k" by this have "e \ \(snd ` (Some -` refresh_1 f ` pred A a p))" using assms(2) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) also have "snd ` Some -` refresh_1 f ` pred A a p = snd ` Some -` refresh_1 f ` (pred A a p \ dom (refresh_1 f))" unfolding dom_def image_def Int_def by auto metis also have "\ = snd ` the ` refresh_1 f ` (pred A a p \ dom (refresh_1 f))" unfolding dom_def by force also have "\ = (snd \ the \ refresh_1 f) ` (pred A a p \ dom (refresh_1 f))" by force also have "\ = (snd \ the \ refresh_1 g) ` (pred A a p \ dom (refresh_1 g))" proof (rule image_cong) show "pred A a p \ dom (refresh_1 f) = pred A a p \ dom (refresh_1 g)" unfolding refresh_1_dom 1(1) by rule show "(snd \ the \ refresh_1 f) q \ (snd \ the \ refresh_1 g) q" if 2: "q \ pred A a p \ dom (refresh_1 g)" for q proof have 3: "\ x \ ran f. \ snd x \ (n, True) \ ran g \ g q = Some (k, c) \ c" for n k c using 1(1, 3) unfolding dom_def ran_def by (auto dest!: Collect_inj) (metis option.sel snd_conv) have 4: "g q = Some (n, True) \ f q = Some (k, c) \ c" for n k c using 1(3) unfolding dom_def by force have 5: "\ x \ ran g. \ snd x \ (k, True) \ ran f \ False" for k using 1(1, 3) unfolding dom_def ran_def by (auto dest!: Collect_inj) (metis option.sel snd_conv) show "(snd \ the \ refresh_1 f) q \ (snd \ the \ refresh_1 g) q" using 1(1, 3) 2 3 unfolding refresh_1_def by (force split: if_splits) show "(snd \ the \ refresh_1 g) q \ (snd \ the \ refresh_1 f) q" using 1(1, 3) 2 4 5 unfolding refresh_1_def by (auto simp: map_option_case split: option.splits if_splits) (force+) qed qed also have "\ = snd ` the ` refresh_1 g ` (pred A a p \ dom (refresh_1 g))" by force also have "\ = snd ` Some -` refresh_1 g ` (pred A a p \ dom (refresh_1 g))" unfolding dom_def by force also have "\ = snd ` Some -` refresh_1 g ` pred A a p" unfolding dom_def image_def Int_def by auto metis also have "\(snd ` (Some -` refresh_1 g ` pred A a p)) \ c" using assms(3) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) finally show "e \ c" by this qed lemma complement_4_language_1: "language (complement_3 A) \ language (complement_4 A)" proof (rule simulation_language) show "alphabet (complement_3 A) \ alphabet (complement_4 A)" unfolding complement_3_def complement_4_def by simp show "\ q \ initial (complement_4 A). (p, q) \ R" if "p \ initial (complement_3 A)" for p using that unfolding complement_3_def complement_4_def R_def by simp - show "\ g' \ succ (complement_4 A) a g. (f', g') \ R" - if "f' \ succ (complement_3 A) a f" "(f, g) \ R" + show "\ g' \ transition (complement_4 A) a g. (f', g') \ R" + if "f' \ transition (complement_3 A) a f" "(f, g) \ R" for a f f' g proof - have 1: "f' \ expand_map (get_3 A (bounds_3 A a (refresh_1 f)))" using that(1) unfolding complement_3_def complement_succ_3_def by auto have 2: "dom f = dom g" "\ p \ dom f. fst (the (f p)) \ fst (the (g p))" "\ p \ dom f. snd (the (f p)) \ snd (the (g p))" using that(2) unfolding R_def by auto have "dom f' = dom (get_3 A (bounds_3 A a (refresh_1 f)))" using expand_map_dom 1 by this also have "\ = dom (bounds_3 A a (refresh_1 f))" by simp finally have 3: "dom f' = dom (bounds_3 A a (refresh_1 f))" by this define g' where "g' p \ do { (k, c) \ bounds_3 A a (refresh_1 g) p; (l, d) \ f' p; Some (if even k = even l then k else k - 1, d) }" for p have 4: "g' p = do { kc \ bounds_3 A a (refresh_1 g) p; ld \ f' p; Some (if even (fst kc) = even (fst ld) then fst kc else fst kc - 1, snd ld) }" for p unfolding g'_def case_prod_beta by rule have "dom g' = dom (bounds_3 A a (refresh_1 g)) \ dom f'" using 4 bind_eq_Some_conv by fastforce also have "\ = dom f'" using 2 3 by simp finally have 5: "dom g' = dom f'" by this have 6: "(l, d) \ items_3 A p (k, c)" if "bounds_3 A a (refresh_1 f) p = Some (k, c)" "f' p = Some (l, d)" for p k c l d using 1 that unfolding expand_map_alt_def get_3_def by blast show ?thesis unfolding complement_4_def nba.sel complement_succ_4_def comp_apply proof show "(f', g') \ R" unfolding R_def mem_Collect_eq prod.case proof (intro conjI ballI) show "dom f' = dom g'" using 5 by rule next fix p assume 10: "p \ dom f'" have 11: "p \ dom (bounds_3 A a (refresh_1 g))" using 2(1) 3 10 by simp obtain k c where 12: "bounds_3 A a (refresh_1 g) p = Some (k, c)" using 11 by fast obtain l d where 13: "f' p = Some (l, d)" using 10 by auto obtain n e where 14: "bounds_3 A a (refresh_1 f) p = Some (n, e)" using 10 3 by fast have 15: "(l, d) \ items_3 A p (n, e)" using 6 14 13 by this have 16: "n \ k" using bounds_R(1) that(2) 14 12 by this have 17: "l \ k" using 15 16 unfolding items_3_def by simp have 18: "even k \ odd l \ l \ k \ l \ k - 1" by presburger have 19: "e \ c" using bounds_R(2) that(2) 14 12 by this show "fst (the (f' p)) \ fst (the (g' p))" using 17 18 unfolding 4 12 13 by simp show "snd (the (f' p)) \ snd (the (g' p))" using 19 unfolding 4 12 13 by simp qed show "g' \ expand_map (get_4 A (bounds_3 A a (refresh_1 g)))" unfolding expand_map_alt_def mem_Collect_eq proof (intro conjI allI impI) show "dom g' = dom (get_4 A (bounds_3 A a (refresh_1 g)))" using 2(1) 3 5 by simp fix p S xy assume 10: "get_4 A (bounds_3 A a (refresh_1 g)) p = Some S" assume 11: "g' p = Some xy" obtain k c where 12: "bounds_3 A a (refresh_1 g) p = Some (k, c)" "S = items_4 A p (k, c)" using 10 unfolding get_4_def by auto obtain l d where 13: "f' p = Some (l, d)" "xy = (if even k \ even l then k else k - 1, d)" using 11 12 unfolding g'_def by (auto split: bind_splits) obtain n e where 14: "bounds_3 A a (refresh_1 f) p = Some (n, e)" using 13(1) 3 by fast have 15: "(l, d) \ items_3 A p (n, e)" using 6 14 13(1) by this have 16: "n \ k" using bounds_R(1) that(2) 14 12(1) by this have 17: "e \ c" using bounds_R(2) that(2) 14 12(1) by this show "xy \ S" using 15 16 17 unfolding 12(2) 13(2) items_3_def items_4_def by auto qed qed qed show "\ p q. (p, q) \ R \ accepting (complement_3 A) p \ accepting (complement_4 A) q" unfolding complement_3_def complement_4_def R_def map_to_set_def by (auto) (metis domIff eq_snd_iff option.exhaust_sel option.sel) qed lemma complement_4_less: "complement_4 A \ complement_3 A" unfolding less_eq_nba_def unfolding complement_4_def complement_3_def nba.sel unfolding complement_succ_4_def complement_succ_3_def proof (safe intro!: le_funI, unfold comp_apply) fix a f g assume "g \ expand_map (get_4 A (bounds_3 A a (refresh_1 f)))" then show "g \ expand_map (get_3 A (bounds_3 A a (refresh_1 f)))" unfolding get_4_def get_3_def items_4_def items_3_def expand_map_alt_def by blast qed lemma complement_4_language_2: "language (complement_4 A) \ language (complement_3 A)" using language_mono complement_4_less by (auto dest: monoD) lemma complement_4_language: "language (complement_3 A) = language (complement_4 A)" using complement_4_language_1 complement_4_language_2 by blast lemma complement_4_finite[simp]: assumes "finite (nodes A)" shows "finite (nodes (complement_4 A))" proof - have "(nodes (complement_3 A), nodes (complement_2 A)) \ \Id\ set_rel" using complement_3_refine by parametricity auto also have "(nodes (complement_2 A), nodes (complement_1 A)) \ \Id\ set_rel" using complement_2_refine by parametricity auto also have "(nodes (complement_1 A), nodes (complement A)) \ \cs_rel\ set_rel" using complement_1_refine by parametricity auto finally have 1: "(nodes (complement_3 A), nodes (complement A)) \ \cs_rel\ set_rel" by simp have 2: "finite (nodes (complement A))" using complement_finite assms(1) by this have 3: "finite (nodes (complement_3 A))" using finite_set_rel_transfer_back 1 cs_rel_inv_single_valued 2 by this have 4: "nodes (complement_4 A) \ nodes (complement_3 A)" using nodes_mono complement_4_less by (auto dest: monoD) show "finite (nodes (complement_4 A))" using finite_subset 4 3 by this qed lemma complement_4_correct: assumes "finite (nodes A)" shows "language (complement_4 A) = streams (alphabet A) - language A" proof - have "language (complement_4 A) = language (complement_3 A)" using complement_4_language by rule also have "(language (complement_3 A), language (complement_2 A)) \ \\Id\ stream_rel\ set_rel" using complement_3_refine by parametricity auto also have "(language (complement_2 A), language (complement_1 A)) \ \\Id\ stream_rel\ set_rel" using complement_2_refine by parametricity auto also have "(language (complement_1 A), language (complement A)) \ \\Id\ stream_rel\ set_rel" using complement_1_refine by parametricity auto also have "language (complement A) = streams (alphabet A) - language A" using complement_language assms(1) by this finally show "language (complement_4 A) = streams (alphabet A) - language A" by simp qed subsection \Phase 5\ definition refresh_5 :: "'state items \ 'state items nres" where "refresh_5 f \ if \ (p, k, c) \ map_to_set f. c then RETURN f else do { ASSUME (finite (dom f)); FOREACH (map_to_set f) (\ (p, k, c) m. do { ASSERT (p \ dom m); RETURN (m (p \ (k, True))) } ) Map.empty }" definition merge_5 :: "item \ item option \ item" where "merge_5 \ \ (k, c). \ None \ (k, c) | Some (l, d) \ (k \ l, c \ d)" definition bounds_5 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items nres" where "bounds_5 A a f \ do { ASSUME (finite (dom f)); - ASSUME (\ p. finite (succ A a p)); + ASSUME (\ p. finite (transition A a p)); FOREACH (map_to_set f) (\ (p, s) m. - FOREACH (succ A a p) (\ q f. + FOREACH (transition A a p) (\ q f. RETURN (f (q \ merge_5 s (f q)))) m) Map.empty }" definition items_5 :: "('label, 'state) nba \ 'state \ item \ item set" where "items_5 A p \ \ (k, c). do { let values = if accepting A p then Set.filter even {k - 1 .. k} else {k - 1 .. k}; let item = \ l. (l, c \ even l); item ` values }" definition get_5 :: "('label, 'state) nba \ 'state items \ ('state \ item set)" where "get_5 A f \ \ p. map_option (items_5 A p) (f p)" definition expand_5 :: "('a \ 'b set) \ ('a \ 'b) set nres" where "expand_5 f \ FOREACH (map_to_set f) (\ (x, S) X. do { ASSERT (\ g \ X. x \ dom g); ASSERT (\ a \ S. \ b \ S. a \ b \ (\ y. (\ g. g (x \ y)) ` X) a \ (\ y. (\ g. g (x \ y)) ` X) b = {}); RETURN (\ y \ S. (\ g. g (x \ y)) ` X) }) {Map.empty}" definition complement_succ_5 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set nres" where "complement_succ_5 A a f \ do { f \ refresh_5 f; f \ bounds_5 A a f; ASSUME (finite (dom f)); expand_5 (get_5 A f) }" lemma bounds_3_empty: "bounds_3 A a Map.empty = Map.empty" unfolding bounds_3_def Let_def by auto lemma bounds_3_update: "bounds_3 A a (f (p \ s)) = - override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) (succ A a p)" + override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) (transition A a p)" proof note fun_upd_image[simp] fix q show "bounds_3 A a (f (p \ s)) q = - override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) (succ A a p) q" - proof (cases "q \ succ A a p") + override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) (transition A a p) q" + proof (cases "q \ transition A a p") case True define S where "S \ Some -` f ` (pred A a q - {p})" have 1: "Some -` f (p := Some s) ` pred A a q = insert s S" using True unfolding S_def by auto have 2: "Some -` f (p := None) ` pred A a q = S" unfolding S_def by auto have "bounds_3 A a (f (p \ s)) q = Some (\(fst ` (insert s S)), \(snd ` (insert s S)))" unfolding bounds_3_def 1 by simp also have "\ = Some (merge_5 s (bounds_3 A a (f (p := None)) q))" unfolding 2 bounds_3_def merge_5_def by (cases s) (simp_all add: cInf_insert) also have "\ = override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) - (succ A a p) q" using True by simp + (transition A a p) q" using True by simp finally show ?thesis by this next case False then have "pred A a q \ {x. x \ p} = pred A a q" by auto with False show ?thesis by (simp add: bounds_3_def) qed qed lemma refresh_5_refine: "(refresh_5, \ f. RETURN (refresh_1 f)) \ Id \ \Id\ nres_rel" proof safe fix f :: "'a \ item option" have 1: "(\ (p, k, c) \ map_to_set f. c) \ True \ snd ` ran f" unfolding image_def map_to_set_def ran_def by force show "(refresh_5 f, RETURN (refresh_1 f)) \ \Id\ nres_rel" unfolding refresh_5_def refresh_1_def 1 by (refine_vcg FOREACH_rule_map_eq[where X = "\ m. map_option (apsnd \) \ m"]) (auto) qed lemma bounds_5_refine: "(bounds_5 A a, \ f. RETURN (bounds_3 A a f)) \ Id \ \Id\ nres_rel" unfolding bounds_5_def by (refine_vcg FOREACH_rule_map_eq[where X = "bounds_3 A a"] FOREACH_rule_insert_eq) (auto simp: override_on_insert bounds_3_empty bounds_3_update) lemma items_5_refine: "items_5 = items_4" unfolding items_5_def items_4_def by (intro ext) (auto split: if_splits) lemma get_5_refine: "get_5 = get_4" unfolding get_5_def get_4_def items_5_refine by rule lemma expand_5_refine: "(expand_5 f, ASSERT (finite (dom f)) \ RETURN (expand_map f)) \ \Id\ nres_rel" unfolding expand_5_def by (refine_vcg FOREACH_rule_map_eq[where X = expand_map]) (auto dest!: expand_map_dom map_upd_eqD1) lemma complement_succ_5_refine: "(complement_succ_5, RETURN \\\ complement_succ_4) \ Id \ Id \ Id \ \Id\ nres_rel" unfolding complement_succ_5_def complement_succ_4_def get_5_refine comp_apply by (refine_vcg vcg1[OF refresh_5_refine] vcg1[OF bounds_5_refine] vcg0[OF expand_5_refine]) (auto) subsection \Phase 6\ definition expand_map_get_6 :: "('label, 'state) nba \ 'state items \ 'state items set nres" where "expand_map_get_6 A f \ FOREACH (map_to_set f) (\ (k, v) X. do { ASSERT (\ g \ X. k \ dom g); ASSERT (\ a \ (items_5 A k v). \ b \ (items_5 A k v). a \ b \ (\ y. (\ g. g (k \ y)) ` X) a \ (\ y. (\ g. g (k \ y)) ` X) b = {}); RETURN (\ y \ items_5 A k v. (\ g. g (k \ y)) ` X) }) {Map.empty}" lemma expand_map_get_6_refine: "(expand_map_get_6, expand_5 \\ get_5) \ Id \ Id \ \Id\ nres_rel" unfolding expand_map_get_6_def expand_5_def get_5_def by (auto intro: FOREACH_rule_map_map[param_fo]) definition complement_succ_6 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set nres" where "complement_succ_6 A a f \ do { f \ refresh_5 f; f \ bounds_5 A a f; ASSUME (finite (dom f)); expand_map_get_6 A f }" lemma complement_succ_6_refine: "(complement_succ_6, complement_succ_5) \ Id \ Id \ Id \ \Id\ nres_rel" unfolding complement_succ_6_def complement_succ_5_def by (refine_vcg vcg2[OF expand_map_get_6_refine]) (auto intro: refine_IdI) subsection \Phase 7\ interpretation autoref_syn by this context fixes fi f assumes fi[autoref_rules]: "(fi, f) \ state_rel" begin private lemma [simp]: "finite (dom f)" using list_map_rel_finite fi unfolding finite_map_rel_def by force schematic_goal refresh_7: "(?f :: ?'a, refresh_5 f) \ ?R" unfolding refresh_5_def by (autoref_monadic (plain)) end concrete_definition refresh_7 uses refresh_7 lemma refresh_7_refine: "(\ f. RETURN (refresh_7 f), refresh_5) \ state_rel \ \state_rel\ nres_rel" using refresh_7.refine by fast context fixes A :: "('label, nat) nba" fixes succi a fi f - assumes succi[autoref_rules]: "(succi, succ A a) \ nat_rel \ \nat_rel\ list_set_rel" + assumes succi[autoref_rules]: "(succi, transition A a) \ nat_rel \ \nat_rel\ list_set_rel" assumes fi[autoref_rules]: "(fi, f) \ state_rel" begin - private lemma [simp]: "finite (succ A a p)" + private lemma [simp]: "finite (transition A a p)" using list_set_rel_finite succi[param_fo] unfolding finite_set_rel_def by blast private lemma [simp]: "finite (dom f)" using fi by force - private lemma [autoref_op_pat]: "succ A a \ OP (succ A a)" by simp + private lemma [autoref_op_pat]: "transition A a \ OP (transition A a)" by simp private lemma [autoref_rules]: "(min, min) \ nat_rel \ nat_rel \ nat_rel" by simp schematic_goal bounds_7: notes ty_REL[where R = "\nat_rel, item_rel\ dflt_ahm_rel", autoref_tyrel] shows "(?f :: ?'a, bounds_5 A a f) \ ?R" unfolding bounds_5_def merge_5_def sup_bool_def inf_nat_def by (autoref_monadic (plain)) end concrete_definition bounds_7 uses bounds_7 - lemma bounds_7_refine: "(si, succ A a) \ nat_rel \ \nat_rel\ list_set_rel \ + lemma bounds_7_refine: "(si, transition A a) \ nat_rel \ \nat_rel\ list_set_rel \ (\ p. RETURN (bounds_7 si p), bounds_5 A a) \ state_rel \ \\nat_rel, item_rel\ dflt_ahm_rel\ nres_rel" using bounds_7.refine by auto context fixes A :: "('label, nat) nba" fixes acci assumes [autoref_rules]: "(acci, accepting A) \ nat_rel \ bool_rel" begin private lemma [autoref_op_pat]: "accepting A \ OP (accepting A)" by simp private lemma [autoref_rules]: "((dvd), (dvd)) \ nat_rel \ nat_rel \ bool_rel" by simp private lemma [autoref_rules]: "(\ k l. upt k (Suc l), atLeastAtMost) \ nat_rel \ nat_rel \ \nat_rel\ list_set_rel" by (auto simp: list_set_rel_def in_br_conv) schematic_goal items_7: "(?f :: ?'a, items_5 A) \ ?R" unfolding items_5_def Let_def Set.filter_def by autoref end concrete_definition items_7 uses items_7 (* TODO: use generic expand_map implementation *) context fixes A :: "('label, nat) nba" fixes ai fixes fi f assumes ai: "(ai, accepting A) \ nat_rel \ bool_rel" assumes fi[autoref_rules]: "(fi, f) \ \nat_rel, item_rel\ dflt_ahm_rel" begin private lemma [simp]: "finite (dom f)" using dflt_ahm_rel_finite_nat fi unfolding finite_map_rel_def by force private lemma [simp]: assumes "\ m. m \ S \ x \ dom m" shows "inj_on (\ m. m (x \ y)) S" using assms unfolding dom_def inj_on_def by (auto) (metis fun_upd_triv fun_upd_upd) private lemmas [simp] = op_map_update_def[abs_def] private lemma [autoref_op_pat]: "items_5 A \ OP (items_5 A)" by simp private lemmas [autoref_rules] = items_7.refine[OF ai] schematic_goal expand_map_get_7: "(?f, expand_map_get_6 A f) \ \\state_rel\ list_set_rel\ nres_rel" unfolding expand_map_get_6_def by (autoref_monadic (plain)) end concrete_definition expand_map_get_7 uses expand_map_get_7 lemma expand_map_get_7_refine: assumes "(ai, accepting A) \ nat_rel \ bool_rel" shows "(\ fi. RETURN (expand_map_get_7 ai fi), \ f. ASSUME (finite (dom f)) \ expand_map_get_6 A f) \ \nat_rel, item_rel\ dflt_ahm_rel \ \\state_rel\ list_set_rel\ nres_rel" using expand_map_get_7.refine[OF assms] by auto context fixes A :: "('label, nat) nba" fixes a :: "'label" fixes p :: "nat items" fixes Ai fixes ai fixes pi assumes Ai: "(Ai, A) \ \Id, Id\ nbai_nba_rel" assumes ai: "(ai, a) \ Id" assumes pi[autoref_rules]: "(pi, p) \ state_rel" begin private lemmas succi = nbai_nba_param(4)[THEN fun_relD, OF Ai, THEN fun_relD, OF ai] private lemmas acceptingi = nbai_nba_param(5)[THEN fun_relD, OF Ai] private lemma [autoref_op_pat]: "(\ g. ASSUME (finite (dom g)) \ expand_map_get_6 A g) \ OP (\ g. ASSUME (finite (dom g)) \ expand_map_get_6 A g)" by simp private lemma [autoref_op_pat]: "bounds_5 A a \ OP (bounds_5 A a)" by simp private lemmas [autoref_rules] = refresh_7_refine bounds_7_refine[OF succi] expand_map_get_7_refine[OF acceptingi] schematic_goal complement_succ_7: "(?f :: ?'a, complement_succ_6 A a p) \ ?R" unfolding complement_succ_6_def by (autoref_monadic (plain)) end concrete_definition complement_succ_7 uses complement_succ_7 lemma complement_succ_7_refine: "(RETURN \\\ complement_succ_7, complement_succ_6) \ \Id, Id\ nbai_nba_rel \ Id \ state_rel \ \\state_rel\ list_set_rel\ nres_rel" using complement_succ_7.refine unfolding comp_apply by parametricity context fixes A :: "('label, nat) nba" fixes Ai fixes n ni :: nat assumes Ai: "(Ai, A) \ \Id, Id\ nbai_nba_rel" assumes ni[autoref_rules]: "(ni, n) \ Id" begin private lemma [autoref_op_pat]: "initial A \ OP (initial A)" by simp private lemmas [autoref_rules] = nbai_nba_param(3)[THEN fun_relD, OF Ai] schematic_goal complement_initial_7: "(?f, {(Some \ (const (2 * n, False))) |` initial A}) \ \state_rel\ list_set_rel" by autoref end concrete_definition complement_initial_7 uses complement_initial_7 schematic_goal complement_accepting_7: "(?f, \ f. \ (p, k, c) \ map_to_set f. \ c) \ state_rel \ bool_rel" by autoref concrete_definition complement_accepting_7 uses complement_accepting_7 definition complement_7 :: "('label, nat) nbai \ nat \ ('label, state) nbai" where "complement_7 Ai ni \ nbai (alphabeti Ai) (complement_initial_7 Ai ni) (complement_succ_7 Ai) (complement_accepting_7)" lemma complement_7_refine[autoref_rules]: assumes "(Ai, A) \ \Id, Id\ nbai_nba_rel" assumes "(ni, (OP card ::: \Id\ ahs_rel bhc \ nat_rel) $ ((OP nodes ::: \Id, Id\ nbai_nba_rel \ \Id\ ahs_rel bhc) $ A)) \ nat_rel" shows "(complement_7 Ai ni, (OP complement_4 ::: \Id, Id\ nbai_nba_rel \ \Id, state_rel\ nbai_nba_rel) $ A) \ \Id, state_rel\ nbai_nba_rel" proof - note complement_succ_7_refine also note complement_succ_6_refine also note complement_succ_5_refine finally have 1: "(complement_succ_7, complement_succ_4) \ \Id, Id\ nbai_nba_rel \ Id \ state_rel \ \state_rel\ list_set_rel" unfolding nres_rel_comp unfolding nres_rel_def unfolding fun_rel_def by auto show ?thesis unfolding complement_7_def complement_4_def using 1 complement_initial_7.refine complement_accepting_7.refine assms unfolding autoref_tag_defs by parametricity qed end \ No newline at end of file diff --git a/thys/Buchi_Complementation/Graph.thy b/thys/Buchi_Complementation/Graph.thy --- a/thys/Buchi_Complementation/Graph.thy +++ b/thys/Buchi_Complementation/Graph.thy @@ -1,205 +1,205 @@ section \Run Graphs\ theory Graph imports "Transition_Systems_and_Automata.NBA" begin type_synonym 'state node = "nat \ 'state" abbreviation "ginitial A \ {0} \ initial A" abbreviation "gaccepting A \ accepting A \ snd" global_interpretation graph: transition_system_initial "const" - "\ u (k, p). w !! k \ alphabet A \ u \ {Suc k} \ succ A (w !! k) p \ V" + "\ u (k, p). w !! k \ alphabet A \ u \ {Suc k} \ transition A (w !! k) p \ V" "\ v. v \ ginitial A \ V" for A w V defines gpath = graph.path and grun = graph.run and greachable = graph.reachable and gnodes = graph.nodes by this text \We disable rules that are degenerate due to @{term "execute = const"}.\ declare graph.reachable.execute[rule del] declare graph.nodes.execute[rule del] abbreviation "gtarget \ graph.target" abbreviation "gstates \ graph.states" abbreviation "gtrace \ graph.trace" abbreviation gsuccessors :: "('label, 'state) nba \ 'label stream \ 'state node set \ 'state node \ 'state node set" where "gsuccessors A w V \ graph.successors TYPE('label) w A V" abbreviation "gusuccessors A w \ gsuccessors A w UNIV" abbreviation "gupath A w \ gpath A w UNIV" abbreviation "gurun A w \ grun A w UNIV" abbreviation "gureachable A w \ greachable A w UNIV" abbreviation "gunodes A w \ gnodes A w UNIV" lemma gtarget_alt_def: "gtarget r v = last (v # r)" using fold_const by this lemma gstates_alt_def: "gstates r v = r" by simp lemma gtrace_alt_def: "gtrace r v = r" by simp lemma gpath_elim[elim?]: assumes "gpath A w V s v" obtains r k p where "s = [Suc k ..< Suc k + length r] || r" "v = (k, p)" proof - obtain t r where 1: "s = t || r" "length t = length r" using zip_map_fst_snd[of s] by (metis length_map) obtain k p where 2: "v = (k, p)" by force have 3: "t = [Suc k ..< Suc k + length r]" using assms 1 2 proof (induct arbitrary: t r k p) case (nil v) then show ?case by (metis add_0_right le_add1 length_0_conv length_zip min.idem upt_conv_Nil) next case (cons u v s) have 1: "t || r = (hd t, hd r) # (tl t || tl r)" by (metis cons.prems(1) hd_Cons_tl neq_Nil_conv zip.simps(1) zip_Cons_Cons zip_Nil) have 2: "s = tl t || tl r" using cons 1 by simp have "t = hd t # tl t" using cons(4) by (metis hd_Cons_tl list.simps(3) zip_Nil) also have "hd t = Suc k" using "1" cons.hyps(1) cons.prems(1) cons.prems(3) by auto also have "tl t = [Suc (Suc k) ..< Suc (Suc k) + length (tl r)]" using cons(3)[OF 2] using "1" \hd t = Suc k\ cons.prems(1) cons.prems(2) by auto finally show ?case using cons.prems(2) upt_rec by auto qed show ?thesis using that 1 2 3 by simp qed lemma gpath_path[symmetric]: "path A (stake (length r) (sdrop k w) || r) p \ gpath A w UNIV ([Suc k ..< Suc k + length r] || r) (k, p)" proof (induct r arbitrary: k p) case (Nil) show ?case by auto next case (Cons q r) have 1: "path A (stake (length r) (sdrop (Suc k) w) || r) q \ gpath A w UNIV ([Suc (Suc k) ..< Suc k + length (q # r)] || r) (Suc k, q)" using Cons[of "Suc k" "q"] by simp have "stake (length (q # r)) (sdrop k w) || q # r = (w !! k, q) # (stake (length r) (sdrop (Suc k) w) || r)" by simp also have "path A \ p \ gpath A w UNIV ((Suc k, q) # ([Suc (Suc k) ..< Suc k + length (q # r)] || r)) (k, p)" using 1 by auto also have "(Suc k, q) # ([Suc (Suc k) ..< Suc k + length (q # r)] || r) = Suc k # [Suc (Suc k) ..< Suc k + length (q # r)] || q # r" unfolding zip_Cons_Cons by rule also have "Suc k # [Suc (Suc k) ..< Suc k + length (q # r)] = [Suc k ..< Suc k + length (q # r)]" by (simp add: upt_rec) finally show ?case by this qed lemma grun_elim[elim?]: assumes "grun A w V s v" obtains r k p where "s = fromN (Suc k) ||| r" "v = (k, p)" proof - obtain t r where 1: "s = t ||| r" using szip_smap by metis obtain k p where 2: "v = (k, p)" by force have 3: "t = fromN (Suc k)" using assms unfolding 1 2 by (coinduction arbitrary: t r k p) (force iff: eq_scons elim: graph.run.cases) show ?thesis using that 1 2 3 by simp qed lemma run_grun: assumes "run A (sdrop k w ||| r) p" shows "gurun A w (fromN (Suc k) ||| r) (k, p)" using assms by (coinduction arbitrary: k p r) (auto elim: nba.run.cases) lemma grun_run: assumes "grun A w V (fromN (Suc k) ||| r) (k, p)" shows "run A (sdrop k w ||| r) p" proof - have 2: "\ ka wa. sdrop k (stl w :: 'a stream) = sdrop ka wa \ P ka wa" if "P (Suc k) w" for P k w using that by (metis sdrop.simps(2)) show ?thesis using assms by (coinduction arbitrary: k p w r) (auto intro!: 2 elim: graph.run.cases) qed lemma greachable_reachable: fixes l q k p defines "u \ (l, q)" defines "v \ (k, p)" assumes "u \ greachable A w V v" shows "q \ reachable A p" using assms(3, 1, 2) proof (induct arbitrary: l q k p) case reflexive then show ?case by auto next case (execute u) have 1: "q \ successors A (snd u)" using execute by auto have "snd u \ reachable A p" using execute by auto also have "q \ reachable A (snd u)" using 1 by blast finally show ?case by this qed lemma gnodes_nodes: "gnodes A w V \ UNIV \ nodes A" proof fix v assume "v \ gnodes A w V" then show "v \ UNIV \ nodes A" by induct auto qed lemma gpath_subset: assumes "gpath A w V r v" assumes "set (gstates r v) \ U" shows "gpath A w U r v" using assms by induct auto lemma grun_subset: assumes "grun A w V r v" assumes "sset (gtrace r v) \ U" shows "grun A w U r v" using assms proof (coinduction arbitrary: r v) case (run a s r v) have 1: "grun A w V s a" using run(1, 2) by fastforce have 2: "a \ gusuccessors A w v" using run(1, 2) by fastforce show ?case using 1 2 run(1, 3) by force qed lemma greachable_subset: "greachable A w V v \ insert v V" proof fix u assume "u \ greachable A w V v" then show "u \ insert v V" by induct auto qed lemma gtrace_infinite: assumes "grun A w V r v" shows "infinite (sset (gtrace r v))" using assms by (metis grun_elim gtrace_alt_def infinite_Ici sset_fromN sset_szip_finite) lemma infinite_greachable_gtrace: assumes "grun A w V r v" assumes "u \ sset (gtrace r v)" shows "infinite (greachable A w V u)" proof - obtain i where 1: "u = gtrace r v !! i" using sset_range imageE assms(2) by metis have 2: "gtarget (stake (Suc i) r) v = u" unfolding 1 sscan_snth by rule have "infinite (sset (sdrop (Suc i) (gtrace r v)))" using gtrace_infinite[OF assms(1)] by (metis List.finite_set finite_Un sset_shift stake_sdrop) also have "sdrop (Suc i) (gtrace r v) = gtrace (sdrop (Suc i) r) (gtarget (stake (Suc i) r) v)" by simp also have "sset \ \ greachable A w V u" using assms(1) 2 by (metis graph.reachable.reflexive graph.reachable_trace graph.run_sdrop) finally show ?thesis by this qed lemma finite_nodes_gsuccessors: assumes "finite (nodes A)" assumes "v \ gunodes A w" shows "finite (gusuccessors A w v)" proof - have "gusuccessors A w v \ gureachable A w v" by rule also have "\ \ gunodes A w" using assms(2) by blast also have "\ \ UNIV \ nodes A" using gnodes_nodes by this finally have 3: "gusuccessors A w v \ UNIV \ nodes A" by this have "gusuccessors A w v \ {Suc (fst v)} \ nodes A" using 3 by auto also have "finite \" using assms(1) by simp finally show ?thesis by this qed end diff --git a/thys/Buchi_Complementation/code/Complementation.sml b/thys/Buchi_Complementation/code/Complementation.sml --- a/thys/Buchi_Complementation/code/Complementation.sml +++ b/thys/Buchi_Complementation/code/Complementation.sml @@ -1,17 +1,17 @@ open Complementation open Automaton fun write_automaton path automaton = let fun t (p, (a, q)) = Int.toString (integer_of_nat p) ^ " " ^ a ^ " " ^ Int.toString (integer_of_nat q) ^ "\n" val out = TextIO.openOut path fun write [] = () | write (x :: xs) = (TextIO.output (out, t x); write xs) - val _ = write (transei automaton) + val _ = write (transitionei automaton) val _ = TextIO.closeOut out in () end val path = hd (CommandLine.arguments ()) (* TODO: maybe we want to optimize the representation, nbae_nba_impl includes list lookups *) val _ = write_automaton path (complement_impl (nbaei_nbai automaton)) diff --git a/thys/LTL_Master_Theorem/Code_Export.thy b/thys/LTL_Master_Theorem/Code_Export.thy --- a/thys/LTL_Master_Theorem/Code_Export.thy +++ b/thys/LTL_Master_Theorem/Code_Export.thy @@ -1,55 +1,55 @@ (* Author: Benedikt Seidl License: BSD *) section \Code export to Standard ML\ theory Code_Export imports "LTL_to_DRA/DRA_Implementation" LTL.Code_Equations "HOL-Library.Code_Target_Numeral" begin subsection \LTL to DRA\ export_code ltlc_to_draei checking declare ltl_to_dra.af_letter\<^sub>F_lifted_semantics [code] declare ltl_to_dra.af_letter\<^sub>G_lifted_semantics [code] declare ltl_to_dra.af_letter\<^sub>\_lifted_semantics [code] definition atoms_ltlc_list_literals :: "String.literal ltlc \ String.literal list" where "atoms_ltlc_list_literals = atoms_ltlc_list" definition ltlc_to_draei_literals :: "String.literal ltlc \ (String.literal set, nat) draei" where "ltlc_to_draei_literals = ltlc_to_draei" definition sort_transitions :: "(nat \ String.literal set \ nat) list \ (nat \ String.literal set \ nat) list" where "sort_transitions = sort_key fst" export_code True_ltlc Iff_ltlc ltlc_to_draei_literals - alphabetei initialei transei acceptingei + alphabetei initialei transitionei acceptingei integer_of_nat atoms_ltlc_list_literals sort_transitions set in SML module_name LTL file_prefix LTL_to_DRA subsection \LTL to NBA\ (* TODO *) subsection \LTL to LDBA\ (* TODO *) end diff --git a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy --- a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy +++ b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Construction.thy @@ -1,440 +1,440 @@ (* Author: Benedikt Seidl License: BSD *) section \Constructing DRAs for LTL Formulas\ theory DRA_Construction imports Transition_Functions "../Quotient_Type" "../Omega_Words_Fun_Stream" "../Logical_Characterization/Master_Theorem" Transition_Systems_and_Automata.DBA_Combine Transition_Systems_and_Automata.DCA_Combine Transition_Systems_and_Automata.DRA_Combine begin \ \We use prefix and suffix on infinite words.\ hide_const Sublist.prefix Sublist.suffix locale dra_construction = transition_functions eq + quotient eq Rep Abs for eq :: "'a ltln \ 'a ltln \ bool" (infix "\" 75) and Rep :: "'ltlq \ 'a ltln" and Abs :: "'a ltln \ 'ltlq" begin subsection \Lifting Setup\ abbreviation true\<^sub>n_lifted :: "'ltlq" ("\true\<^sub>n") where "\true\<^sub>n \ Abs true\<^sub>n" abbreviation false\<^sub>n_lifted :: "'ltlq" ("\false\<^sub>n") where "\false\<^sub>n \ Abs false\<^sub>n" abbreviation af_letter_lifted :: "'a set \ 'ltlq \ 'ltlq" ("\afletter") where "\afletter \ \ \ Abs (af_letter (Rep \) \)" abbreviation af_lifted :: "'ltlq \ 'a set list \ 'ltlq" ("\af") where "\af \ w \ fold \afletter w \" abbreviation GF_advice_lifted :: "'ltlq \ 'a ltln set \ 'ltlq" ("_\[_]\<^sub>\" [90,60] 89) where "\\[X]\<^sub>\ \ Abs ((Rep \)[X]\<^sub>\)" lemma af_letter_lifted_semantics: "\afletter \ (Abs \) = Abs (af_letter \ \)" by (metis Rep_Abs_eq af_letter_congruent Abs_eq) lemma af_lifted_semantics: "\af (Abs \) w = Abs (af \ w)" by (induction w rule: rev_induct) (auto simp: Abs_eq, insert Rep_Abs_eq af_letter_congruent eq_sym, blast) lemma af_lifted_range: "range (\af (Abs \)) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" using af_lifted_semantics af_nested_prop_atoms by blast definition af_letter\<^sub>F_lifted :: "'a ltln \ 'a set \ 'ltlq \ 'ltlq" ("\afletter\<^sub>F") where "\afletter\<^sub>F \ \ \ \ Abs (af_letter\<^sub>F \ (Rep \) \)" definition af_letter\<^sub>G_lifted :: "'a ltln \ 'a set \ 'ltlq \ 'ltlq" ("\afletter\<^sub>G") where "\afletter\<^sub>G \ \ \ \ Abs (af_letter\<^sub>G \ (Rep \) \)" lemma af_letter\<^sub>F_lifted_semantics: "\afletter\<^sub>F \ \ (Abs \) = Abs (af_letter\<^sub>F \ \ \)" by (metis af_letter\<^sub>F_lifted_def Rep_inverse af_letter\<^sub>F_def af_letter_congruent Abs_eq) lemma af_letter\<^sub>G_lifted_semantics: "\afletter\<^sub>G \ \ (Abs \) = Abs (af_letter\<^sub>G \ \ \)" by (metis af_letter\<^sub>G_lifted_def Rep_inverse af_letter\<^sub>G_def af_letter_congruent Abs_eq) abbreviation af\<^sub>F_lifted :: "'a ltln \ 'ltlq \ 'a set list \ 'ltlq" ("\af\<^sub>F") where "\af\<^sub>F \ \ w \ fold (\afletter\<^sub>F \) w \" abbreviation af\<^sub>G_lifted :: "'a ltln \ 'ltlq \ 'a set list \ 'ltlq" ("\af\<^sub>G") where "\af\<^sub>G \ \ w \ fold (\afletter\<^sub>G \) w \" lemma af\<^sub>F_lifted_semantics: "\af\<^sub>F \ (Abs \) w = Abs (af\<^sub>F \ \ w)" by (induction w rule: rev_induct) (auto simp: af_letter\<^sub>F_lifted_semantics) lemma af\<^sub>G_lifted_semantics: "\af\<^sub>G \ (Abs \) w = Abs (af\<^sub>G \ \ w)" by (induction w rule: rev_induct) (auto simp: af_letter\<^sub>G_lifted_semantics) lemma af\<^sub>F_lifted_range: "nested_prop_atoms \ \ nested_prop_atoms (F\<^sub>n \) \ range (\af\<^sub>F \ (Abs \)) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (F\<^sub>n \)}" using af\<^sub>F_lifted_semantics af\<^sub>F_nested_prop_atoms by blast lemma af\<^sub>G_lifted_range: "nested_prop_atoms \ \ nested_prop_atoms (G\<^sub>n \) \ range (\af\<^sub>G \ (Abs \)) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (G\<^sub>n \)}" using af\<^sub>G_lifted_semantics af\<^sub>G_nested_prop_atoms by blast definition af_letter\<^sub>\_lifted :: "'a ltln set \ 'a set \ 'ltlq \ 'ltlq \ 'ltlq \ 'ltlq" ("\afletter\<^sub>\") where "\afletter\<^sub>\ X \ p \ (Abs (fst (af_letter\<^sub>\ X (Rep (fst p), Rep (snd p)) \)), Abs (snd (af_letter\<^sub>\ X (Rep (fst p), Rep (snd p)) \)))" abbreviation af\<^sub>\_lifted :: "'a ltln set \ 'ltlq \ 'ltlq \ 'a set list \ 'ltlq \ 'ltlq" ("\af\<^sub>\") where "\af\<^sub>\ X p w \ fold (\afletter\<^sub>\ X) w p" lemma af_letter\<^sub>\_lifted_semantics: "\afletter\<^sub>\ X \ (Abs x, Abs y) = (Abs (fst (af_letter\<^sub>\ X (x, y) \)), Abs (snd (af_letter\<^sub>\ X (x, y) \)))" by (simp add: af_letter\<^sub>\_def af_letter\<^sub>\_lifted_def) (insert GF_advice_congruent Rep_Abs_eq Rep_inverse af_letter_lifted_semantics eq_trans Abs_eq, blast) lemma af\<^sub>\_lifted_semantics: "\af\<^sub>\ X (Abs \, Abs \) w = (Abs (fst (af\<^sub>\ X (\, \) w)), Abs (snd (af\<^sub>\ X (\, \) w)))" apply (induction w rule: rev_induct) apply (auto simp: af_letter\<^sub>\_lifted_def af_letter\<^sub>\_lifted_semantics af_letter_lifted_semantics) by (metis (no_types, hide_lams) af_letter\<^sub>\_lifted_def af\<^sub>\_fst af_letter\<^sub>\_lifted_semantics eq_fst_iff prod.sel(2)) subsection \Büchi automata for basic languages\ definition \\<^sub>\ :: "'a ltln \ ('a set, 'ltlq) dba" where "\\<^sub>\ \ = dba UNIV (Abs \) \afletter (\\. \ = \true\<^sub>n)" definition \\<^sub>\_GF :: "'a ltln \ ('a set, 'ltlq) dba" where "\\<^sub>\_GF \ = dba UNIV (Abs (F\<^sub>n \)) (\afletter\<^sub>F \) (\\. \ = \true\<^sub>n)" definition \\<^sub>\ :: "'a ltln \ ('a set, 'ltlq) dca" where "\\<^sub>\ \ = dca UNIV (Abs \) \afletter (\\. \ = \false\<^sub>n)" definition \\<^sub>\_FG :: "'a ltln \ ('a set, 'ltlq) dca" where "\\<^sub>\_FG \ = dca UNIV (Abs (G\<^sub>n \)) (\afletter\<^sub>G \) (\\. \ = \false\<^sub>n)" lemma dba_run: "DBA.run (dba UNIV p \ \) (to_stream w) p" unfolding DBA.run_def by (rule transition_system.snth_run) fastforce lemma dca_run: "DCA.run (dca UNIV p \ \) (to_stream w) p" unfolding DCA.run_def by (rule transition_system.snth_run) fastforce lemma \\<^sub>\_language: "\ \ \LTL \ to_stream w \ DBA.language (\\<^sub>\ \) \ w \\<^sub>n \" proof - assume "\ \ \LTL" then have "w \\<^sub>n \ \ (\n. \k\n. af \ (w[0 \ k]) \ true\<^sub>n)" by (meson af_\LTL af_prefix_true le_cases) also have "\ \ (\n. \k\n. af \ (w[0 \ Suc k]) \ true\<^sub>n)" by (meson af_prefix_true le_SucI order_refl) also have "\ \ infs (\\. \ = \true\<^sub>n) (DBA.trace (\\<^sub>\ \) (to_stream w) (Abs \))" - by (simp add: infs_snth \\<^sub>\_def DBA.succ_def af_lifted_semantics Abs_eq[symmetric] af_letter_lifted_semantics) + by (simp add: infs_snth \\<^sub>\_def DBA.transition_def af_lifted_semantics Abs_eq[symmetric] af_letter_lifted_semantics) also have "\ \ to_stream w \ DBA.language (\\<^sub>\ \)" unfolding \\<^sub>\_def dba.initial_def dba.accepting_def by (metis dba_run DBA.language DBA.language_elim dba.sel(2) dba.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_GF_language: "\ \ \LTL \ to_stream w \ DBA.language (\\<^sub>\_GF \) \ w \\<^sub>n G\<^sub>n (F\<^sub>n \)" proof - assume "\ \ \LTL" then have "w \\<^sub>n G\<^sub>n (F\<^sub>n \) \ (\n. \k. af (F\<^sub>n \) (w[n \ k]) \\<^sub>L true\<^sub>n)" using ltl_lang_equivalence.af_\LTL_GF by blast also have "\ \ (\n. \k>n. af\<^sub>F \ (F\<^sub>n \) (w[0 \ k]) \ true\<^sub>n)" using af\<^sub>F_semantics_ltr af\<^sub>F_semantics_rtl using \\ \ \LTL\ af_\LTL_GF calculation by blast also have "\ \ (\n. \k\n. af\<^sub>F \ (F\<^sub>n \) (w[0 \ Suc k]) \ true\<^sub>n)" by (metis less_Suc_eq_le less_imp_Suc_add) also have "\ \ infs (\\. \ = \true\<^sub>n) (DBA.trace (\\<^sub>\_GF \) (to_stream w) (Abs (F\<^sub>n \)))" - by (simp add: infs_snth \\<^sub>\_GF_def DBA.succ_def af\<^sub>F_lifted_semantics Abs_eq[symmetric] af_letter\<^sub>F_lifted_semantics) + by (simp add: infs_snth \\<^sub>\_GF_def DBA.transition_def af\<^sub>F_lifted_semantics Abs_eq[symmetric] af_letter\<^sub>F_lifted_semantics) also have "\ \ to_stream w \ DBA.language (\\<^sub>\_GF \)" unfolding \\<^sub>\_GF_def dba.initial_def dba.accepting_def by (metis dba_run DBA.language DBA.language_elim dba.sel(2) dba.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_language: "\ \ \LTL \ to_stream w \ DCA.language (\\<^sub>\ \) \ w \\<^sub>n \" proof - assume "\ \ \LTL" then have "w \\<^sub>n \ \ (\n. \k\n. \ af \ (w[0 \ k]) \ false\<^sub>n)" by (meson af_\LTL af_prefix_false le_cases order_refl) also have "\ \ (\n. \k\n. \ af \ (w[0 \ Suc k]) \ false\<^sub>n)" by (meson af_prefix_false le_SucI order_refl) also have "\ \ fins (\\. \ = \false\<^sub>n) (DCA.trace (\\<^sub>\ \) (to_stream w) (Abs \))" - by (simp add: infs_snth \\<^sub>\_def DBA.succ_def af_lifted_semantics Abs_eq[symmetric] af_letter_lifted_semantics) + by (simp add: infs_snth \\<^sub>\_def DBA.transition_def af_lifted_semantics Abs_eq[symmetric] af_letter_lifted_semantics) also have "\ \ to_stream w \ DCA.language (\\<^sub>\ \)" unfolding \\<^sub>\_def dca.initial_def dca.rejecting_def by (metis dca_run DCA.language DCA.language_elim dca.sel(2) dca.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_FG_language: "\ \ \LTL \ to_stream w \ DCA.language (\\<^sub>\_FG \) \ w \\<^sub>n F\<^sub>n (G\<^sub>n \)" proof - assume "\ \ \LTL" then have "w \\<^sub>n F\<^sub>n (G\<^sub>n \) \ (\k. \j. \ af (G\<^sub>n \) (w[k \ j]) \\<^sub>L false\<^sub>n)" using ltl_lang_equivalence.af_\LTL_FG by blast also have "\ \ (\n. \k>n. \ af\<^sub>G \ (G\<^sub>n \) (w[0 \ k]) \ false\<^sub>n)" using af\<^sub>G_semantics_ltr af\<^sub>G_semantics_rtl using \\ \ \LTL\ af_\LTL_FG calculation by blast also have "\ \ (\n. \k\n. \ af\<^sub>G \ (G\<^sub>n \) (w[0 \ Suc k]) \ false\<^sub>n)" by (metis less_Suc_eq_le less_imp_Suc_add) also have "\ \ fins (\\. \ = \false\<^sub>n) (DCA.trace (\\<^sub>\_FG \) (to_stream w) (Abs (G\<^sub>n \)))" - by (simp add: infs_snth \\<^sub>\_FG_def DBA.succ_def af\<^sub>G_lifted_semantics Abs_eq[symmetric] af_letter\<^sub>G_lifted_semantics) + by (simp add: infs_snth \\<^sub>\_FG_def DBA.transition_def af\<^sub>G_lifted_semantics Abs_eq[symmetric] af_letter\<^sub>G_lifted_semantics) also have "\ \ to_stream w \ DCA.language (\\<^sub>\_FG \)" unfolding \\<^sub>\_FG_def dca.initial_def dca.rejecting_def by (metis dca_run DCA.language DCA.language_elim dca.sel(2) dca.sel(4)) finally show ?thesis by simp qed lemma \\<^sub>\_nodes: "DBA.nodes (\\<^sub>\ \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" unfolding \\<^sub>\_def transition_system_initial.nodes_alt_def using af_lifted_semantics af_nested_prop_atoms by fastforce lemma \\<^sub>\_GF_nodes: "DBA.nodes (\\<^sub>\_GF \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (F\<^sub>n \)}" unfolding \\<^sub>\_GF_def DBA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def using af\<^sub>F_nested_prop_atoms[of "F\<^sub>n \"] by (auto simp: af\<^sub>F_lifted_semantics) lemma \\<^sub>\_nodes: "DCA.nodes (\\<^sub>\ \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" unfolding \\<^sub>\_def transition_system_initial.nodes_alt_def using af_lifted_semantics af_nested_prop_atoms by fastforce lemma \\<^sub>\_FG_nodes: "DCA.nodes (\\<^sub>\_FG \) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms (G\<^sub>n \)}" unfolding \\<^sub>\_FG_def DCA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def using af\<^sub>G_nested_prop_atoms[of "G\<^sub>n \"] by (auto simp: af\<^sub>G_lifted_semantics) subsection \A DCA checking the GF-advice Function\ definition \ :: "'a ltln \ 'a ltln set \ ('a set, 'ltlq \ 'ltlq) dca" where "\ \ X = dca UNIV (Abs \, Abs (\[X]\<^sub>\)) (\afletter\<^sub>\ X) (\p. snd p = \false\<^sub>n)" lemma \_language: "to_stream w \ DCA.language (\ \ X) \ (\i. suffix i w \\<^sub>n af \ (prefix i w)[X]\<^sub>\)" proof - have "(\i. suffix i w \\<^sub>n af \ (prefix i w)[X]\<^sub>\) \ (\m. \k\m. \ snd (af\<^sub>\ X (\, \[X]\<^sub>\) (prefix (Suc k) w)) \ false\<^sub>n)" using af\<^sub>\_semantics_ltr af\<^sub>\_semantics_rtl by blast also have "\ \ fins (\p. snd p = \false\<^sub>n) (DCA.trace (\ \ X) (to_stream w) (Abs \, Abs (\[X]\<^sub>\)))" - by(simp add: infs_snth \_def DCA.succ_def af\<^sub>\_lifted_semantics af_letter\<^sub>\_lifted_semantics Abs_eq) + by(simp add: infs_snth \_def DCA.transition_def af\<^sub>\_lifted_semantics af_letter\<^sub>\_lifted_semantics Abs_eq) also have "\ \ to_stream w \ DCA.language (\ \ X)" by (simp add: \_def dca.initial_def dca.rejecting_def DCA.language_def dca_run) finally show ?thesis by blast qed lemma \_nodes: "DCA.nodes (\ \ X) \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms \} \ {Abs \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X}" unfolding \_def DCA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def apply (auto simp add: af\<^sub>\_lifted_semantics af_letter\<^sub>\_lifted_semantics) using af\<^sub>\_fst_nested_prop_atoms apply force using GF_advice_nested_prop_atoms\<^sub>\ af\<^sub>\_snd_nested_prop_atoms dra_construction_axioms by fastforce subsection \A DRA for each combination of sets X and Y\ lemma dba_language: "(\w. to_stream w \ DBA.language \ \ w \\<^sub>n \) \ DBA.language \ = {w. to_omega w \\<^sub>n \}" by (metis (mono_tags, lifting) Collect_cong DBA.language_def mem_Collect_eq to_stream_to_omega) lemma dca_language: "(\w. to_stream w \ DCA.language \ \ w \\<^sub>n \) \ DCA.language \ = {w. to_omega w \\<^sub>n \}" by (metis (mono_tags, lifting) Collect_cong DCA.language_def mem_Collect_eq to_stream_to_omega) definition \\<^sub>1 :: "'a ltln \ 'a ltln list \ ('a set, 'ltlq \ 'ltlq) dca" where "\\<^sub>1 \ xs = \ \ (set xs)" lemma \\<^sub>1_language: "to_omega ` DCA.language (\\<^sub>1 \ xs) = L\<^sub>1 \ (set xs)" by (simp add: \\<^sub>1_def L\<^sub>1_def set_eq_iff \_language) lemma \\<^sub>1_alphabet: "DCA.alphabet (\\<^sub>1 \ xs) = UNIV" unfolding \\<^sub>1_def \_def by simp definition \\<^sub>2 :: "'a ltln list \ 'a ltln list \ ('a set, 'ltlq list degen) dba" where "\\<^sub>2 xs ys = dbail (map (\\. \\<^sub>\_GF (\[set ys]\<^sub>\)) xs)" lemma \\<^sub>2_language: "to_omega ` DBA.language (\\<^sub>2 xs ys) = L\<^sub>2 (set xs) (set ys)" by (simp add: \\<^sub>2_def L\<^sub>2_def set_eq_iff dba_language[OF \\<^sub>\_GF_language[OF FG_advice_\LTL(1)]]) lemma \\<^sub>2_alphabet: "DBA.alphabet (\\<^sub>2 xs ys) = UNIV" by (simp add: \\<^sub>2_def \\<^sub>\_GF_def dbail_def dbgail_def) definition \\<^sub>3 :: "'a ltln list \ 'a ltln list \ ('a set, 'ltlq list) dca" where "\\<^sub>3 xs ys = dcail (map (\\. \\<^sub>\_FG (\[set xs]\<^sub>\)) ys)" lemma \\<^sub>3_language: "to_omega ` DCA.language (\\<^sub>3 xs ys) = L\<^sub>3 (set xs) (set ys)" by (simp add: \\<^sub>3_def L\<^sub>3_def set_eq_iff dca_language[OF \\<^sub>\_FG_language[OF GF_advice_\LTL(1)]]) lemma \\<^sub>3_alphabet: "DCA.alphabet (\\<^sub>3 xs ys) = UNIV" by (simp add: \\<^sub>3_def \\<^sub>\_FG_def dcail_def) definition "\' \ xs ys = dbcrai (\\<^sub>2 xs ys) (dcai (\\<^sub>1 \ xs) (\\<^sub>3 xs ys))" lemma \'_language: "to_omega ` DRA.language (\' \ xs ys) = (L\<^sub>1 \ (set xs) \ L\<^sub>2 (set xs) (set ys) \ L\<^sub>3 (set xs) (set ys))" by (simp add: \'_def \\<^sub>1_language \\<^sub>2_language \\<^sub>3_language) fastforce lemma \'_alphabet: "DRA.alphabet (\' \ xs ys) = UNIV" by (simp add: \'_def dbcrai_def dcai_def \\<^sub>1_alphabet \\<^sub>2_alphabet \\<^sub>3_alphabet) subsection \A DRA for @{term "L(\)"}\ definition "ltl_to_dra \ = draul (map (\(xs, ys). \' \ xs ys) (advice_sets \))" lemma ltl_to_dra_language: "to_omega ` DRA.language (ltl_to_dra \) = language_ltln \" proof - have 1: "INTER (set (map (\(x, y). \' \ x y) (advice_sets \))) dra.alphabet = UNION (set (map (\(x, y). \' \ x y) (advice_sets \))) dra.alphabet" by (induction "advice_sets \") (metis advice_sets_not_empty, simp add: \'_alphabet split_def advice_sets_not_empty) have "language_ltln \ = \ {(L\<^sub>1 \ X \ L\<^sub>2 X Y \ L\<^sub>3 X Y) | X Y. X \ subformulas\<^sub>\ \ \ Y \ subformulas\<^sub>\ \}" unfolding master_theorem_language by auto also have "\ = \ {L\<^sub>1 \ (set xs) \ L\<^sub>2 (set xs) (set ys) \ L\<^sub>3 (set xs) (set ys) | xs ys. (xs, ys) \ set (advice_sets \)}" unfolding advice_sets_subformulas by (metis (no_types, lifting)) also have "\ = \ {to_omega ` DRA.language (\' \ xs ys) | xs ys. (xs, ys) \ set (advice_sets \)}" by (simp add: \'_language) finally show ?thesis unfolding ltl_to_dra_def draul_language[OF 1] by auto qed lemma ltl_to_dra_alphabet: "alphabet (ltl_to_dra \) = UNIV" by (auto simp: ltl_to_dra_def draul_def \'_alphabet split: prod.split) (insert advice_sets_subformulas, blast) subsection \Setting the Alphabet of a DRA\ definition dra_set_alphabet :: "('a set, 'b) dra \ 'a set set \ ('a set, 'b) dra" where - "dra_set_alphabet \ \ = dra \ (initial \) (succ \) (accepting \)" + "dra_set_alphabet \ \ = dra \ (initial \) (transition \) (accepting \)" lemma dra_set_alphabet_language: "\ \ alphabet \ \ language (dra_set_alphabet \ \) = language \ \ {s. sset s \ \}" by (auto simp add: dra_set_alphabet_def language_def set_eq_iff DRA.run_alt_def) lemma dra_set_alphabet_alphabet[simp]: "alphabet (dra_set_alphabet \ \) = \" unfolding dra_set_alphabet_def by simp lemma dra_set_alphabet_nodes: "\ \ alphabet \ \ DRA.nodes (dra_set_alphabet \ \) \ DRA.nodes \" unfolding dra_set_alphabet_def DRA.nodes_def transition_system_initial.nodes_alt_def transition_system.reachable_alt_def by auto (metis DRA.path_alt_def DRA.path_def dra.sel(1) dra.sel(3) subset_trans) subsection \A DRA for @{term "L(\)"} with a finite alphabet\ definition "ltl_to_dra_alphabet \ Ap = dra_set_alphabet (ltl_to_dra \) (Pow Ap)" lemma ltl_to_dra_alphabet_language: assumes "atoms_ltln \ \ Ap" shows "to_omega ` language (ltl_to_dra_alphabet \ Ap) = language_ltln \ \ {w. range w \ Pow Ap}" proof - have 1: "Pow Ap \ alphabet (ltl_to_dra \)" unfolding ltl_to_dra_alphabet by simp show ?thesis unfolding ltl_to_dra_alphabet_def dra_set_alphabet_language[OF 1] by (simp add: ltl_to_dra_language sset_range) force qed lemma ltl_to_dra_alphabet_alphabet[simp]: "alphabet (ltl_to_dra_alphabet \ Ap) = Pow Ap" unfolding ltl_to_dra_alphabet_def by simp lemma ltl_to_dra_alphabet_nodes: "DRA.nodes (ltl_to_dra_alphabet \ Ap) \ DRA.nodes (ltl_to_dra \)" unfolding ltl_to_dra_alphabet_def by (rule dra_set_alphabet_nodes) (simp add: ltl_to_dra_alphabet) end end diff --git a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Implementation.thy b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Implementation.thy --- a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Implementation.thy +++ b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Implementation.thy @@ -1,132 +1,132 @@ (* Author: Benedikt Seidl License: BSD *) section \Implementation of the DRA Construction\ theory DRA_Implementation imports DRA_Instantiation LTL.Rewriting Transition_Systems_and_Automata.DRA_Translate begin text \We convert the implicit automaton to its explicit representation and afterwards proof the final correctness theorem and the overall size bound.\ subsection \Generating the Explicit Automaton\ definition dra_to_drai :: "('a, 'b) dra \ 'a list \ ('a, 'b) drai" where - "dra_to_drai \ \ = drai \ (initial \) (succ \) (accepting \)" + "dra_to_drai \ \ = drai \ (initial \) (transition \) (accepting \)" lemma dra_to_drai_language: "set \ = alphabet \ \ language (drai_dra (dra_to_drai \ \)) = language \" by (simp add: dra_to_drai_def drai_dra_def) definition drai_to_draei :: "('a, 'b) drai \ ('a, nat) draei" where "drai_to_draei = to_draei_impl (=) bot 2" fun atoms_ltlc_list :: "'a ltlc \ 'a list" where "atoms_ltlc_list true\<^sub>c = []" | "atoms_ltlc_list false\<^sub>c = []" | "atoms_ltlc_list prop\<^sub>c(q) = [q]" | "atoms_ltlc_list (not\<^sub>c \) = atoms_ltlc_list \" | "atoms_ltlc_list (\ and\<^sub>c \) = List.union (atoms_ltlc_list \) (atoms_ltlc_list \)" | "atoms_ltlc_list (\ or\<^sub>c \) = List.union (atoms_ltlc_list \) (atoms_ltlc_list \)" | "atoms_ltlc_list (\ implies\<^sub>c \) = List.union (atoms_ltlc_list \) (atoms_ltlc_list \)" | "atoms_ltlc_list (X\<^sub>c \) = atoms_ltlc_list \" | "atoms_ltlc_list (F\<^sub>c \) = atoms_ltlc_list \" | "atoms_ltlc_list (G\<^sub>c \) = atoms_ltlc_list \" | "atoms_ltlc_list (\ U\<^sub>c \) = List.union (atoms_ltlc_list \) (atoms_ltlc_list \)" | "atoms_ltlc_list (\ R\<^sub>c \) = List.union (atoms_ltlc_list \) (atoms_ltlc_list \)" | "atoms_ltlc_list (\ W\<^sub>c \) = List.union (atoms_ltlc_list \) (atoms_ltlc_list \)" | "atoms_ltlc_list (\ M\<^sub>c \) = List.union (atoms_ltlc_list \) (atoms_ltlc_list \)" lemma atoms_ltlc_list_set: "set (atoms_ltlc_list \) = atoms_ltlc \" by (induction \) simp_all lemma atoms_ltlc_list_distinct: "distinct (atoms_ltlc_list \)" by (induction \) simp_all definition ltl_alphabet :: "'a ltlc \ 'a set list" where "ltl_alphabet \ = map set (subseqs (atoms_ltlc_list \))" definition ltlc_to_draei :: "'a ltlc \ ('a set, nat) draei" where "ltlc_to_draei \ = drai_to_draei (dra_to_drai (ltl_to_dra_alphabet (simplify Slow (ltlc_to_ltln \)) (atoms_ltlc \)) (ltl_alphabet \))" subsection \Final Proof of Correctness\ lemma dra_to_drai_rel: assumes "(\, alphabet A) \ \Id\ list_set_rel" shows "(dra_to_drai A \, A) \ \Id, Id\drai_dra_rel" proof - have "(A, A) \ \Id, Id\dra_rel" by simp - then have "(dra_to_drai A \, dra (alphabet A) (initial A) (succ A) (accepting A)) \ \Id, Id\drai_dra_rel" + then have "(dra_to_drai A \, dra (alphabet A) (initial A) (transition A) (accepting A)) \ \Id, Id\drai_dra_rel" unfolding dra_to_drai_def using assms by parametricity then show ?thesis by simp qed lemma draei_language: assumes 1: "(\, alphabet A) \ \Id\ list_set_rel" and "finite (DRA.nodes A)" shows "DRA.language (drae_dra (draei_drae (drai_to_draei (dra_to_drai A \)))) = DRA.language A" proof - have "(dra_to_drai A \, A) \ \Id, Id\drai_dra_rel" using dra_to_drai_rel 1 by fast then have "(drai_to_draei (dra_to_drai A \), to_draei A) \ \Id_on (dra.alphabet A), rel (dra_to_drai A \) A (=) bot 2\ draei_dra_rel" unfolding drai_to_draei_def using to_draei_impl_refine[unfolded autoref_tag_defs] apply parametricity by (simp_all add: assms is_bounded_hashcode_def bot_nat_def is_valid_def_hm_size_def) then have "(DRA.language ((drae_dra \ draei_drae) (drai_to_draei (dra_to_drai A \))), DRA.language (id (to_draei A))) \ \\Id_on (dra.alphabet A)\ stream_rel\ set_rel" by parametricity then show ?thesis by (simp add: to_draei_def) qed lemma ltl_to_dra_alphabet_rel: "(ltl_alphabet \, alphabet (ltl_to_dra_alphabet \ (atoms_ltlc \))) \ \Id\ list_set_rel" unfolding ltl_to_dra.ltl_to_dra_alphabet_alphabet ltl_alphabet_def by (simp add: list_set_rel_def atoms_ltlc_list_set atoms_ltlc_list_distinct in_br_conv subseqs_powset distinct_set_subseqs) lemma ltl_to_dra_alphabet_nodes_finite: "finite (DRA.nodes (ltl_to_dra_alphabet \ Ap))" using ltl_to_dra.ltl_to_dra_alphabet_nodes ltl_to_dra_nodes_finite finite_subset by fast lemma ltlc_to_ltln_simplify_atoms: "atoms_ltln (simplify Slow (ltlc_to_ltln \)) \ atoms_ltlc \" using ltlc_to_ltln_atoms simplify_atoms by fast theorem final_correctness: "to_omega ` language (drae_dra (draei_drae (ltlc_to_draei \))) = language_ltlc \ \ {w. range w \ Pow (atoms_ltlc \)}" unfolding ltlc_to_draei_def unfolding draei_language[OF ltl_to_dra_alphabet_rel ltl_to_dra_alphabet_nodes_finite] unfolding ltl_to_dra.ltl_to_dra_alphabet_language[OF ltlc_to_ltln_simplify_atoms] unfolding ltlc_to_ltln_atoms language_ltln_def language_ltlc_def ltlc_to_ltln_semantics simplify_correct .. end diff --git a/thys/LTL_Master_Theorem/code/hoa.sml b/thys/LTL_Master_Theorem/code/hoa.sml --- a/thys/LTL_Master_Theorem/code/hoa.sml +++ b/thys/LTL_Master_Theorem/code/hoa.sml @@ -1,136 +1,136 @@ (* Serialization of DRAs in the HOA format. Benedikt Seidl *) open LTL structure HOA : sig val serialize : string ltlc -> (string set, nat) draei -> string end = struct (* Metadata *) val tool = "LTL to DRA translation based on the Master Theorem" val version = "0.1" val properties = ["transition-labels", "state-acc", "deterministic"] val acc_name = "Rabin" fun quote s = "\"" ^ s ^ "\"" fun bracket true s = "(" ^ s ^ ")" | bracket false s = s fun nat_to_string n = IntInf.toString (integer_of_nat n) fun ltlc_to_string b True_ltlc = "true" | ltlc_to_string b False_ltlc = "false" | ltlc_to_string b (Prop_ltlc s) = s | ltlc_to_string b (Not_ltlc x) = "!" ^ ltlc_to_string true x | ltlc_to_string b (And_ltlc (x1, x2)) = bracket b (ltlc_to_string true x1 ^ " & " ^ ltlc_to_string true x2) | ltlc_to_string b (Or_ltlc (x1, x2)) = bracket b (ltlc_to_string true x1 ^ " | " ^ ltlc_to_string true x2) | ltlc_to_string b (Implies_ltlc (x1, x2)) = bracket b (ltlc_to_string true x1 ^ " -> " ^ ltlc_to_string true x2) | ltlc_to_string b (Next_ltlc x) = bracket b ("X " ^ ltlc_to_string true x) | ltlc_to_string b (Final_ltlc x) = bracket b ("F " ^ ltlc_to_string true x) | ltlc_to_string b (Global_ltlc x) = bracket b ("G " ^ ltlc_to_string true x) | ltlc_to_string b (Until_ltlc (x1, x2)) = bracket b (ltlc_to_string true x1 ^ " U " ^ ltlc_to_string true x2) | ltlc_to_string b (Release_ltlc (x1, x2)) = bracket b (ltlc_to_string true x1 ^ " R " ^ ltlc_to_string true x2) | ltlc_to_string b (WeakUntil_ltlc (x1, x2)) = bracket b (ltlc_to_string true x1 ^ " W " ^ ltlc_to_string true x2) | ltlc_to_string b (StrongRelease_ltlc (x1, x2)) = bracket b (ltlc_to_string true x1 ^ " M " ^ ltlc_to_string true x2) fun mapi f xs = let fun inner (x, (i, xs)) = (i + 1, f (i, x) :: xs) in List.rev (#2 (foldl inner (0, []) xs)) end (* TODO formalize in Isabelle *) fun group_sorted f xs = let fun inner (t, ([], acc)) = ([t], acc) | inner (t, (x :: xs, acc)) = if f (t, x) then ([t], (x :: xs) :: acc) else (t :: x :: xs, acc) val (grp, acc) = foldl inner ([], []) xs in case grp of [] => acc | x :: xs => (x :: xs) :: acc end (* Header *) fun serialize_aps aps = Int.toString (List.length aps) ^ " " ^ String.concatWith " " (List.map quote aps) fun serialize_rabin_pair (i, _) = "Inf(" ^ Int.toString (2 * i) ^ ") & Fin(" ^ Int.toString (2 * i + 1) ^ ")" fun serialize_acceptance acc = Int.toString (2 * List.length acc) ^ " " ^ String.concatWith " | " (mapi serialize_rabin_pair acc) fun serialize_header phi aut states = "HOA: v1\n" ^ "tool: " ^ quote tool ^ " " ^ quote version ^ "\n" ^ "name: " ^ quote ("DRA for " ^ ltlc_to_string false phi) ^ "\n" ^ "properties: " ^ String.concatWith " " properties ^ "\n" ^ "States: " ^ Int.toString (List.length states) ^ "\n" ^ "Start: " ^ nat_to_string (initialei aut) ^ "\n" ^ "AP: " ^ serialize_aps (atoms_ltlc_list_literals phi) ^ "\n" ^ "Acceptance: " ^ serialize_acceptance (acceptingei aut) ^ "\n" ^ "acc-name: " ^ acc_name ^ "\n" (* Body *) fun iterate_aps phi f = let fun inner (i, a) = if f a then Int.toString i else "!" ^ Int.toString i in case atoms_ltlc_list_literals phi of [] => "[t]" | xs => "[" ^ String.concatWith "&" (mapi inner xs) ^ "]" end fun serialize_label phi (Set xs) = iterate_aps phi (fn a => List.exists (fn b => a = b) xs) | serialize_label phi (Coset xs) = iterate_aps phi (fn a => List.all (fn b => a <> b) xs) fun serialize_transition phi (from, (label, to)) = serialize_label phi label ^ " " ^ nat_to_string to ^ "\n" fun serialize_state_labels acc state = let fun inner (i, (inf, fin)) = (if (List.exists (fn j => state = j) inf) then [Int.toString (2 * i)] else []) @ (if (List.exists (fn j => state = j) fin) then [Int.toString (2 * i + 1)] else []) in "{" ^ String.concatWith " " (List.concat (mapi inner acc)) ^ "}" end fun serialize_state phi aut state = "State: " ^ nat_to_string (#1 (List.hd state)) ^ " " ^ serialize_state_labels (acceptingei aut) (#1 (List.hd state)) ^ "\n" ^ String.concat (List.map (serialize_transition phi) state) fun serialize_body phi aut states = String.concatWith "\n" (List.map (serialize_state phi aut) states) (* Automaton *) fun serialize phi aut = let val states = group_sorted (fn (x, y) => integer_of_nat (#1 x) < - integer_of_nat (#1 y)) (List.rev (sort_transitions (transei aut))) + integer_of_nat (#1 y)) (List.rev (sort_transitions (transitionei aut))) in serialize_header phi aut states ^ "\n--BODY--\n" ^ serialize_body phi aut states ^ "--END--" end end; diff --git a/thys/Transition_Systems_and_Automata/Automata/DBA/DBA.thy b/thys/Transition_Systems_and_Automata/Automata/DBA/DBA.thy --- a/thys/Transition_Systems_and_Automata/Automata/DBA/DBA.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DBA/DBA.thy @@ -1,62 +1,62 @@ section \Deterministic Büchi Automata\ theory DBA imports "../../Basic/Sequence_Zip" "../../Basic/Acceptance" "../../Transition_Systems/Transition_System" "../../Transition_Systems/Transition_System_Extra" "../../Transition_Systems/Transition_System_Construction" begin datatype ('label, 'state) dba = dba (alphabet: "'label set") (initial: "'state") - (succ: "'label \ 'state \ 'state") + (transition: "'label \ 'state \ 'state") (accepting: "'state pred") global_interpretation dba: transition_system_initial - "succ A" "\ a p. a \ alphabet A" "\ p. p = initial A" + "transition A" "\ a p. a \ alphabet A" "\ p. p = initial A" for A defines path = dba.path and run = dba.run and reachable = dba.reachable and nodes = dba.nodes and enableds = dba.enableds and paths = dba.paths and runs = dba.runs by this abbreviation target where "target \ dba.target" abbreviation states where "states \ dba.states" abbreviation trace where "trace \ dba.trace" abbreviation successors :: "('label, 'state) dba \ 'state \ 'state set" where "successors \ dba.successors TYPE('label)" lemma path_alt_def: "path A w p \ set w \ alphabet A" unfolding lists_iff_set[symmetric] proof show "w \ lists (alphabet A)" if "path A w p" using that by (induct arbitrary: p) (auto) show "path A w p" if "w \ lists (alphabet A)" using that by (induct arbitrary: p) (auto) qed lemma run_alt_def: "run A w p \ sset w \ alphabet A" unfolding streams_iff_sset[symmetric] proof show "w \ streams (alphabet A)" if "run A w p" using that by (coinduction arbitrary: w p) (force elim: dba.run.cases) show "run A w p" if "w \ streams (alphabet A)" using that by (coinduction arbitrary: w p) (force elim: streams.cases) qed definition language :: "('label, 'state) dba \ 'label stream set" where "language A \ {w. run A w (initial A) \ infs (accepting A) (trace A w (initial A))}" lemma language[intro]: assumes "run A w (initial A)" "infs (accepting A) (trace A w (initial A))" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains "run A w (initial A)" "infs (accepting A) (trace A w (initial A))" using assms unfolding language_def by auto lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def run_alt_def using sset_streams by auto end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DBA/DBA_Combine.thy @@ -1,303 +1,303 @@ section \Deterministic Büchi Automata Combinations\ theory DBA_Combine imports "DBA" "DGBA" begin definition dbgail :: "('label, 'state) dba list \ ('label, 'state list) dgba" where "dbgail AA \ dgba (INTER (set AA) dba.alphabet) (map dba.initial AA) - (\ a pp. map2 (\ A p. dba.succ A a p) AA pp) + (\ a pp. map2 (\ A p. dba.transition A a p) AA pp) (map (\ k pp. dba.accepting (AA ! k) (pp ! k)) [0 ..< length AA])" lemma dbgail_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w pp) = dba.trace (AA ! k) w (pp ! k)" using assms unfolding dbgail_def by (coinduction arbitrary: w pp) (force) lemma dbgail_nodes_length: assumes "pp \ DGBA.nodes (dbgail AA)" shows "length pp = length AA" using assms unfolding dbgail_def by induct auto lemma dbgail_nodes[intro]: assumes "pp \ DGBA.nodes (dbgail AA)" "k < length pp" shows "pp ! k \ DBA.nodes (AA ! k)" using assms unfolding dbgail_def by induct auto lemma dbgail_nodes_finite[intro]: assumes "list_all (finite \ DBA.nodes) AA" shows "finite (DGBA.nodes (dbgail AA))" proof (rule finite_subset) show "DGBA.nodes (dbgail AA) \ listset (map DBA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dbgail_nodes_length) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms by (simp add: list.pred_map) qed lemma dbgail_nodes_card: assumes "list_all (finite \ DBA.nodes) AA" shows "card (DGBA.nodes (dbgail AA)) \ prod_list (map (card \ DBA.nodes) AA)" proof - have "card (DGBA.nodes (dbgail AA)) \ card (listset (map DBA.nodes AA))" proof (rule card_mono) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms by (simp add: list.pred_map) show "DGBA.nodes (dbgail AA) \ listset (map DBA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dbgail_nodes_length) qed also have "\ = prod_list (map (card \ DBA.nodes) AA)" by simp finally show ?thesis by this qed lemma dbgail_language[simp]: "DGBA.language (dbgail AA) = INTER (set AA) DBA.language" proof safe fix w A assume 1: "w \ DGBA.language (dbgail AA)" "A \ set AA" obtain 2: "dgba.run (dbgail AA) w (dgba.initial (dbgail AA))" "gen infs (dgba.accepting (dbgail AA)) (dgba.trace (dbgail AA) w (dgba.initial (dbgail AA)))" using 1(1) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(2) unfolding in_set_conv_nth by auto have 4: "(\ pp. dba.accepting A (pp ! k)) \ set (dgba.accepting (dbgail AA))" using 3 unfolding dbgail_def by auto show "w \ DBA.language A" proof show "dba.run A w (dba.initial A)" using 1(2) 2(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbgail_def by auto have "True \ infs (\ pp. dba.accepting A (pp ! k)) (dgba.trace (dbgail AA) w (map dba.initial AA))" using 2(2) 4 unfolding dbgail_def by auto also have "\ \ infs (dba.accepting A) (smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w (map dba.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w (map dba.initial AA)) = dba.trace (AA ! k) w (map dba.initial AA ! k)" using 3(2) by (fastforce intro: dbgail_trace_smap) also have "\ = dba.trace A w (dba.initial A)" using 3 by auto finally show "infs (dba.accepting A) (dba.trace A w (dba.initial A))" by simp qed next fix w assume 1: "w \ INTER (set AA) DBA.language" have 2: "dba.run A w (dba.initial A)" "infs (dba.accepting A) (dba.trace A w (dba.initial A))" if "A \ set AA" for A using 1 that by auto show "w \ DGBA.language (dbgail AA)" proof (intro DGBA.language ballI gen) show "dgba.run (dbgail AA) w (dgba.initial (dbgail AA))" using 2(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbgail_def by auto next fix P assume 3: "P \ set (dgba.accepting (dbgail AA))" obtain k where 4: "P = (\ pp. dba.accepting (AA ! k) (pp ! k))" "k < length AA" using 3 unfolding dbgail_def by auto have "True \ infs (dba.accepting (AA ! k)) (dba.trace (AA ! k) w (map dba.initial AA ! k))" using 2(2) 4(2) by auto also have "dba.trace (AA ! k) w (map dba.initial AA ! k) = smap (\ pp. pp ! k) (dgba.trace (dbgail AA) w (map dba.initial AA))" using 4(2) by (fastforce intro: dbgail_trace_smap[symmetric]) also have "infs (dba.accepting (AA ! k)) \ \ infs P (dgba.trace (dbgail AA) w (map dba.initial AA))" unfolding 4(1) by (simp add: comp_def) also have "map dba.initial AA = dgba.initial (dbgail AA)" unfolding dbgail_def by simp finally show "infs P (dgba.trace (dbgail AA) w (dgba.initial (dbgail AA)))" by simp qed qed definition dbail :: "('label, 'state) dba list \ ('label, 'state list degen) dba" where "dbail = dgbad \ dbgail" lemma dbail_nodes_finite[intro]: assumes "list_all (finite \ DBA.nodes) AA" shows "finite (DBA.nodes (dbail AA))" using dbgail_nodes_finite assms unfolding dbail_def by auto lemma dbail_nodes_card: assumes"list_all (finite \ DBA.nodes) AA" shows "card (DBA.nodes (dbail AA)) \ max 1 (length AA) * prod_list (map (card \ DBA.nodes) AA)" proof - have "card (DBA.nodes (dbail AA)) \ max 1 (length (dgba.accepting (dbgail AA))) * card (DGBA.nodes (dbgail AA))" unfolding dbail_def using dgbad_nodes_card by simp also have "length (dgba.accepting (dbgail AA)) = length AA" unfolding dbgail_def by simp also have "card (DGBA.nodes (dbgail AA)) \ prod_list (map (card \ DBA.nodes) AA)" using dbgail_nodes_card assms by this finally show ?thesis by simp qed lemma dbail_language[simp]: "DBA.language (dbail AA) = INTER (set AA) DBA.language" unfolding dbail_def using dgbad_language dbgail_language by auto definition dbau :: "('label, 'state\<^sub>1) dba \ ('label, 'state\<^sub>2) dba \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) dba" where "dbau A B \ dba (dba.alphabet A \ dba.alphabet B) (dba.initial A, dba.initial B) - (\ a (p, q). (dba.succ A a p, dba.succ B a q)) + (\ a (p, q). (dba.transition A a p, dba.transition B a q)) (\ (p, q). dba.accepting A p \ dba.accepting B q)" (* TODO: can these be extracted as more general theorems about sscan? *) lemma dbau_fst[iff]: "infs (P \ fst) (dba.trace (dbau A B) w (p, q)) \ infs P (dba.trace A w p)" proof - let ?t = "dba.trace (dbau A B) w (p, q)" have "infs (P \ fst) ?t \ infs P (smap fst ?t)" by (simp add: comp_def) also have "smap fst ?t = dba.trace A w p" unfolding dbau_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbau_snd[iff]: "infs (P \ snd) (dba.trace (dbau A B) w (p, q)) \ infs P (dba.trace B w q)" proof - let ?t = "dba.trace (dbau A B) w (p, q)" have "infs (P \ snd) ?t \ infs P (smap snd ?t)" by (simp add: comp_def) also have "smap snd ?t = dba.trace B w q" unfolding dbau_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbau_nodes_fst[intro]: assumes "dba.alphabet A = dba.alphabet B" shows "fst ` DBA.nodes (dbau A B) \ DBA.nodes A" proof (rule subsetI, erule imageE) fix pq p assume "pq \ DBA.nodes (dbau A B)" "p = fst pq" then show "p \ DBA.nodes A" using assms unfolding dbau_def by (induct arbitrary: p) (auto) qed lemma dbau_nodes_snd[intro]: assumes "dba.alphabet A = dba.alphabet B" shows "snd ` DBA.nodes (dbau A B) \ DBA.nodes B" proof (rule subsetI, erule imageE) fix pq q assume "pq \ DBA.nodes (dbau A B)" "q = snd pq" then show "q \ DBA.nodes B" using assms unfolding dbau_def by (induct arbitrary: q) (auto) qed lemma dbau_nodes_finite[intro]: assumes "dba.alphabet A = dba.alphabet B" assumes "finite (DBA.nodes A)" "finite (DBA.nodes B)" shows "finite (DBA.nodes (dbau A B))" proof (rule finite_subset) show "DBA.nodes (dbau A B) \ DBA.nodes A \ DBA.nodes B" using dbau_nodes_fst[OF assms(1)] dbau_nodes_snd[OF assms(1)] unfolding image_subset_iff by force show "finite (DBA.nodes A \ DBA.nodes B)" using assms(2, 3) by simp qed lemma dbau_nodes_card[intro]: assumes "dba.alphabet A = dba.alphabet B" assumes "finite (DBA.nodes A)" "finite (DBA.nodes B)" shows "card (DBA.nodes (dbau A B)) \ card (DBA.nodes A) * card (DBA.nodes B)" proof - have "card (DBA.nodes (dbau A B)) \ card (DBA.nodes A \ DBA.nodes B)" proof (rule card_mono) show "finite (DBA.nodes A \ DBA.nodes B)" using assms(2, 3) by simp show "DBA.nodes (dbau A B) \ DBA.nodes A \ DBA.nodes B" using dbau_nodes_fst[OF assms(1)] dbau_nodes_snd[OF assms(1)] unfolding image_subset_iff by force qed also have "\ = card (DBA.nodes A) * card (DBA.nodes B)" using card_cartesian_product by this finally show ?thesis by this qed lemma dbau_language[simp]: assumes "dba.alphabet A = dba.alphabet B" shows "DBA.language (dbau A B) = DBA.language A \ DBA.language B" proof - have 1: "dba.alphabet (dbau A B) = dba.alphabet A \ dba.alphabet B" unfolding dbau_def by simp have 2: "dba.initial (dbau A B) = (dba.initial A, dba.initial B)" unfolding dbau_def by simp have 3: "dba.accepting (dbau A B) = (\ pq. (dba.accepting A \ fst) pq \ (dba.accepting B \ snd) pq)" unfolding dbau_def by auto have 4: "infs (dba.accepting (dbau A B)) (DBA.trace (dbau A B) w (p, q)) \ infs (dba.accepting A) (DBA.trace A w p) \ infs (dba.accepting B) (DBA.trace B w q)" for w p q unfolding 3 by blast show ?thesis using assms unfolding DBA.language_def DBA.run_alt_def 1 2 4 by auto qed definition dbaul :: "('label, 'state) dba list \ ('label, 'state list) dba" where "dbaul AA \ dba (UNION (set AA) dba.alphabet) (map dba.initial AA) - (\ a pp. map2 (\ A p. dba.succ A a p) AA pp) + (\ a pp. map2 (\ A p. dba.transition A a p) AA pp) (\ pp. \ k < length AA. dba.accepting (AA ! k) (pp ! k))" lemma dbaul_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dba.trace (dbaul AA) w pp) = dba.trace (AA ! k) w (pp ! k)" using assms unfolding dbaul_def by (coinduction arbitrary: w pp) (force) lemma dbaul_nodes_length: assumes "pp \ DBA.nodes (dbaul AA)" shows "length pp = length AA" using assms unfolding dbaul_def by induct auto lemma dbaul_nodes[intro]: assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" assumes "pp \ DBA.nodes (dbaul AA)" "k < length pp" shows "pp ! k \ DBA.nodes (AA ! k)" using assms(2, 3, 1) unfolding dbaul_def by induct force+ lemma dbaul_nodes_finite[intro]: assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" assumes "list_all (finite \ DBA.nodes) AA" shows "finite (DBA.nodes (dbaul AA))" proof (rule finite_subset) show "DBA.nodes (dbaul AA) \ listset (map DBA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dbaul_nodes_length) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms(2) by (simp add: list.pred_map) qed lemma dbaul_nodes_card: assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" assumes "list_all (finite \ DBA.nodes) AA" shows "card (DBA.nodes (dbaul AA)) \ prod_list (map (card \ DBA.nodes) AA)" proof - have "card (DBA.nodes (dbaul AA)) \ card (listset (map DBA.nodes AA))" proof (rule card_mono) have "finite (listset (map DBA.nodes AA)) \ list_all finite (map DBA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DBA.nodes AA))" using assms(2) by (simp add: list.pred_map) show "DBA.nodes (dbaul AA) \ listset (map DBA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dbaul_nodes_length) qed also have "\ = prod_list (map (card \ DBA.nodes) AA)" by simp finally show ?thesis by this qed lemma dbaul_language[simp]: assumes "INTER (set AA) dba.alphabet = UNION (set AA) dba.alphabet" shows "DBA.language (dbaul AA) = UNION (set AA) DBA.language" proof safe fix w assume 1: "w \ DBA.language (dbaul AA)" obtain 2: "dba.run (dbaul AA) w (dba.initial (dbaul AA))" "infs (dba.accepting (dbaul AA)) (dba.trace (dbaul AA) w (dba.initial (dbaul AA)))" using 1 by rule obtain k where 3: "k < length AA" "infs (\ pp. dba.accepting (AA ! k) (pp ! k)) (dba.trace (dbaul AA) w (map dba.initial AA))" using 2(2) unfolding dbaul_def by auto show "w \ UNION (set AA) DBA.language" proof (intro UN_I DBA.language) show "AA ! k \ set AA" using 3(1) by simp show "dba.run (AA ! k) w (dba.initial (AA ! k))" using assms 2(1) 3(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbaul_def by force have "True \ infs (\ pp. dba.accepting (AA ! k) (pp ! k)) (dba.trace (dbaul AA) w (map dba.initial AA))" using 3(2) by auto also have "\ \ infs (dba.accepting (AA ! k)) (smap (\ pp. pp ! k) (dba.trace (dbaul AA) w (map dba.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dba.trace (dbaul AA) w (map dba.initial AA)) = dba.trace (AA ! k) w (map dba.initial AA ! k)" using 3(1) by (fastforce intro: dbaul_trace_smap) also have "\ = dba.trace (AA ! k) w (dba.initial (AA ! k))" using 3 by auto finally show "infs (dba.accepting (AA ! k)) (dba.trace (AA ! k) w (dba.initial (AA ! k)))" by simp qed next fix A w assume 1: "A \ set AA" "w \ DBA.language A" obtain 2: "dba.run A w (dba.initial A)" "infs (dba.accepting A) (dba.trace A w (dba.initial A))" using 1(2) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(1) unfolding in_set_conv_nth by auto show "w \ DBA.language (dbaul AA)" proof show "dba.run (dbaul AA) w (dba.initial (dbaul AA))" using 1(1) 2(1) unfolding DBA.run_alt_def DGBA.run_alt_def dbaul_def by auto have "True \ infs (dba.accepting (AA ! k)) (dba.trace (AA ! k) w (map dba.initial AA ! k))" using 2(2) 3 by auto also have "dba.trace (AA ! k) w (map dba.initial AA ! k) = smap (\ pp. pp ! k) (dba.trace (dbaul AA) w (map dba.initial AA))" using 3(2) by (fastforce intro: dbaul_trace_smap[symmetric]) also have "infs (dba.accepting (AA ! k)) \ \ infs (\ pp. dba.accepting (AA ! k) (pp ! k)) (dba.trace (dbaul AA) w (map dba.initial AA))" by (simp add: comp_def) finally show "infs (dba.accepting (dbaul AA)) (dba.trace (dbaul AA) w (dba.initial (dbaul AA)))" using 3(2) unfolding dbaul_def by auto qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DBA/DGBA.thy b/thys/Transition_Systems_and_Automata/Automata/DBA/DGBA.thy --- a/thys/Transition_Systems_and_Automata/Automata/DBA/DGBA.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DBA/DGBA.thy @@ -1,146 +1,146 @@ section \Deterministic Generalized Büchi Automata\ theory DGBA imports "DBA" "../../Basic/Degeneralization" begin datatype ('label, 'state) dgba = dgba (alphabet: "'label set") (initial: "'state") - (succ: "'label \ 'state \ 'state") + (transition: "'label \ 'state \ 'state") (accepting: "'state pred gen") global_interpretation dgba: transition_system_initial - "succ A" "\ a p. a \ alphabet A" "\ p. p = initial A" + "transition A" "\ a p. a \ alphabet A" "\ p. p = initial A" for A defines path = dgba.path and run = dgba.run and reachable = dgba.reachable and nodes = dgba.nodes and enableds = dgba.enableds and paths = dgba.paths and runs = dgba.runs by this abbreviation target where "target \ dgba.target" abbreviation states where "states \ dgba.states" abbreviation trace where "trace \ dgba.trace" abbreviation successors :: "('label, 'state) dgba \ 'state \ 'state set" where "successors \ dgba.successors TYPE('label)" lemma path_alt_def: "path A w p \ set w \ alphabet A" unfolding lists_iff_set[symmetric] proof show "w \ lists (alphabet A)" if "path A w p" using that by (induct arbitrary: p) (auto) show "path A w p" if "w \ lists (alphabet A)" using that by (induct arbitrary: p) (auto) qed lemma run_alt_def: "run A w p \ sset w \ alphabet A" unfolding streams_iff_sset[symmetric] proof show "w \ streams (alphabet A)" if "run A w p" using that by (coinduction arbitrary: w p) (force elim: dgba.run.cases) show "run A w p" if "w \ streams (alphabet A)" using that by (coinduction arbitrary: w p) (force elim: streams.cases) qed definition language :: "('label, 'state) dgba \ 'label stream set" where "language A \ {w. run A w (initial A) \ gen infs (accepting A) (trace A w (initial A))}" lemma language[intro]: assumes "run A w (initial A)" "gen infs (accepting A) (trace A w (initial A))" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains "run A w (initial A)" "gen infs (accepting A) (trace A w (initial A))" using assms unfolding language_def by auto lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def run_alt_def using sset_streams by auto definition dgbad :: "('label, 'state) dgba \ ('label, 'state degen) dba" where "dgbad A \ dba (alphabet A) (initial A, 0) - (\ a (p, k). (succ A a p, count (accepting A) p k)) + (\ a (p, k). (transition A a p, count (accepting A) p k)) (degen (accepting A))" lemma dgbad_simps[simp]: "dba.alphabet (dgbad A) = alphabet A" "dba.initial (dgbad A) = (initial A, 0)" - "dba.succ (dgbad A) a (p, k) = (succ A a p, count (accepting A) p k)" + "dba.transition (dgbad A) a (p, k) = (transition A a p, count (accepting A) p k)" "dba.accepting (dgbad A) = degen (accepting A)" unfolding dgbad_def by auto lemma dgbad_target[simp]: "dba.target (dgbad A) w (p, k) = (dgba.target A w p, fold (count (accepting A)) (butlast (p # dgba.states A w p)) k)" by (induct w arbitrary: p k) (auto) lemma dgbad_states[simp]: "dba.states (dgbad A) w (p, k) = dgba.states A w p || scan (count (accepting A)) (p # dgba.states A w p) k" by (induct w arbitrary: p k) (auto) lemma dgbad_trace[simp]: "dba.trace (dgbad A) w (p, k) = dgba.trace A w p ||| sscan (count (accepting A)) (p ## dgba.trace A w p) k" by (coinduction arbitrary: w p k) (auto) lemma dgbad_path[iff]: "dba.path (dgbad A) w (p, k) \ dgba.path A w p" unfolding DBA.path_alt_def DGBA.path_alt_def by simp lemma dgbad_run[iff]: "dba.run (dgbad A) w (p, k) \ dgba.run A w p" unfolding DBA.run_alt_def DGBA.run_alt_def by simp (* TODO: revise *) lemma dgbad_nodes_fst[simp]: "fst ` DBA.nodes (dgbad A) = DGBA.nodes A" unfolding dba.nodes_alt_def dba.reachable_alt_def unfolding dgba.nodes_alt_def dgba.reachable_alt_def unfolding image_def by simp lemma dgbad_nodes_snd_empty: assumes "accepting A = []" shows "snd ` DBA.nodes (dgbad A) \ {0}" proof - - have 2: "snd (dba.succ (dgbad A) a (p, k)) = 0" for a p k using assms by auto + have 2: "snd (dba.transition (dgbad A) a (p, k)) = 0" for a p k using assms by auto show ?thesis using 2 by (auto elim: dba.nodes.cases) qed lemma dgbad_nodes_snd_nonempty: assumes "accepting A \ []" shows "snd ` DBA.nodes (dgbad A) \ {0 ..< length (accepting A)}" proof - have 1: "snd (dba.initial (dgbad A)) < length (accepting A)" using assms by simp - have 2: "snd (dba.succ (dgbad A) a (p, k)) < length (accepting A)" for a p k + have 2: "snd (dba.transition (dgbad A) a (p, k)) < length (accepting A)" for a p k using assms by auto show ?thesis using 1 2 by (auto elim: dba.nodes.cases) qed lemma dgbad_nodes_empty: assumes "accepting A = []" shows "DBA.nodes (dgbad A) = DGBA.nodes A \ {0}" proof - have "(p, k) \ DBA.nodes (dgbad A) \ p \ fst ` DBA.nodes (dgbad A) \ k = 0" for p k using dgbad_nodes_snd_empty[OF assms] by (force simp del: dgbad_nodes_fst) then show ?thesis by auto qed lemma dgbad_nodes_nonempty: assumes "accepting A \ []" shows "DBA.nodes (dgbad A) \ DGBA.nodes A \ {0 ..< length (accepting A)}" using subset_fst_snd dgbad_nodes_fst[of A] dgbad_nodes_snd_nonempty[OF assms] by blast lemma dgbad_nodes: "DBA.nodes (dgbad A) \ DGBA.nodes A \ {0 ..< max 1 (length (accepting A))}" using dgbad_nodes_empty dgbad_nodes_nonempty by force lemma dgbad_language[simp]: "DBA.language (dgbad A) = DGBA.language A" by force lemma dgbad_nodes_finite[iff]: "finite (DBA.nodes (dgbad A)) \ finite (DGBA.nodes A)" proof show "finite (DGBA.nodes A)" if "finite (DBA.nodes (dgbad A))" using that by (auto simp flip: dgbad_nodes_fst) show "finite (DBA.nodes (dgbad A))" if "finite (DGBA.nodes A)" using dgbad_nodes that finite_subset by fastforce qed lemma dgbad_nodes_card: "card (DBA.nodes (dgbad A)) \ max 1 (length (accepting A)) * card (DGBA.nodes A)" proof (cases "finite (DGBA.nodes A)") case True have "card (DBA.nodes (dgbad A)) \ card (DGBA.nodes A \ {0 ..< max 1 (length (accepting A))})" using dgbad_nodes True by (blast intro: card_mono) also have "\ = max 1 (length (accepting A)) * card (DGBA.nodes A)" unfolding card_cartesian_product by simp finally show ?thesis by this next case False then have "card (DGBA.nodes A) = 0" "card (DBA.nodes (dgbad A)) = 0" by auto then show ?thesis by simp qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DCA/DCA.thy b/thys/Transition_Systems_and_Automata/Automata/DCA/DCA.thy --- a/thys/Transition_Systems_and_Automata/Automata/DCA/DCA.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DCA/DCA.thy @@ -1,62 +1,62 @@ section \Deterministic Co-Büchi Automata\ theory DCA imports "../../Basic/Sequence_Zip" "../../Basic/Acceptance" "../../Transition_Systems/Transition_System" "../../Transition_Systems/Transition_System_Extra" "../../Transition_Systems/Transition_System_Construction" begin datatype ('label, 'state) dca = dca (alphabet: "'label set") (initial: "'state") - (succ: "'label \ 'state \ 'state") + (transition: "'label \ 'state \ 'state") (rejecting: "'state \ bool") global_interpretation dca: transition_system_initial - "succ A" "\ a p. a \ alphabet A" "\ p. p = initial A" + "transition A" "\ a p. a \ alphabet A" "\ p. p = initial A" for A defines path = dca.path and run = dca.run and reachable = dca.reachable and nodes = dca.nodes and enableds = dca.enableds and paths = dca.paths and runs = dca.runs by this abbreviation target where "target \ dca.target" abbreviation states where "states \ dca.states" abbreviation trace where "trace \ dca.trace" abbreviation successors :: "('label, 'state) dca \ 'state \ 'state set" where "successors \ dca.successors TYPE('label)" lemma path_alt_def: "path A w p \ set w \ alphabet A" unfolding lists_iff_set[symmetric] proof show "w \ lists (alphabet A)" if "path A w p" using that by (induct arbitrary: p) (auto) show "path A w p" if "w \ lists (alphabet A)" using that by (induct arbitrary: p) (auto) qed lemma run_alt_def: "run A w p \ sset w \ alphabet A" unfolding streams_iff_sset[symmetric] proof show "w \ streams (alphabet A)" if "run A w p" using that by (coinduction arbitrary: w p) (force elim: dca.run.cases) show "run A w p" if "w \ streams (alphabet A)" using that by (coinduction arbitrary: w p) (force elim: streams.cases) qed definition language :: "('label, 'state) dca \ 'label stream set" where "language A \ {w. run A w (initial A) \ fins (rejecting A) (trace A w (initial A))}" lemma language[intro]: assumes "run A w (initial A)" "fins (rejecting A) (trace A w (initial A))" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains "run A w (initial A)" "fins (rejecting A) (trace A w (initial A))" using assms unfolding language_def by auto lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def run_alt_def using sset_streams by auto end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DCA/DCA_Combine.thy @@ -1,298 +1,298 @@ section \Deterministic Co-Büchi Automata Combinations\ theory DCA_Combine imports "DCA" "DGCA" begin definition dcai :: "('label, 'state\<^sub>1) dca \ ('label, 'state\<^sub>2) dca \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) dca" where "dcai A B \ dca (dca.alphabet A \ dca.alphabet B) (dca.initial A, dca.initial B) - (\ a (p, q). (dca.succ A a p, dca.succ B a q)) + (\ a (p, q). (dca.transition A a p, dca.transition B a q)) (\ (p, q). dca.rejecting A p \ dca.rejecting B q)" lemma dcai_fst[iff]: "infs (P \ fst) (dca.trace (dcai A B) w (p, q)) \ infs P (dca.trace A w p)" proof - let ?t = "dca.trace (dcai A B) w (p, q)" have "infs (P \ fst) ?t \ infs P (smap fst ?t)" by (simp add: comp_def) also have "smap fst ?t = dca.trace A w p" unfolding dcai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dcai_snd[iff]: "infs (P \ snd) (dca.trace (dcai A B) w (p, q)) \ infs P (dca.trace B w q)" proof - let ?t = "dca.trace (dcai A B) w (p, q)" have "infs (P \ snd) ?t \ infs P (smap snd ?t)" by (simp add: comp_def) also have "smap snd ?t = dca.trace B w q" unfolding dcai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dcai_nodes_fst[intro]: "fst ` DCA.nodes (dcai A B) \ DCA.nodes A" proof (rule subsetI, erule imageE) fix pq p assume "pq \ DCA.nodes (dcai A B)" "p = fst pq" then show "p \ DCA.nodes A" unfolding dcai_def by (induct arbitrary: p) (auto) qed lemma dcai_nodes_snd[intro]: "snd ` DCA.nodes (dcai A B) \ DCA.nodes B" proof (rule subsetI, erule imageE) fix pq q assume "pq \ DCA.nodes (dcai A B)" "q = snd pq" then show "q \ DCA.nodes B" unfolding dcai_def by (induct arbitrary: q) (auto) qed lemma dcai_nodes_finite[intro]: assumes "finite (DCA.nodes A)" "finite (DCA.nodes B)" shows "finite (DCA.nodes (dcai A B))" proof (rule finite_subset) show "DCA.nodes (dcai A B) \ DCA.nodes A \ DCA.nodes B" using dcai_nodes_fst dcai_nodes_snd unfolding image_subset_iff by force show "finite (DCA.nodes A \ DCA.nodes B)" using assms by simp qed lemma dcai_nodes_card[intro]: assumes "finite (DCA.nodes A)" "finite (DCA.nodes B)" shows "card (DCA.nodes (dcai A B)) \ card (DCA.nodes A) * card (DCA.nodes B)" proof - have "card (DCA.nodes (dcai A B)) \ card (DCA.nodes A \ DCA.nodes B)" proof (rule card_mono) show "finite (DCA.nodes A \ DCA.nodes B)" using assms by simp show "DCA.nodes (dcai A B) \ DCA.nodes A \ DCA.nodes B" using dcai_nodes_fst dcai_nodes_snd unfolding image_subset_iff by force qed also have "\ = card (DCA.nodes A) * card (DCA.nodes B)" using card_cartesian_product by this finally show ?thesis by this qed lemma dcai_language[simp]: "DCA.language (dcai A B) = DCA.language A \ DCA.language B" proof - have 1: "dca.alphabet (dcai A B) = dca.alphabet A \ dca.alphabet B" unfolding dcai_def by simp have 2: "dca.initial (dcai A B) = (dca.initial A, dca.initial B)" unfolding dcai_def by simp have 3: "dca.rejecting (dcai A B) = (\ pq. (dca.rejecting A \ fst) pq \ (dca.rejecting B \ snd) pq)" unfolding dcai_def by auto have 4: "infs (dca.rejecting (dcai A B)) (DCA.trace (dcai A B) w (p, q)) \ infs (dca.rejecting A) (DCA.trace A w p) \ infs (dca.rejecting B) (DCA.trace B w q)" for w p q unfolding 3 by blast show ?thesis unfolding DCA.language_def DCA.run_alt_def 1 2 4 by auto qed definition dcail :: "('label, 'state) dca list \ ('label, 'state list) dca" where "dcail AA \ dca (INTER (set AA) dca.alphabet) (map dca.initial AA) - (\ a pp. map2 (\ A p. dca.succ A a p) AA pp) + (\ a pp. map2 (\ A p. dca.transition A a p) AA pp) (\ pp. \ k < length AA. dca.rejecting (AA ! k) (pp ! k))" lemma dcail_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dca.trace (dcail AA) w pp) = dca.trace (AA ! k) w (pp ! k)" using assms unfolding dcail_def by (coinduction arbitrary: w pp) (force) lemma dcail_nodes_length: assumes "pp \ DCA.nodes (dcail AA)" shows "length pp = length AA" using assms unfolding dcail_def by induct auto lemma dcail_nodes[intro]: assumes "pp \ DCA.nodes (dcail AA)" "k < length pp" shows "pp ! k \ DCA.nodes (AA ! k)" using assms unfolding dcail_def by induct auto lemma dcail_finite[intro]: assumes "list_all (finite \ DCA.nodes) AA" shows "finite (DCA.nodes (dcail AA))" proof (rule finite_subset) show "DCA.nodes (dcail AA) \ listset (map DCA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dcail_nodes_length) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms by (simp add: list.pred_map) qed lemma dcail_nodes_card: assumes "list_all (finite \ DCA.nodes) AA" shows "card (DCA.nodes (dcail AA)) \ prod_list (map (card \ DCA.nodes) AA)" proof - have "card (DCA.nodes (dcail AA)) \ card (listset (map DCA.nodes AA))" proof (rule card_mono) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms by (simp add: list.pred_map) show "DCA.nodes (dcail AA) \ listset (map DCA.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth dcail_nodes_length) qed also have "\ = prod_list (map (card \ DCA.nodes) AA)" by simp finally show ?thesis by this qed lemma dcail_language[simp]: "DCA.language (dcail AA) = INTER (set AA) DCA.language" proof safe fix A w assume 1: "w \ DCA.language (dcail AA)" "A \ set AA" obtain 2: "dca.run (dcail AA) w (dca.initial (dcail AA))" "fins (dca.rejecting (dcail AA)) (dca.trace (dcail AA) w (dca.initial (dcail AA)))" using 1(1) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(2) unfolding in_set_conv_nth by auto have 4: "fins (\ pp. dca.rejecting A (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" using 2(2) 3 unfolding dcail_def by auto show "w \ DCA.language A" proof show "dca.run A w (dca.initial A)" using 1(2) 2(1) unfolding DCA.run_alt_def dcail_def by auto have "True \ fins (\ pp. dca.rejecting A (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" using 4 by simp also have "\ \ fins (dca.rejecting A) (smap (\ pp. pp ! k) (dca.trace (dcail AA) w (map dca.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dca.trace (dcail AA) w (map dca.initial AA)) = dca.trace (AA ! k) w (map dca.initial AA ! k)" using 3(2) by (fastforce intro: dcail_trace_smap) also have "\ = dca.trace A w (dca.initial A)" using 3 by auto finally show "fins (dca.rejecting A) (DCA.trace A w (dca.initial A))" by simp qed next fix w assume 1: "w \ INTER (set AA) DCA.language" have 2: "dca.run A w (dca.initial A)" "fins (dca.rejecting A) (dca.trace A w (dca.initial A))" if "A \ set AA" for A using 1 that by auto have 3: "fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" if "k < length AA" for k proof - have "True \ fins (dca.rejecting (AA ! k)) (dca.trace (AA ! k) w (map dca.initial AA ! k))" using 2(2) that by auto also have "dca.trace (AA ! k) w (map dca.initial AA ! k) = smap (\ pp. pp ! k) (dca.trace (dcail AA) w (map dca.initial AA))" using that by (fastforce intro: dcail_trace_smap[symmetric]) also have "infs (dca.rejecting (AA ! k)) \ \ infs (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dca.trace (dcail AA) w (map dca.initial AA))" by (simp add: comp_def) finally show ?thesis by simp qed show "w \ DCA.language (dcail AA)" proof show "dca.run (dcail AA) w (dca.initial (dcail AA))" using 2(1) unfolding DCA.run_alt_def dcail_def by auto show "fins (dca.rejecting (dcail AA)) (dca.trace (dcail AA) w (dca.initial (dcail AA)))" using 3 unfolding dcail_def by auto qed qed definition dcgaul :: "('label, 'state) dca list \ ('label, 'state list) dgca" where "dcgaul AA \ dgca (UNION (set AA) dca.alphabet) (map dca.initial AA) - (\ a pp. map2 (\ A p. dca.succ A a p) AA pp) + (\ a pp. map2 (\ A p. dca.transition A a p) AA pp) (map (\ k pp. dca.rejecting (AA ! k) (pp ! k)) [0 ..< length AA])" lemma dcgaul_trace_smap: assumes "length pp = length AA" "k < length AA" shows "smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w pp) = dca.trace (AA ! k) w (pp ! k)" using assms unfolding dcgaul_def by (coinduction arbitrary: w pp) (force) lemma dcgaul_nodes_length: assumes "pp \ DGCA.nodes (dcgaul AA)" shows "length pp = length AA" using assms unfolding dcgaul_def by induct auto lemma dcgaul_nodes[intro]: assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" assumes "pp \ DGCA.nodes (dcgaul AA)" "k < length pp" shows "pp ! k \ DCA.nodes (AA ! k)" using assms(2, 3, 1) unfolding dcgaul_def by induct force+ lemma dcgaul_nodes_finite[intro]: assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" assumes "list_all (finite \ DCA.nodes) AA" shows "finite (DGCA.nodes (dcgaul AA))" proof (rule finite_subset) show "DGCA.nodes (dcgaul AA) \ listset (map DCA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dcgaul_nodes_length) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms(2) by (simp add: list.pred_map) qed lemma dcgaul_nodes_card: assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" assumes "list_all (finite \ DCA.nodes) AA" shows "card (DGCA.nodes (dcgaul AA)) \ prod_list (map (card \ DCA.nodes) AA)" proof - have "card (DGCA.nodes (dcgaul AA)) \ card (listset (map DCA.nodes AA))" proof (rule card_mono) have "finite (listset (map DCA.nodes AA)) \ list_all finite (map DCA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DCA.nodes AA))" using assms(2) by (simp add: list.pred_map) show "DGCA.nodes (dcgaul AA) \ listset (map DCA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth dcgaul_nodes_length) qed also have "\ = prod_list (map (card \ DCA.nodes) AA)" by simp finally show ?thesis by this qed lemma dcgaul_language[simp]: assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" shows "DGCA.language (dcgaul AA) = UNION (set AA) DCA.language" proof safe fix w assume 1: "w \ DGCA.language (dcgaul AA)" obtain k where 2: "dgca.run (dcgaul AA) w (dgca.initial (dcgaul AA))" "k < length AA" "fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (dgca.initial (dcgaul AA)))" using 1 unfolding dcgaul_def by force show "w \ UNION (set AA) DCA.language" proof (intro UN_I DCA.language) show "AA ! k \ set AA" using 2(2) by simp show "dca.run (AA ! k) w (dca.initial (AA ! k))" using assms 2(1, 2) unfolding DCA.run_alt_def DGCA.run_alt_def dcgaul_def by force have "True \ fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (map dca.initial AA))" using 2(3) unfolding dcgaul_def by auto also have "\ \ fins (dca.rejecting (AA ! k)) (smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w (map dca.initial AA)))" by (simp add: comp_def) also have "smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w (map dca.initial AA)) = dca.trace (AA ! k) w (map dca.initial AA ! k)" using 2(2) by (fastforce intro: dcgaul_trace_smap) also have "\ = dca.trace (AA ! k) w (dca.initial (AA ! k))" using 2(2) by auto finally show "fins (dca.rejecting (AA ! k)) (dca.trace (AA ! k) w (dca.initial (AA ! k)))" by simp qed next fix A w assume 1: "A \ set AA" "w \ DCA.language A" obtain 2: "dca.run A w (dca.initial A)" "fins (dca.rejecting A) (dca.trace A w (dca.initial A))" using 1(2) by rule obtain k where 3: "A = AA ! k" "k < length AA" using 1(1) unfolding in_set_conv_nth by auto show "w \ DGCA.language (dcgaul AA)" proof (intro DGCA.language bexI cogen) show "dgca.run (dcgaul AA) w (dgca.initial (dcgaul AA))" using 1(1) 2(1) unfolding DCA.run_alt_def DGCA.run_alt_def dcgaul_def by auto have "True \ fins (dca.rejecting (AA ! k)) (dca.trace (AA ! k) w (map dca.initial AA ! k))" using 2(2) 3 by auto also have "dca.trace (AA ! k) w (map dca.initial AA ! k) = smap (\ pp. pp ! k) (dgca.trace (dcgaul AA) w (map dca.initial AA))" using 3(2) by (fastforce intro: dcgaul_trace_smap[symmetric]) also have "fins (dca.rejecting (AA ! k)) \ \ fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (map dca.initial AA))" by (simp add: comp_def) also have "map dca.initial AA = dgca.initial (dcgaul AA)" unfolding dcgaul_def by simp finally show "fins (\ pp. dca.rejecting (AA ! k) (pp ! k)) (dgca.trace (dcgaul AA) w (dgca.initial (dcgaul AA)))" by simp show "(\ pp. dca.rejecting (AA ! k) (pp ! k)) \ set (dgca.rejecting (dcgaul AA))" unfolding dcgaul_def using 3(2) by simp qed qed definition dcaul :: "('label, 'state) dca list \ ('label, 'state list degen) dca" where "dcaul = dgcad \ dcgaul" lemma dcaul_nodes_finite[intro]: assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" assumes "list_all (finite \ DCA.nodes) AA" shows "finite (DCA.nodes (dcaul AA))" using dcgaul_nodes_finite assms unfolding dcaul_def by auto lemma dcaul_nodes_card: assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" assumes "list_all (finite \ DCA.nodes) AA" shows "card (DCA.nodes (dcaul AA)) \ max 1 (length AA) * prod_list (map (card \ DCA.nodes) AA)" proof - have "card (DCA.nodes (dcaul AA)) \ max 1 (length (dgca.rejecting (dcgaul AA))) * card (DGCA.nodes (dcgaul AA))" unfolding dcaul_def using dgcad_nodes_card by simp also have "length (dgca.rejecting (dcgaul AA)) = length AA" unfolding dcgaul_def by simp also have "card (DGCA.nodes (dcgaul AA)) \ prod_list (map (card \ DCA.nodes) AA)" using dcgaul_nodes_card assms by this finally show ?thesis by simp qed lemma dcaul_language[simp]: assumes "INTER (set AA) dca.alphabet = UNION (set AA) dca.alphabet" shows "DCA.language (dcaul AA) = UNION (set AA) DCA.language" unfolding dcaul_def using dgcad_language dcgaul_language[OF assms] by auto end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DCA/DGCA.thy b/thys/Transition_Systems_and_Automata/Automata/DCA/DGCA.thy --- a/thys/Transition_Systems_and_Automata/Automata/DCA/DGCA.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DCA/DGCA.thy @@ -1,146 +1,146 @@ section \Deterministic Co-Generalized Co-Büchi Automata\ theory DGCA imports "DCA" "../../Basic/Degeneralization" begin datatype ('label, 'state) dgca = dgca (alphabet: "'label set") (initial: "'state") - (succ: "'label \ 'state \ 'state") + (transition: "'label \ 'state \ 'state") (rejecting: "'state pred gen") global_interpretation dgca: transition_system_initial - "succ A" "\ a p. a \ alphabet A" "\ p. p = initial A" + "transition A" "\ a p. a \ alphabet A" "\ p. p = initial A" for A defines path = dgca.path and run = dgca.run and reachable = dgca.reachable and nodes = dgca.nodes and enableds = dgca.enableds and paths = dgca.paths and runs = dgca.runs by this abbreviation target where "target \ dgca.target" abbreviation states where "states \ dgca.states" abbreviation trace where "trace \ dgca.trace" abbreviation successors :: "('label, 'state) dgca \ 'state \ 'state set" where "successors \ dgca.successors TYPE('label)" lemma path_alt_def: "path A w p \ set w \ alphabet A" unfolding lists_iff_set[symmetric] proof show "w \ lists (alphabet A)" if "path A w p" using that by (induct arbitrary: p) (auto) show "path A w p" if "w \ lists (alphabet A)" using that by (induct arbitrary: p) (auto) qed lemma run_alt_def: "run A w p \ sset w \ alphabet A" unfolding streams_iff_sset[symmetric] proof show "w \ streams (alphabet A)" if "run A w p" using that by (coinduction arbitrary: w p) (force elim: dgca.run.cases) show "run A w p" if "w \ streams (alphabet A)" using that by (coinduction arbitrary: w p) (force elim: streams.cases) qed definition language :: "('label, 'state) dgca \ 'label stream set" where "language A \ {w. run A w (initial A) \ cogen fins (rejecting A) (trace A w (initial A))}" lemma language[intro]: assumes "run A w (initial A)" "cogen fins (rejecting A) (trace A w (initial A))" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains "run A w (initial A)" "cogen fins (rejecting A) (trace A w (initial A))" using assms unfolding language_def by auto lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def run_alt_def using sset_streams by auto definition dgcad :: "('label, 'state) dgca \ ('label, 'state degen) dca" where "dgcad A \ dca (alphabet A) (initial A, 0) - (\ a (p, k). (succ A a p, count (rejecting A) p k)) + (\ a (p, k). (transition A a p, count (rejecting A) p k)) (degen (rejecting A))" lemma dgcad_simps[simp]: "dca.alphabet (dgcad A) = alphabet A" "dca.initial (dgcad A) = (initial A, 0)" - "dca.succ (dgcad A) a (p, k) = (succ A a p, count (rejecting A) p k)" + "dca.transition (dgcad A) a (p, k) = (transition A a p, count (rejecting A) p k)" "dca.rejecting (dgcad A) = degen (rejecting A)" unfolding dgcad_def by auto lemma dgcad_target[simp]: "dca.target (dgcad A) w (p, k) = (dgca.target A w p, fold (count (rejecting A)) (butlast (p # dgca.states A w p)) k)" by (induct w arbitrary: p k) (auto) lemma dgcad_states[simp]: "dca.states (dgcad A) w (p, k) = dgca.states A w p || scan (count (rejecting A)) (p # dgca.states A w p) k" by (induct w arbitrary: p k) (auto) lemma dgcad_trace[simp]: "dca.trace (dgcad A) w (p, k) = dgca.trace A w p ||| sscan (count (rejecting A)) (p ## dgca.trace A w p) k" by (coinduction arbitrary: w p k) (auto) lemma dgcad_path[iff]: "dca.path (dgcad A) w (p, k) \ dgca.path A w p" unfolding DCA.path_alt_def DGCA.path_alt_def by simp lemma dgcad_run[iff]: "dca.run (dgcad A) w (p, k) \ dgca.run A w p" unfolding DCA.run_alt_def DGCA.run_alt_def by simp (* TODO: revise *) lemma dgcad_nodes_fst[simp]: "fst ` DCA.nodes (dgcad A) = DGCA.nodes A" unfolding dca.nodes_alt_def dca.reachable_alt_def unfolding dgca.nodes_alt_def dgca.reachable_alt_def unfolding image_def by simp lemma dgcad_nodes_snd_empty: assumes "rejecting A = []" shows "snd ` DCA.nodes (dgcad A) \ {0}" proof - - have 2: "snd (dca.succ (dgcad A) a (p, k)) = 0" for a p k using assms by auto + have 2: "snd (dca.transition (dgcad A) a (p, k)) = 0" for a p k using assms by auto show ?thesis using 2 by (auto elim: dca.nodes.cases) qed lemma dgcad_nodes_snd_nonempty: assumes "rejecting A \ []" shows "snd ` DCA.nodes (dgcad A) \ {0 ..< length (rejecting A)}" proof - have 1: "snd (dca.initial (dgcad A)) < length (rejecting A)" using assms by simp - have 2: "snd (dca.succ (dgcad A) a (p, k)) < length (rejecting A)" for a p k + have 2: "snd (dca.transition (dgcad A) a (p, k)) < length (rejecting A)" for a p k using assms by auto show ?thesis using 1 2 by (auto elim: dca.nodes.cases) qed lemma dgcad_nodes_empty: assumes "rejecting A = []" shows "DCA.nodes (dgcad A) = DGCA.nodes A \ {0}" proof - have "(p, k) \ DCA.nodes (dgcad A) \ p \ fst ` DCA.nodes (dgcad A) \ k = 0" for p k using dgcad_nodes_snd_empty[OF assms] by (force simp del: dgcad_nodes_fst) then show ?thesis by auto qed lemma dgcad_nodes_nonempty: assumes "rejecting A \ []" shows "DCA.nodes (dgcad A) \ DGCA.nodes A \ {0 ..< length (rejecting A)}" using subset_fst_snd dgcad_nodes_fst[of A] dgcad_nodes_snd_nonempty[OF assms] by blast lemma dgcad_nodes: "DCA.nodes (dgcad A) \ DGCA.nodes A \ {0 ..< max 1 (length (rejecting A))}" using dgcad_nodes_empty dgcad_nodes_nonempty by force lemma dgcad_language[simp]: "DCA.language (dgcad A) = DGCA.language A" by force lemma dgcad_nodes_finite[iff]: "finite (DCA.nodes (dgcad A)) \ finite (DGCA.nodes A)" proof show "finite (DGCA.nodes A)" if "finite (DCA.nodes (dgcad A))" using that by (auto simp flip: dgcad_nodes_fst) show "finite (DCA.nodes (dgcad A))" if "finite (DGCA.nodes A)" using dgcad_nodes that finite_subset by fastforce qed lemma dgcad_nodes_card: "card (DCA.nodes (dgcad A)) \ max 1 (length (rejecting A)) * card (DGCA.nodes A)" proof (cases "finite (DGCA.nodes A)") case True have "card (DCA.nodes (dgcad A)) \ card (DGCA.nodes A \ {0 ..< max 1 (length (rejecting A))})" using dgcad_nodes True by (blast intro: card_mono) also have "\ = max 1 (length (rejecting A)) * card (DGCA.nodes A)" unfolding card_cartesian_product by simp finally show ?thesis by this next case False then have "card (DGCA.nodes A) = 0" "card (DCA.nodes (dgcad A)) = 0" by auto then show ?thesis by simp qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA.thy @@ -1,62 +1,62 @@ section \Deterministic Rabin Automata\ theory DRA imports "../../Basic/Sequence_Zip" "../../Basic/Acceptance" "../../Transition_Systems/Transition_System" "../../Transition_Systems/Transition_System_Extra" "../../Transition_Systems/Transition_System_Construction" begin datatype ('label, 'state) dra = dra (alphabet: "'label set") (initial: "'state") - (succ: "'label \ 'state \ 'state") + (transition: "'label \ 'state \ 'state") (accepting: "'state rabin gen") global_interpretation dra: transition_system_initial - "succ A" "\ a p. a \ alphabet A" "\ p. p = initial A" + "transition A" "\ a p. a \ alphabet A" "\ p. p = initial A" for A defines path = dra.path and run = dra.run and reachable = dra.reachable and nodes = dra.nodes and enableds = dra.enableds and paths = dra.paths and runs = dra.runs by this abbreviation target where "target \ dra.target" abbreviation states where "states \ dra.states" abbreviation trace where "trace \ dra.trace" abbreviation successors :: "('label, 'state) dra \ 'state \ 'state set" where "successors \ dra.successors TYPE('label)" lemma path_alt_def: "path A w p \ set w \ alphabet A" unfolding lists_iff_set[symmetric] proof show "w \ lists (alphabet A)" if "path A w p" using that by (induct arbitrary: p) (auto) show "path A w p" if "w \ lists (alphabet A)" using that by (induct arbitrary: p) (auto) qed lemma run_alt_def: "run A w p \ sset w \ alphabet A" unfolding streams_iff_sset[symmetric] proof show "w \ streams (alphabet A)" if "run A w p" using that by (coinduction arbitrary: w p) (force elim: dra.run.cases) show "run A w p" if "w \ streams (alphabet A)" using that by (coinduction arbitrary: w p) (force elim: streams.cases) qed definition language :: "('label, 'state) dra \ 'label stream set" where "language A \ {w. run A w (initial A) \ cogen rabin (accepting A) (trace A w (initial A))}" lemma language[intro]: assumes "run A w (initial A)" "cogen rabin (accepting A) (trace A w (initial A))" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains "run A w (initial A)" "cogen rabin (accepting A) (trace A w (initial A))" using assms unfolding language_def by auto lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def run_alt_def using sset_streams by auto end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Combine.thy @@ -1,223 +1,223 @@ section \Deterministic Rabin Automata Combinations\ theory DRA_Combine imports "DRA" "../DBA/DBA" "../DCA/DCA" begin definition from_dba :: "('label, 'state) dba \ ('label, 'state) dra" where - "from_dba A \ dra (dba.alphabet A) (dba.initial A) (dba.succ A) [(dba.accepting A, bot)]" + "from_dba A \ dra (dba.alphabet A) (dba.initial A) (dba.transition A) [(dba.accepting A, bot)]" lemma from_dba_language[simp]: "DRA.language (from_dba A) = DBA.language A" unfolding DBA.language_def DRA.language_def from_dba_def DBA.run_def DRA.run_def by (auto 0 3) definition from_dca :: "('label, 'state) dca \ ('label, 'state) dra" where - "from_dca A \ dra (dca.alphabet A) (dca.initial A) (dca.succ A) [(top, dca.rejecting A)]" + "from_dca A \ dra (dca.alphabet A) (dca.initial A) (dca.transition A) [(top, dca.rejecting A)]" lemma from_dca_language[simp]: "DRA.language (from_dca A) = DCA.language A" unfolding DCA.language_def DRA.language_def from_dca_def DCA.run_def DRA.run_def by (auto 0 3) definition dbcrai :: "('label, 'state\<^sub>1) dba \ ('label, 'state\<^sub>2) dca \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) dra" where "dbcrai A B \ dra (dba.alphabet A \ dca.alphabet B) (dba.initial A, dca.initial B) - (\ a (p, q). (dba.succ A a p, dca.succ B a q)) + (\ a (p, q). (dba.transition A a p, dca.transition B a q)) [(dba.accepting A \ fst, dca.rejecting B \ snd)]" lemma dbcrai_fst[iff]: "infs (P \ fst) (dra.trace (dbcrai A B) w (p, q)) \ infs P (dba.trace A w p)" proof - let ?t = "dra.trace (dbcrai A B) w (p, q)" have "infs (P \ fst) ?t \ infs P (smap fst ?t)" by (simp add: comp_def) also have "smap fst ?t = dba.trace A w p" unfolding dbcrai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbcrai_snd[iff]: "infs (P \ snd) (dra.trace (dbcrai A B) w (p, q)) \ infs P (dca.trace B w q)" proof - let ?t = "dra.trace (dbcrai A B) w (p, q)" have "infs (P \ snd) ?t \ infs P (smap snd ?t)" by (simp add: comp_def) also have "smap snd ?t = dca.trace B w q" unfolding dbcrai_def by (coinduction arbitrary: w p q) (auto) finally show ?thesis by this qed lemma dbcrai_nodes_fst[intro]: "fst ` DRA.nodes (dbcrai A B) \ DBA.nodes A" proof (rule subsetI, erule imageE) fix pq p assume "pq \ DRA.nodes (dbcrai A B)" "p = fst pq" then show "p \ DBA.nodes A" unfolding dbcrai_def by (induct arbitrary: p) (auto) qed lemma dbcrai_nodes_snd[intro]: "snd ` DRA.nodes (dbcrai A B) \ DCA.nodes B" proof (rule subsetI, erule imageE) fix pq q assume "pq \ DRA.nodes (dbcrai A B)" "q = snd pq" then show "q \ DCA.nodes B" unfolding dbcrai_def by (induct arbitrary: q) (auto) qed lemma dbcrai_nodes_finite[intro]: assumes "finite (DBA.nodes A)" "finite (DCA.nodes B)" shows "finite (DRA.nodes (dbcrai A B))" proof (rule finite_subset) show "DRA.nodes (dbcrai A B) \ DBA.nodes A \ DCA.nodes B" using dbcrai_nodes_fst dbcrai_nodes_snd unfolding image_subset_iff by force show "finite (DBA.nodes A \ DCA.nodes B)" using assms by simp qed lemma dbcrai_nodes_card[intro]: assumes "finite (DBA.nodes A)" "finite (DCA.nodes B)" shows "card (DRA.nodes (dbcrai A B)) \ card (DBA.nodes A) * card (DCA.nodes B)" proof - have "card (DRA.nodes (dbcrai A B)) \ card (DBA.nodes A \ DCA.nodes B)" proof (rule card_mono) show "finite (DBA.nodes A \ DCA.nodes B)" using assms by simp show "DRA.nodes (dbcrai A B) \ DBA.nodes A \ DCA.nodes B" using dbcrai_nodes_fst dbcrai_nodes_snd unfolding image_subset_iff by force qed also have "\ = card (DBA.nodes A) * card (DCA.nodes B)" using card_cartesian_product by this finally show ?thesis by this qed lemma dbcrai_language[simp]: "DRA.language (dbcrai A B) = DBA.language A \ DCA.language B" proof - have 1: "dra.alphabet (dbcrai A B) = dba.alphabet A \ dca.alphabet B" unfolding dbcrai_def by simp have 2: "dra.initial (dbcrai A B) = (dba.initial A, dca.initial B)" unfolding dbcrai_def by simp have 3: "dra.accepting (dbcrai A B) = [(dba.accepting A \ fst, dca.rejecting B \ snd)]" unfolding dbcrai_def by simp have 4: "cogen rabin (dra.accepting (dbcrai A B)) (DRA.trace (dbcrai A B) w (p, q)) \ infs (dba.accepting A) (DBA.trace A w p) \ fins (rejecting B) (DCA.trace B w q)" for w p q unfolding cogen_def rabin_def 3 by simp show ?thesis unfolding DRA.language_def DBA.language_def DCA.language_def unfolding DRA.run_alt_def DBA.run_alt_def DCA.run_alt_def unfolding 1 2 4 by auto qed abbreviation (input) "get k pp \ pp ! k" definition draul :: "('label, 'state) dra list \ ('label, 'state list) dra" where "draul AA \ dra (UNION (set AA) dra.alphabet) (map dra.initial AA) - (\ a pp. map2 (\ A p. dra.succ A a p) AA pp) + (\ a pp. map2 (\ A p. dra.transition A a p) AA pp) (do { k \ [0 ..< length AA]; (f, g) \ dra.accepting (AA ! k); [(f \ get k, g \ get k)] })" lemma draul_get: assumes "length pp = length AA" "k < length AA" shows "infs (P \ get k) (dra.trace (draul AA) w pp) \ infs P (dra.trace (AA ! k) w (pp ! k))" proof - have "infs (P \ get k) (dra.trace (draul AA) w pp) \ infs P (smap (get k) (dra.trace (draul AA) w pp))" by (simp add: comp_def) also have "smap (get k) (dra.trace (draul AA) w pp) = dra.trace (AA ! k) w (pp ! k)" using assms unfolding draul_def by (coinduction arbitrary: w pp) (force) finally show ?thesis by this qed lemma draul_nodes_length: assumes "pp \ DRA.nodes (draul AA)" shows "length pp = length AA" using assms unfolding draul_def by induct auto lemma draul_nodes[intro]: assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" assumes "pp \ DRA.nodes (draul AA)" "k < length pp" shows "pp ! k \ DRA.nodes (AA ! k)" using assms(2, 3, 1) unfolding draul_def by induct force+ lemma draul_nodes_finite[intro]: assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" assumes "list_all (finite \ DRA.nodes) AA" shows "finite (DRA.nodes (draul AA))" proof (rule finite_subset) show "DRA.nodes (draul AA) \ listset (map DRA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth draul_nodes_length) have "finite (listset (map DRA.nodes AA)) \ list_all finite (map DRA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DRA.nodes AA))" using assms(2) by (simp add: list.pred_map) qed lemma draul_nodes_card: assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" assumes "list_all (finite \ DRA.nodes) AA" shows "card (DRA.nodes (draul AA)) \ prod_list (map (card \ DRA.nodes) AA)" proof - have "card (DRA.nodes (draul AA)) \ card (listset (map DRA.nodes AA))" proof (rule card_mono) have "finite (listset (map DRA.nodes AA)) \ list_all finite (map DRA.nodes AA)" by (rule listset_finite) (auto simp: list_all_iff) then show "finite (listset (map DRA.nodes AA))" using assms(2) by (simp add: list.pred_map) show "DRA.nodes (draul AA) \ listset (map DRA.nodes AA)" using assms(1) by (force simp: listset_member list_all2_conv_all_nth draul_nodes_length) qed also have "\ = prod_list (map (card \ DRA.nodes) AA)" by simp finally show ?thesis by this qed lemma draul_language[simp]: assumes "INTER (set AA) dra.alphabet = UNION (set AA) dra.alphabet" shows "DRA.language (draul AA) = UNION (set AA) DRA.language" proof safe fix w assume 1: "w \ DRA.language (draul AA)" obtain 2: "dra.run (draul AA) w (dra.initial (draul AA))" "cogen rabin (dra.accepting (draul AA)) (dra.trace (draul AA) w (dra.initial (draul AA)))" using 1 by rule obtain I F where 3: "(I, F) \ set (dra.accepting (draul AA))" "infs I (dra.trace (draul AA) w (dra.initial (draul AA)))" "fins F (dra.trace (draul AA) w (dra.initial (draul AA)))" using 2(2) by blast obtain k P Q where 4: "k < length AA" "I = P \ get k" "F = Q \ get k" "(P, Q) \ set (dra.accepting (AA ! k))" using 3(1) unfolding draul_def List.bind_def by (auto simp: comp_def) show "w \ UNION (set AA) DRA.language" proof (intro UN_I DRA.language cogen rabin) show "AA ! k \ set AA" using 4(1) by auto show "DRA.run (AA ! k) w (dra.initial (AA ! k))" using assms 2(1) 4(1) unfolding DRA.run_alt_def draul_def by force show "(P, Q) \ set (dra.accepting (AA ! k))" using 4(4) by this show "(P, Q) = (P, Q)" by rule have "True \ infs (P \ get k) (dra.trace (draul AA) w (map dra.initial AA))" using 3(2) unfolding draul_def 4(2) by simp also have "\ \ infs P (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(1) by (force intro!: draul_get) also have "map dra.initial AA ! k = dra.initial (AA ! k)" using 4(1) by simp finally show "infs P (dra.trace (AA ! k) w (dra.initial (AA ! k)))" by simp have "False \ infs (Q \ get k) (dra.trace (draul AA) w (map dra.initial AA))" using 3(3) unfolding draul_def 4(3) by simp also have "\ \ infs Q (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(1) by (force intro!: draul_get) also have "map dra.initial AA ! k = dra.initial (AA ! k)" using 4(1) by simp finally show "fins Q (dra.trace (AA ! k) w (dra.initial (AA ! k)))" by simp qed next fix A w assume 1: "A \ set AA" "w \ DRA.language A" obtain 2: "dra.run A w (dra.initial A)" "cogen rabin (dra.accepting A) (dra.trace A w (dra.initial A))" using 1(2) by rule obtain I F where 3: "(I, F) \ set (dra.accepting A)" "infs I (dra.trace A w (dra.initial A))" "fins F (dra.trace A w (dra.initial A))" using 2(2) by blast obtain k where 4: "A = AA ! k" "k < length AA" using 1(1) unfolding in_set_conv_nth by auto show "w \ DRA.language (draul AA)" proof (intro DRA.language cogen rabin) show "dra.run (draul AA) w (dra.initial (draul AA))" using 1(1) 2(1) unfolding DRA.run_alt_def draul_def by auto show "(I \ get k, F \ get k) \ set (dra.accepting (draul AA))" unfolding draul_def List.bind_def using 3(1) 4 by (force simp: comp_def) show "(I \ get k, F \ get k) = (I \ get k, F \ get k)" by rule have "infs (I \ get k) (dra.trace (draul AA) w (dra.initial (draul AA))) \ infs (I \ get k) (dra.trace (draul AA) w (map dra.initial AA))" unfolding draul_def by simp also have "\ \ infs I (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(2) by (force intro!: draul_get) also have "\ \ True" using 3(2) 4 by simp finally show "infs (I \ get k) (dra.trace (draul AA) w (dra.initial (draul AA)))" by simp have "infs (F \ get k) (dra.trace (draul AA) w (dra.initial (draul AA))) \ infs (F \ get k) (dra.trace (draul AA) w (map dra.initial AA))" unfolding draul_def by simp also have "\ \ infs F (dra.trace (AA ! k) w (map dra.initial AA ! k))" using 4(2) by (force intro!: draul_get) also have "\ \ False" using 3(3) 4 by simp finally show "fins (F \ get k) (dra.trace (draul AA) w (dra.initial (draul AA)))" by simp qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Explicit.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Explicit.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Explicit.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Explicit.thy @@ -1,208 +1,208 @@ section \Explicit Deterministic Rabin Automata\ theory DRA_Explicit imports DRA_Nodes begin datatype ('label, 'state) drae = drae (alphabete: "'label set") (initiale: "'state") - (transe: "('state \ 'label \ 'state) set") + (transitione: "('state \ 'label \ 'state) set") (acceptinge: "('state set \ 'state set) list") definition drae_rel where [to_relAPP]: "drae_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabete A\<^sub>1, alphabete A\<^sub>2) \ \L\ set_rel \ (initiale A\<^sub>1, initiale A\<^sub>2) \ S \ - (transe A\<^sub>1, transe A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ + (transitione A\<^sub>1, transitione A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ (acceptinge A\<^sub>1, acceptinge A\<^sub>2) \ \\S\ set_rel \\<^sub>r \S\ set_rel\ list_rel}" lemma drae_param[param, autoref_rules]: "(drae, drae) \ \L\ set_rel \ S \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ \\S\ set_rel \\<^sub>r \S\ set_rel\ list_rel \ \L, S\ drae_rel" "(alphabete, alphabete) \ \L, S\ drae_rel \ \L\ set_rel" "(initiale, initiale) \ \L, S\ drae_rel \ S" - "(transe, transe) \ \L, S\ drae_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel" + "(transitione, transitione) \ \L, S\ drae_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel" "(acceptinge, acceptinge) \ \L, S\ drae_rel \ \\S\ set_rel \\<^sub>r \S\ set_rel\ list_rel" unfolding drae_rel_def by auto lemma drae_rel_id[simp]: "\Id, Id\ drae_rel = Id" unfolding drae_rel_def using drae.expand by auto lemma drae_rel_comp[simp]: "\L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ drae_rel = \L\<^sub>1, S\<^sub>1\ drae_rel O \L\<^sub>2, S\<^sub>2\ drae_rel" proof safe fix A B assume 1: "(A, B) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ drae_rel" obtain a b c d where 2: "(alphabete A, a) \ \L\<^sub>1\ set_rel" "(a, alphabete B) \ \L\<^sub>2\ set_rel" "(initiale A, b) \ S\<^sub>1" "(b, initiale B) \ S\<^sub>2" - "(transe A, c) \ \S\<^sub>1 \\<^sub>r L\<^sub>1 \\<^sub>r S\<^sub>1\ set_rel" "(c, transe B) \ \S\<^sub>2 \\<^sub>r L\<^sub>2 \\<^sub>r S\<^sub>2\ set_rel" + "(transitione A, c) \ \S\<^sub>1 \\<^sub>r L\<^sub>1 \\<^sub>r S\<^sub>1\ set_rel" "(c, transitione B) \ \S\<^sub>2 \\<^sub>r L\<^sub>2 \\<^sub>r S\<^sub>2\ set_rel" "(acceptinge A, d) \ \\S\<^sub>1\ set_rel \\<^sub>r \S\<^sub>1\ set_rel\ list_rel" "(d, acceptinge B) \ \\S\<^sub>2\ set_rel \\<^sub>r \S\<^sub>2\ set_rel\ list_rel" using 1 unfolding drae_rel_def prod_rel_compp set_rel_compp by auto show "(A, B) \ \L\<^sub>1, S\<^sub>1\ drae_rel O \L\<^sub>2, S\<^sub>2\ drae_rel" proof show "(A, drae a b c d) \ \L\<^sub>1, S\<^sub>1\ drae_rel" using 2 unfolding drae_rel_def by auto show "(drae a b c d, B) \ \L\<^sub>2, S\<^sub>2\ drae_rel" using 2 unfolding drae_rel_def by auto qed next show "(A, C) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ drae_rel" if "(A, B) \ \L\<^sub>1, S\<^sub>1\ drae_rel" "(B, C) \ \L\<^sub>2, S\<^sub>2\ drae_rel" for A B C using that unfolding drae_rel_def prod_rel_compp set_rel_compp by auto qed (* TODO: why do we need all this setup? can't i_of_rel do the trick? *) consts i_drae_scheme :: "interface \ interface \ interface" context begin interpretation autoref_syn by this lemma drae_scheme_itype[autoref_itype]: "drae ::\<^sub>i \L\\<^sub>i i_set \\<^sub>i S \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set \\<^sub>i \\\S\\<^sub>i i_set, \S\\<^sub>i i_set\\<^sub>i i_prod\\<^sub>i i_list \\<^sub>i \L, S\\<^sub>i i_drae_scheme" "alphabete ::\<^sub>i \L, S\\<^sub>i i_drae_scheme \\<^sub>i \L\\<^sub>i i_set" "initiale ::\<^sub>i \L, S\\<^sub>i i_drae_scheme \\<^sub>i S" - "transe ::\<^sub>i \L, S\\<^sub>i i_drae_scheme \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set" + "transitione ::\<^sub>i \L, S\\<^sub>i i_drae_scheme \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set" "acceptinge ::\<^sub>i \L, S\\<^sub>i i_drae_scheme \\<^sub>i \\\S\\<^sub>i i_set, \S\\<^sub>i i_set\\<^sub>i i_prod\\<^sub>i i_list" by auto end datatype ('label, 'state) draei = draei (alphabetei: "'label list") (initialei: "'state") - (transei: "('state \ 'label \ 'state) list") + (transitionei: "('state \ 'label \ 'state) list") (acceptingei: "('state list \ 'state list) list") definition draei_rel where [to_relAPP]: "draei_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabetei A\<^sub>1, alphabetei A\<^sub>2) \ \L\ list_rel \ (initialei A\<^sub>1, initialei A\<^sub>2) \ S \ - (transei A\<^sub>1, transei A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ + (transitionei A\<^sub>1, transitionei A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ (acceptingei A\<^sub>1, acceptingei A\<^sub>2) \ \\S\ list_rel \\<^sub>r \S\ list_rel\ list_rel}" lemma draei_param[param, autoref_rules]: "(draei, draei) \ \L\ list_rel \ S \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ \\S\ list_rel \\<^sub>r \S\ list_rel\ list_rel \ \L, S\ draei_rel" "(alphabetei, alphabetei) \ \L, S\ draei_rel \ \L\ list_rel" "(initialei, initialei) \ \L, S\ draei_rel \ S" - "(transei, transei) \ \L, S\ draei_rel \ \S \\<^sub>r L \\<^sub>r S\ list_rel" + "(transitionei, transitionei) \ \L, S\ draei_rel \ \S \\<^sub>r L \\<^sub>r S\ list_rel" "(acceptingei, acceptingei) \ \L, S\ draei_rel \ \\S\ list_rel \\<^sub>r \S\ list_rel\ list_rel" unfolding draei_rel_def by auto definition draei_drae_rel where [to_relAPP]: "draei_drae_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabetei A\<^sub>1, alphabete A\<^sub>2) \ \L\ list_set_rel \ (initialei A\<^sub>1, initiale A\<^sub>2) \ S \ - (transei A\<^sub>1, transe A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ + (transitionei A\<^sub>1, transitione A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ (acceptingei A\<^sub>1, acceptinge A\<^sub>2) \ \\S\ list_set_rel \\<^sub>r \S\ list_set_rel\ list_rel}" lemmas [autoref_rel_intf] = REL_INTFI[of draei_drae_rel i_drae_scheme] lemma draei_drae_param[param, autoref_rules]: "(draei, drae) \ \L\ list_set_rel \ S \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ \\S\ list_set_rel \\<^sub>r \S\ list_set_rel\ list_rel \ \L, S\ draei_drae_rel" "(alphabetei, alphabete) \ \L, S\ draei_drae_rel \ \L\ list_set_rel" "(initialei, initiale) \ \L, S\ draei_drae_rel \ S" - "(transei, transe) \ \L, S\ draei_drae_rel \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel" + "(transitionei, transitione) \ \L, S\ draei_drae_rel \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel" "(acceptingei, acceptinge) \ \L, S\ draei_drae_rel \ \\S\ list_set_rel \\<^sub>r \S\ list_set_rel\ list_rel" unfolding draei_drae_rel_def by auto definition draei_drae where "draei_drae A \ drae (set (alphabetei A)) (initialei A) - (set (transei A)) (map (map_prod set set) (acceptingei A))" + (set (transitionei A)) (map (map_prod set set) (acceptingei A))" lemma draei_drae_id_param[param]: "(draei_drae, id) \ \L, S\ draei_drae_rel \ \L, S\ drae_rel" proof fix Ai A assume 1: "(Ai, A) \ \L, S\ draei_drae_rel" have 2: "draei_drae Ai = drae (set (alphabetei Ai)) (initialei Ai) - (set (transei Ai)) (map (map_prod set set) (acceptingei Ai))" unfolding draei_drae_def by rule + (set (transitionei Ai)) (map (map_prod set set) (acceptingei Ai))" unfolding draei_drae_def by rule have 3: "id A = drae (id (alphabete A)) (initiale A) - (id (transe A)) (map (map_prod id id) (acceptinge A))" by simp + (id (transitione A)) (map (map_prod id id) (acceptinge A))" by simp show "(draei_drae Ai, id A) \ \L, S\ drae_rel" unfolding 2 3 using 1 by parametricity qed abbreviation "transitions L S s \ \ a \ L. \ p \ S. {p} \ {a} \ {s a p}" abbreviation "succs T a p \ the_elem ((T `` {p}) `` {a})" definition wft :: "'label set \ 'state set \ ('state \ 'label \ 'state) set \ bool" where "wft L S T \ \ a \ L. \ p \ S. is_singleton ((T `` {p}) `` {a})" lemma wft_param[param]: assumes "bijective S" "bijective L" shows "(wft, wft) \ \L\ set_rel \ \S\ set_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ bool_rel" using assms unfolding wft_def by parametricity lemma wft_transitions: "wft L S (transitions L S s)" unfolding wft_def is_singleton_def by auto definition dra_drae where "dra_drae A \ drae (alphabet A) (initial A) - (transitions (alphabet A) (nodes A) (succ A)) + (transitions (alphabet A) (nodes A) (transition A)) (map (\ (P, Q). (Set.filter P (nodes A), Set.filter Q (nodes A))) (accepting A))" definition drae_dra where "drae_dra A \ dra (alphabete A) (initiale A) - (succs (transe A)) (map (\ (I, F). (\ p. p \ I, \ p. p \ F)) (acceptinge A))" + (succs (transitione A)) (map (\ (I, F). (\ p. p \ I, \ p. p \ F)) (acceptinge A))" lemma set_rel_Domain_Range[intro!, simp]: "(Domain A, Range A) \ \A\ set_rel" unfolding set_rel_def by auto lemma dra_drae_param[param]: "(dra_drae, dra_drae) \ \L, S\ dra_rel \ \L, S\ drae_rel" unfolding dra_drae_def by parametricity lemma drae_dra_param[param]: assumes "bijective L" "bijective S" - assumes "wft (Range L) (Range S) (transe B)" + assumes "wft (Range L) (Range S) (transitione B)" assumes [param]: "(A, B) \ \L, S\ drae_rel" shows "(drae_dra A, drae_dra B) \ \L, S\ dra_rel" proof - - have 1: "(wft (Domain L) (Domain S) (transe A), wft (Range L) (Range S) (transe B)) \ bool_rel" + have 1: "(wft (Domain L) (Domain S) (transitione A), wft (Range L) (Range S) (transitione B)) \ bool_rel" using assms(1, 2) by parametricity auto - have 2: "wft (Domain L) (Domain S) (transe A)" using assms(3) 1 by simp + have 2: "wft (Domain L) (Domain S) (transitione A)" using assms(3) 1 by simp show ?thesis using assms(1 - 3) 2 assms(2)[unfolded bijective_alt] unfolding drae_dra_def wft_def by parametricity force+ qed lemma succs_transitions_param[param]: "(succs \ transitions L S, id) \ (Id_on L \ Id_on S \ Id_on S) \ (Id_on L \ Id_on S \ Id_on S)" proof fix f g assume 1[param]: "(f, g) \ Id_on L \ Id_on S \ Id_on S" show "((succs \ transitions L S) f, id g) \ Id_on L \ Id_on S \ Id_on S" proof safe fix a p assume 2: "a \ L" "p \ S" have "(succs \ transitions L S) f a p = succs (transitions L S f) a p" by simp also have "(transitions L S f `` {p}) `` {a} = {f a p}" using 2 by auto also have "the_elem \ = f a p" by simp also have "(\, g a p) \ Id_on S" using 2 by parametricity auto finally show "(succs \ transitions L S) f a p = id g a p" by simp show "id g a p \ S" using 1[param_fo] 2 by simp qed qed lemma drae_dra_dra_drae_param[param]: "((drae_dra \ dra_drae) A, id A) \ \Id_on (alphabet A), Id_on (nodes A)\ dra_rel" proof - have [param]: "(\ (P, Q). (\ p. p \ Set.filter P (nodes A), \ p. p \ Set.filter Q (nodes A)), id) \ pred_rel (Id_on (nodes A)) \\<^sub>r pred_rel (Id_on (nodes A)) \ rabin_rel (Id_on (nodes A))" unfolding fun_rel_def Id_on_def by auto have "(drae_dra \ dra_drae) A = dra (alphabet A) (initial A) - ((succs \ transitions (alphabet A) (nodes A)) (succ A)) + ((succs \ transitions (alphabet A) (nodes A)) (transition A)) (map (\ (P, Q). (\ p. p \ Set.filter P (nodes A), \ p. p \ Set.filter Q (nodes A))) (accepting A))" unfolding drae_dra_def dra_drae_def by auto - also have "(\, dra (alphabet A) (initial A) (id (succ A)) (map id (accepting A))) \ + also have "(\, dra (alphabet A) (initial A) (id (transition A)) (map id (accepting A))) \ \Id_on (alphabet A), Id_on (nodes A)\ dra_rel" using dra_rel_eq by parametricity auto - also have "dra (alphabet A) (initial A) (id (succ A)) (map id (accepting A)) = id A" by simp + also have "dra (alphabet A) (initial A) (id (transition A)) (map id (accepting A)) = id A" by simp finally show ?thesis by this qed definition draei_dra_rel where [to_relAPP]: "draei_dra_rel L S \ {(Ae, A). (drae_dra (draei_drae Ae), A) \ \L, S\ dra_rel}" lemma draei_dra_id[param]: "(drae_dra \ draei_drae, id) \ \L, S\ draei_dra_rel \ \L, S\ dra_rel" unfolding draei_dra_rel_def by auto (* schematic_goal drae_dra_impl: "(?f, drae_dra) \ \Id, Id\ draei_drae_rel \ \Id, Id\ drai_dra_rel" unfolding drae_dra_def by (autoref (trace)) concrete_definition drae_dra_impl uses drae_dra_impl *) end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Implement.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Implement.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Implement.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Implement.thy @@ -1,81 +1,81 @@ section \Implementation of Deterministic Rabin Automata\ theory DRA_Implement imports "DRA_Refine" "../../Basic/Implement" begin datatype ('label, 'state) drai = drai (alphabeti: "'label list") (initiali: "'state") - (succi: "'label \ 'state \ 'state") + (transitioni: "'label \ 'state \ 'state") (acceptingi: "'state rabin gen") definition drai_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) drai \ ('label\<^sub>2, 'state\<^sub>2) drai) set" where [to_relAPP]: "drai_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabeti A\<^sub>1, alphabeti A\<^sub>2) \ \L\ list_rel \ (initiali A\<^sub>1, initiali A\<^sub>2) \ S \ - (succi A\<^sub>1, succi A\<^sub>2) \ L \ S \ S \ + (transitioni A\<^sub>1, transitioni A\<^sub>2) \ L \ S \ S \ (acceptingi A\<^sub>1, acceptingi A\<^sub>2) \ \rabin_rel S\ list_rel}" lemma drai_param[param]: "(drai, drai) \ \L\ list_rel \ S \ (L \ S \ S) \ \rabin_rel S\ list_rel \ \L, S\ drai_rel" "(alphabeti, alphabeti) \ \L, S\ drai_rel \ \L\ list_rel" "(initiali, initiali) \ \L, S\ drai_rel \ S" - "(succi, succi) \ \L, S\ drai_rel \ L \ S \ S" + "(transitioni, transitioni) \ \L, S\ drai_rel \ L \ S \ S" "(acceptingi, acceptingi) \ \L, S\ drai_rel \ \rabin_rel S\ list_rel" unfolding drai_rel_def fun_rel_def by auto definition drai_dra_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) drai \ ('label\<^sub>2, 'state\<^sub>2) dra) set" where [to_relAPP]: "drai_dra_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabeti A\<^sub>1, alphabet A\<^sub>2) \ \L\ list_set_rel \ (initiali A\<^sub>1, initial A\<^sub>2) \ S \ - (succi A\<^sub>1, succ A\<^sub>2) \ L \ S \ S \ + (transitioni A\<^sub>1, transition A\<^sub>2) \ L \ S \ S \ (acceptingi A\<^sub>1, accepting A\<^sub>2) \ \rabin_rel S\ list_rel}" lemma drai_dra_param[param, autoref_rules]: "(drai, dra) \ \L\ list_set_rel \ S \ (L \ S \ S) \ \rabin_rel S\ list_rel \ \L, S\ drai_dra_rel" "(alphabeti, alphabet) \ \L, S\ drai_dra_rel \ \L\ list_set_rel" "(initiali, initial) \ \L, S\ drai_dra_rel \ S" - "(succi, succ) \ \L, S\ drai_dra_rel \ L \ S \ S" + "(transitioni, transition) \ \L, S\ drai_dra_rel \ L \ S \ S" "(acceptingi, accepting) \ \L, S\ drai_dra_rel \ \rabin_rel S\ list_rel" unfolding drai_dra_rel_def fun_rel_def by auto definition drai_dra :: "('label, 'state) drai \ ('label, 'state) dra" where - "drai_dra A \ dra (set (alphabeti A)) (initiali A) (succi A) (acceptingi A)" + "drai_dra A \ dra (set (alphabeti A)) (initiali A) (transitioni A) (acceptingi A)" definition drai_invar :: "('label, 'state) drai \ bool" where "drai_invar A \ distinct (alphabeti A)" lemma drai_dra_id_param[param]: "(drai_dra, id) \ \L, S\ drai_dra_rel \ \L, S\ dra_rel" proof fix Ai A assume 1: "(Ai, A) \ \L, S\ drai_dra_rel" - have 2: "drai_dra Ai = dra (set (alphabeti Ai)) (initiali Ai) (succi Ai) (acceptingi Ai)" + have 2: "drai_dra Ai = dra (set (alphabeti Ai)) (initiali Ai) (transitioni Ai) (acceptingi Ai)" unfolding drai_dra_def by rule - have 3: "id A = dra (id (alphabet A)) (initial A) (succ A) (accepting A)" by simp + have 3: "id A = dra (id (alphabet A)) (initial A) (transition A) (accepting A)" by simp show "(drai_dra Ai, id A) \ \L, S\ dra_rel" unfolding 2 3 using 1 by parametricity qed lemma drai_dra_br: "\Id, Id\ drai_dra_rel = br drai_dra drai_invar" proof safe show "(A, B) \ \Id, Id\ drai_dra_rel" if "(A, B) \ br drai_dra drai_invar" for A and B :: "('a, 'b) dra" using that unfolding drai_dra_rel_def drai_dra_def drai_invar_def by (auto simp: in_br_conv list_set_rel_def) show "(A, B) \ br drai_dra drai_invar" if "(A, B) \ \Id, Id\ drai_dra_rel" for A and B :: "('a, 'b) dra" proof - have 1: "(drai_dra A, id B) \ \Id, Id\ dra_rel" using that by parametricity have 2: "drai_invar A" using drai_dra_param(2 - 5)[param_fo, OF that] by (auto simp: in_br_conv list_set_rel_def drai_invar_def) show ?thesis using 1 2 unfolding in_br_conv by auto qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Nodes.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Nodes.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Nodes.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Nodes.thy @@ -1,77 +1,77 @@ section \Exploration of Deterministic Rabin Automata\ theory DRA_Nodes imports DFS_Framework.Reachable_Nodes DRA_Implement begin definition dra_G :: "('label, 'state) dra \ 'state graph_rec" where "dra_G A \ \ g_V = UNIV, g_E = E_of_succ (successors A), g_V0 = {initial A} \" lemma dra_G_graph[simp]: "graph (dra_G A)" unfolding dra_G_def graph_def by simp lemma dra_G_reachable_nodes: "op_reachable (dra_G A) = nodes A" unfolding op_reachable_def dra_G_def graph_rec.simps E_of_succ_def proof safe show "p \ nodes A" if "(initial A, p) \ {(u, v). v \ successors A u}\<^sup>*" for p using that by induct auto show "(initial A, p) \ {(u, v). v \ successors A u}\<^sup>*" if "p \ nodes A" for p using that by (induct) (auto intro: rtrancl_into_rtrancl) qed context begin interpretation autoref_syn by this lemma dra_G_ahs: "dra_G A = \ g_V = UNIV, g_E = E_of_succ (\ p. CAST - ((\ a. succ A a p ::: S) ` alphabet A ::: \S\ ahs_rel bhc)), g_V0 = {initial A} \" + ((\ a. transition A a p ::: S) ` alphabet A ::: \S\ ahs_rel bhc)), g_V0 = {initial A} \" unfolding dra_G_def CAST_def id_apply E_of_succ_def autoref_tag_defs by auto schematic_goal drai_Gi: notes map2set_to_list[autoref_ga_rules] fixes S :: "('statei \ 'state) set" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ drai_dra_rel" shows "(?f :: ?'a, RETURN (dra_G A)) \ ?A" unfolding dra_G_ahs[where S = S and bhc = bhc] by (autoref_monadic (plain)) concrete_definition drai_Gi uses drai_Gi (* TODO: why are term local.drai_Gi and term BA_Nodes.drai_Gi not the same *) lemma drai_Gi_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" shows "(DRA_Nodes.drai_Gi seq bhc hms, dra_G) \ \L, S\ drai_dra_rel \ \unit_rel, S\ g_impl_rel_ext" using drai_Gi.refine[THEN RETURN_nres_relD] assms unfolding autoref_tag_defs by blast schematic_goal dra_nodes: fixes S :: "('statei \ 'state) set" assumes [simp]: "finite ((g_E (dra_G A))\<^sup>* `` g_V0 (dra_G A))" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ drai_dra_rel" shows "(?f :: ?'a, op_reachable (dra_G A)) \ ?R" by autoref concrete_definition dra_nodes uses dra_nodes lemma dra_nodes_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_PRECOND (finite (nodes A))" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" assumes "(Ai, A) \ \L, S\ drai_dra_rel" shows "(DRA_Nodes.dra_nodes seq bhc hms Ai, (OP nodes ::: \L, S\ drai_dra_rel \ \S\ ahs_rel bhc) $ A) \ \S\ ahs_rel bhc" proof - have "finite ((g_E (dra_G A))\<^sup>* `` g_V0 (dra_G A))" using assms(1) unfolding autoref_tag_defs dra_G_reachable_nodes[symmetric] by simp then show ?thesis using dra_nodes.refine assms unfolding autoref_tag_defs dra_G_reachable_nodes[symmetric] by blast qed end end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Refine.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Refine.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Refine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Refine.thy @@ -1,88 +1,88 @@ section \Relations on Deterministic Rabin Automata\ theory DRA_Refine imports "DRA" "../../Basic/Acceptance_Refine" "../../Transition_Systems/Transition_System_Refine" begin definition dra_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) dra \ ('label\<^sub>2, 'state\<^sub>2) dra) set" where [to_relAPP]: "dra_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabet A\<^sub>1, alphabet A\<^sub>2) \ \L\ set_rel \ (initial A\<^sub>1, initial A\<^sub>2) \ S \ - (succ A\<^sub>1, succ A\<^sub>2) \ L \ S \ S \ + (transition A\<^sub>1, transition A\<^sub>2) \ L \ S \ S \ (accepting A\<^sub>1, accepting A\<^sub>2) \ \rabin_rel S\ list_rel}" lemma dra_param[param]: "(dra, dra) \ \L\ set_rel \ S \ (L \ S \ S) \ \rabin_rel S\ list_rel \ \L, S\ dra_rel" "(alphabet, alphabet) \ \L, S\ dra_rel \ \L\ set_rel" "(initial, initial) \ \L, S\ dra_rel \ S" - "(succ, succ) \ \L, S\ dra_rel \ L \ S \ S" + "(transition, transition) \ \L, S\ dra_rel \ L \ S \ S" "(accepting, accepting) \ \L, S\ dra_rel \ \rabin_rel S\ list_rel" unfolding dra_rel_def fun_rel_def by auto lemma dra_rel_id[simp]: "\Id, Id\ dra_rel = Id" unfolding dra_rel_def using dra.expand by auto lemma dra_rel_comp[trans]: assumes [param]: "(A, B) \ \L\<^sub>1, S\<^sub>1\ dra_rel" "(B, C) \ \L\<^sub>2, S\<^sub>2\ dra_rel" shows "(A, C) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ dra_rel" proof - have "(accepting A, accepting B) \ \rabin_rel S\<^sub>1\ list_rel" by parametricity also have "(accepting B, accepting C) \ \rabin_rel S\<^sub>2\ list_rel" by parametricity finally have 1: "(accepting A, accepting C) \ \rabin_rel S\<^sub>1 O rabin_rel S\<^sub>2\ list_rel" by simp have 2: "rabin_rel S\<^sub>1 O rabin_rel S\<^sub>2 \ rabin_rel (S\<^sub>1 O S\<^sub>2)" by (force simp: fun_rel_def) have 3: "(accepting A, accepting C) \ \rabin_rel (S\<^sub>1 O S\<^sub>2)\ list_rel" using 1 2 list_rel_mono by blast - have "(succ A, succ B) \ L\<^sub>1 \ S\<^sub>1 \ S\<^sub>1" by parametricity - also have "(succ B, succ C) \ L\<^sub>2 \ S\<^sub>2 \ S\<^sub>2" by parametricity - finally have 4: "(succ A, succ C) \ L\<^sub>1 O L\<^sub>2 \ S\<^sub>1 O S\<^sub>2 \ S\<^sub>1 O S\<^sub>2" by this + have "(transition A, transition B) \ L\<^sub>1 \ S\<^sub>1 \ S\<^sub>1" by parametricity + also have "(transition B, transition C) \ L\<^sub>2 \ S\<^sub>2 \ S\<^sub>2" by parametricity + finally have 4: "(transition A, transition C) \ L\<^sub>1 O L\<^sub>2 \ S\<^sub>1 O S\<^sub>2 \ S\<^sub>1 O S\<^sub>2" by this show ?thesis unfolding dra_rel_def mem_Collect_eq prod.case set_rel_compp using 3 4 using dra_param(2 - 5)[THEN fun_relD, OF assms(1)] using dra_param(2 - 5)[THEN fun_relD, OF assms(2)] by auto qed lemma dra_rel_converse[simp]: "(\L, S\ dra_rel)\ = \L\, S\\ dra_rel" proof - have 1: "\L\ set_rel = (\L\\ set_rel)\" by simp have 2: "\S\ set_rel = (\S\\ set_rel)\" by simp have 3: "L \ S \ S = (L\ \ S\ \ S\)\" by simp have 4: "\rabin_rel S\ list_rel = (\rabin_rel (S\)\ list_rel)\" by simp show ?thesis unfolding dra_rel_def unfolding 3 unfolding 1 2 4 by fastforce qed lemma dra_rel_eq: "(A, A) \ \Id_on (alphabet A), Id_on (nodes A)\ dra_rel" unfolding dra_rel_def prod_rel_def using list_all2_same[to_set] by auto lemma enableds_param[param]: "(enableds, enableds) \ \L, S\ dra_rel \ S \ \L\ set_rel" using dra_param(2, 4) unfolding dra.enableds_def fun_rel_def set_rel_def by fastforce lemma paths_param[param]: "(paths, paths) \ \L, S\ dra_rel \ S \ \\L\ list_rel\ set_rel" unfolding paths_def by (intro fun_relI paths_param, fold enableds_def) (parametricity+) lemma runs_param[param]: "(runs, runs) \ \L, S\ dra_rel \ S \ \\L\ stream_rel\ set_rel" unfolding runs_def by (intro fun_relI runs_param, fold enableds_def) (parametricity+) lemma reachable_param[param]: "(reachable, reachable) \ \L, S\ dra_rel \ S \ \S\ set_rel" proof - have 1: "reachable A p = (\ w. target A w p) ` paths A p" for A :: "('label, 'state) dra" and p unfolding dra.reachable_alt_def dra.paths_def by auto show ?thesis unfolding 1 by parametricity qed lemma nodes_param[param]: "(nodes, nodes) \ \L, S\ dra_rel \ \S\ set_rel" proof - have 1: "nodes A = reachable A (initial A)" for A :: "('label, 'state) dra" unfolding dra.nodes_alt_def by simp show ?thesis unfolding 1 by parametricity qed lemma language_param[param]: "(language, language) \ \L, S\ dra_rel \ \\L\ stream_rel\ set_rel" proof - have 1: "language A = (\ w \ runs A (initial A). if cogen rabin (accepting A) (trace A w (initial A)) then {w} else {})" for A :: "('label, 'state) dra" unfolding language_def dra.runs_def image_def by auto show ?thesis unfolding 1 by parametricity qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Translate.thy b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Translate.thy --- a/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Translate.thy +++ b/thys/Transition_Systems_and_Automata/Automata/DRA/DRA_Translate.thy @@ -1,249 +1,249 @@ section \Explore and Enumerate Nodes of Deterministic Rabin Automata\ theory DRA_Translate imports DRA_Explicit begin subsection \Syntax\ (* TODO: this syntax has unnecessarily high inner binding strength, requiring extra parentheses the regular let syntax correctly uses inner binding strength 0: ("(2_ =/ _)" 10) *) no_syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" [1000, 13] 13) syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" 13) section \Image on Explicit Automata\ definition drae_image where "drae_image f A \ drae (alphabete A) (f (initiale A)) - ((\ (p, a, q). (f p, a, f q)) ` transe A) (map (map_prod (image f) (image f)) (acceptinge A))" + ((\ (p, a, q). (f p, a, f q)) ` transitione A) (map (map_prod (image f) (image f)) (acceptinge A))" lemma drae_image_param[param]: "(drae_image, drae_image) \ (S \ T) \ \L, S\ drae_rel \ \L, T\ drae_rel" unfolding drae_image_def by parametricity lemma drae_image_id[simp]: "drae_image id = id" unfolding drae_image_def by auto lemma drae_image_dra_drae: "drae_image f (dra_drae A) = drae (alphabet A) (f (initial A)) - (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}) + (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}) (map (\ (P, Q). (f ` {p \ nodes A. P p}, f ` {p \ nodes A. Q p})) (accepting A))" unfolding dra_drae_def drae_image_def drae.simps Set.filter_def by force section \Exploration and Translation\ definition trans_spec where - "trans_spec A f \ \ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}" + "trans_spec A f \ \ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}" definition trans_algo where "trans_algo N L S f \ FOREACH N (\ p T. do { ASSERT (p \ N); FOREACH L (\ a T. do { ASSERT (a \ L); let q = S a p; ASSERT ((f p, a, f q) \ T); RETURN (insert (f p, a, f q) T) } ) T } ) {}" lemma trans_algo_refine: assumes "finite (nodes A)" "finite (alphabet A)" "inj_on f (nodes A)" - assumes "N = nodes A" "L = alphabet A" "S = succ A" + assumes "N = nodes A" "L = alphabet A" "S = transition A" shows "(trans_algo N L S f, SPEC (HOL.eq (trans_spec A f))) \ \Id\ nres_rel" unfolding trans_algo_def trans_spec_def assms(4-6) proof (refine_vcg FOREACH_rule_insert_eq) show "finite (nodes A)" using assms(1) by this - show "(\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}) = - (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p})" by rule - show "(\ p \ {}. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}) = {}" by simp + show "(\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}) = + (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p})" by rule + show "(\ p \ {}. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}) = {}" by simp fix T x assume 1: "T \ nodes A" "x \ nodes A" "x \ T" show "finite (alphabet A)" using assms(2) by this - show "(\ a \ {}. f ` {x} \ {a} \ f ` {succ A a x}) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}) = - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p})" - "(\ a \ alphabet A. f ` {x} \ {a} \ f ` {succ A a x}) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}) = - (\ p \ insert x T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p})" by auto + show "(\ a \ {}. f ` {x} \ {a} \ f ` {transition A a x}) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}) = + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p})" + "(\ a \ alphabet A. f ` {x} \ {a} \ f ` {transition A a x}) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}) = + (\ p \ insert x T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p})" by auto fix Ta xa assume 2: "Ta \ alphabet A" "xa \ alphabet A" "xa \ Ta" - show "(f x, xa, f (succ A xa x)) \ (\ a \ Ta. f ` {x} \ {a} \ f ` {succ A a x}) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p})" + show "(f x, xa, f (transition A xa x)) \ (\ a \ Ta. f ` {x} \ {a} \ f ` {transition A a x}) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p})" using 1 2(3) assms(3) by (auto dest: inj_onD) - show "(\ a \ insert xa Ta. f ` {x} \ {a} \ f ` {succ A a x}) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}) = - insert (f x, xa, f (succ A xa x)) ((\ a \ Ta. f ` {x} \ {a} \ f ` {succ A a x}) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {succ A a p}))" + show "(\ a \ insert xa Ta. f ` {x} \ {a} \ f ` {transition A a x}) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}) = + insert (f x, xa, f (transition A xa x)) ((\ a \ Ta. f ` {x} \ {a} \ f ` {transition A a x}) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` {transition A a p}))" by simp qed definition to_draei :: "('state, 'label) dra \ ('state, 'label) dra" where "to_draei \ id" (* TODO: make separate implementations for "dra_drae" and "op_set_enumerate \ drae_image" *) schematic_goal to_draei_impl: fixes S :: "('statei \ 'state) set" assumes [simp]: "finite (nodes A)" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ drai_dra_rel" shows "(?f :: ?'a, do { let N = nodes A; f \ op_set_enumerate N; ASSERT (dom f = N); ASSERT (f (initial A) \ None); - ASSERT (\ a \ alphabet A. \ p \ dom f. f (succ A a p) \ None); - T \ trans_algo N (alphabet A) (succ A) (\ x. the (f x)); + ASSERT (\ a \ alphabet A. \ p \ dom f. f (transition A a p) \ None); + T \ trans_algo N (alphabet A) (transition A) (\ x. the (f x)); RETURN (drae (alphabet A) ((\ x. the (f x)) (initial A)) T (map (\ (P, Q). ((\ x. the (f x)) ` {p \ N. P p}, (\ x. the (f x)) ` {p \ N. Q p})) (accepting A))) }) \ ?R" unfolding trans_algo_def by (autoref_monadic (plain)) concrete_definition to_draei_impl uses to_draei_impl lemma to_draei_impl_refine'': fixes S :: "('statei \ 'state) set" assumes "finite (nodes A)" assumes "is_bounded_hashcode S seq bhc" assumes "is_valid_def_hm_size TYPE('statei) hms" assumes "(seq, HOL.eq) \ S \ S \ bool_rel" assumes "(Ai, A) \ \L, S\ drai_dra_rel" shows "(RETURN (to_draei_impl seq bhc hms Ai), do { f \ op_set_enumerate (nodes A); RETURN (drae_image (the \ f) (dra_drae A)) }) \ \\L, nat_rel\ draei_drae_rel\ nres_rel" proof - have 1: "finite (alphabet A)" using drai_dra_param(2)[param_fo, OF assms(5)] list_set_rel_finite unfolding finite_set_rel_def by auto note to_draei_impl.refine[OF assms] also have "(do { let N = nodes A; f \ op_set_enumerate N; ASSERT (dom f = N); ASSERT (f (initial A) \ None); - ASSERT (\ a \ alphabet A. \ p \ dom f. f (succ A a p) \ None); - T \ trans_algo N (alphabet A) (succ A) (\ x. the (f x)); + ASSERT (\ a \ alphabet A. \ p \ dom f. f (transition A a p) \ None); + T \ trans_algo N (alphabet A) (transition A) (\ x. the (f x)); RETURN (drae (alphabet A) ((\ x. the (f x)) (initial A)) T (map (\ (P, Q). ((\ x. the (f x)) ` {p \ N. P p}, (\ x. the (f x)) ` {p \ N. Q p})) (accepting A))) }, do { f \ op_set_enumerate (nodes A); T \ SPEC (HOL.eq (trans_spec A (\ x. the (f x)))); RETURN (drae (alphabet A) ((\ x. the (f x)) (initial A)) T (map (\ (P, Q). ((\ x. the (f x)) ` {p \ nodes A. P p}, (\ x. the (f x)) ` {p \ nodes A. Q p})) (accepting A))) }) \ \Id\ nres_rel" unfolding Let_def comp_apply op_set_enumerate_def using assms(1) 1 by (refine_vcg vcg0[OF trans_algo_refine]) (auto intro!: inj_on_map_the[unfolded comp_apply]) also have "(do { f \ op_set_enumerate (nodes A); T \ SPEC (HOL.eq (trans_spec A (\ x. the (f x)))); RETURN (drae (alphabet A) ((\ x. the (f x)) (initial A)) T (map (\ (P, Q). ((\ x. the (f x)) ` {p \ nodes A. P p}, (\ x. the (f x)) ` {p \ nodes A. Q p})) (accepting A))) }, do { f \ op_set_enumerate (nodes A); RETURN (drae_image (the \ f) (dra_drae A)) }) \ \Id\ nres_rel" unfolding trans_spec_def drae_image_dra_drae by refine_vcg force finally show ?thesis unfolding nres_rel_comp by simp qed (* TODO: generalize L *) context fixes Ai A fixes seq bhc hms fixes S :: "('statei \ 'state) set" assumes a: "finite (nodes A)" assumes b: "is_bounded_hashcode S seq bhc" assumes c: "is_valid_def_hm_size TYPE('statei) hms" assumes d: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes e: "(Ai, A) \ \Id, S\ drai_dra_rel" begin definition f' where "f' \ SOME f'. (to_draei_impl seq bhc hms Ai, drae_image (the \ f') (dra_drae A)) \ \Id, nat_rel\ draei_drae_rel \ dom f' = nodes A \ inj_on f' (nodes A)" lemma 1: "\ f'. (to_draei_impl seq bhc hms Ai, drae_image (the \ f') (dra_drae A)) \ \Id, nat_rel\ draei_drae_rel \ dom f' = nodes A \ inj_on f' (nodes A)" using to_draei_impl_refine''[ OF a b c d e, unfolded op_set_enumerate_def bind_RES_RETURN_eq, THEN nres_relD, THEN RETURN_ref_SPECD] by force lemma f'_refine: "(to_draei_impl seq bhc hms Ai, drae_image (the \ f') (dra_drae A)) \ \Id, nat_rel\ draei_drae_rel" using someI_ex[OF 1, folded f'_def] by auto lemma f'_dom: "dom f' = nodes A" using someI_ex[OF 1, folded f'_def] by auto lemma f'_inj: "inj_on f' (nodes A)" using someI_ex[OF 1, folded f'_def] by auto definition f where "f \ the \ f'" definition g where "g = inv_into (nodes A) f" lemma inj_f[intro!, simp]: "inj_on f (nodes A)" using f'_inj f'_dom unfolding f_def by (simp add: inj_on_map_the) lemma inj_g[intro!, simp]: "inj_on g (f ` nodes A)" unfolding g_def by (simp add: inj_on_inv_into) definition rel where "rel \ {(f p, p) |p. p \ nodes A}" lemma rel_alt_def: "rel = (br f (\ p. p \ nodes A))\" unfolding rel_def by (auto simp: in_br_conv) lemma rel_inv_def: "rel = br g (\ k. k \ f ` nodes A)" unfolding rel_alt_def g_def by (auto simp: in_br_conv) lemma rel_domain[simp]: "Domain rel = f ` nodes A" unfolding rel_def by force lemma rel_range[simp]: "Range rel = nodes A" unfolding rel_def by auto lemma [intro!, simp]: "bijective rel" unfolding rel_inv_def by (simp add: bijective_alt) lemma [simp]: "Id_on (f ` nodes A) O rel = rel" unfolding rel_def by auto lemma [simp]: "rel O Id_on (nodes A) = rel" unfolding rel_def by auto lemma [param]: "(f, f) \ Id_on (Range rel) \ Id_on (Domain rel)" unfolding rel_alt_def by auto lemma [param]: "(g, g) \ Id_on (Domain rel) \ Id_on (Range rel)" unfolding rel_inv_def by auto lemma [param]: "(id, f) \ rel \ Id_on (Domain rel)" unfolding rel_alt_def by (auto simp: in_br_conv) lemma [param]: "(f, id) \ Id_on (Range rel) \ rel" unfolding rel_alt_def by (auto simp: in_br_conv) lemma [param]: "(id, g) \ Id_on (Domain rel) \ rel" unfolding rel_inv_def by (auto simp: in_br_conv) lemma [param]: "(g, id) \ rel \ Id_on (Range rel)" unfolding rel_inv_def by (auto simp: in_br_conv) lemma to_draei_impl_refine': "(to_draei_impl seq bhc hms Ai, to_draei A) \ \Id_on (alphabet A), rel\ draei_dra_rel" proof - have 1: "(draei_drae (to_draei_impl seq bhc hms Ai), id (drae_image f (dra_drae A))) \ \Id, nat_rel\ drae_rel" using f'_refine[folded f_def] by parametricity have 2: "(draei_drae (to_draei_impl seq bhc hms Ai), id (drae_image f (dra_drae A))) \ \Id_on (alphabet A), Id_on (f ` nodes A)\ drae_rel" using 1 unfolding drae_rel_def dra_drae_def drae_image_def by auto - have 3: "wft (alphabet A) (nodes A) (transe (dra_drae A))" + have 3: "wft (alphabet A) (nodes A) (transitione (dra_drae A))" using wft_transitions unfolding dra_drae_def drae.sel by this - have 4: "(wft (alphabet A) (f ` nodes A) (transe (drae_image f (dra_drae A))), - wft (alphabet A) (id ` nodes A) (transe (drae_image id (dra_drae A)))) \ bool_rel" + have 4: "(wft (alphabet A) (f ` nodes A) (transitione (drae_image f (dra_drae A))), + wft (alphabet A) (id ` nodes A) (transitione (drae_image id (dra_drae A)))) \ bool_rel" using dra_rel_eq by parametricity auto - have 5: "wft (alphabet A) (f ` nodes A) (transe (drae_image f (dra_drae A)))" using 3 4 by simp + have 5: "wft (alphabet A) (f ` nodes A) (transitione (drae_image f (dra_drae A)))" using 3 4 by simp have "(drae_dra (draei_drae (to_draei_impl seq bhc hms Ai)), drae_dra (id (drae_image f (dra_drae A)))) \ \Id_on (alphabet A), Id_on (f ` nodes A)\ dra_rel" using 2 5 by parametricity auto also have "(drae_dra (id (drae_image f (dra_drae A))), drae_dra (id (drae_image id (dra_drae A)))) \ \Id_on (alphabet A), rel\ dra_rel" using dra_rel_eq 3 by parametricity auto also have "drae_dra (id (drae_image id (dra_drae A))) = (drae_dra \ dra_drae) A" by simp also have "(\, id A) \ \Id_on (alphabet A), Id_on (nodes A)\ dra_rel" by parametricity also have "id A = to_draei A" unfolding to_draei_def by simp finally show ?thesis unfolding draei_dra_rel_def by simp qed end context begin interpretation autoref_syn by this lemma to_draei_impl_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_PRECOND (finite (nodes A))" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" assumes "(Ai, A) \ \Id, S\ drai_dra_rel" shows "(to_draei_impl seq bhc hms Ai, (OP to_draei ::: \Id, S\ drai_dra_rel \ \Id_on (alphabet A), rel Ai A seq bhc hms\ draei_dra_rel) $ A) \ \Id_on (alphabet A), rel Ai A seq bhc hms\ draei_dra_rel" using to_draei_impl_refine' assms unfolding autoref_tag_defs by this end end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA.thy @@ -1,142 +1,142 @@ section \Nondeterministic Büchi Automata\ theory NBA imports "../../Basic/Sequence_Zip" "../../Basic/Acceptance" "../../Transition_Systems/Transition_System" "../../Transition_Systems/Transition_System_Extra" "../../Transition_Systems/Transition_System_Construction" begin datatype ('label, 'state) nba = nba (alphabet: "'label set") (initial: "'state set") - (succ: "'label \ 'state \ 'state set") + (transition: "'label \ 'state \ 'state set") (accepting: "'state pred") global_interpretation nba: transition_system_initial - "\ a p. snd a" "\ a p. fst a \ alphabet A \ snd a \ succ A (fst a) p" "\ p. p \ initial A" + "\ a p. snd a" "\ a p. fst a \ alphabet A \ snd a \ transition A (fst a) p" "\ p. p \ initial A" for A defines path = nba.path and run = nba.run and reachable = nba.reachable and nodes = nba.nodes and enableds = nba.enableds and paths = nba.paths and runs = nba.runs by this abbreviation "target \ nba.target" abbreviation "states \ nba.states" abbreviation "trace \ nba.trace" lemma states_alt_def: "states r p = map snd r" by (induct r arbitrary: p) (auto) lemma trace_alt_def: "trace r p = smap snd r" by (coinduction arbitrary: r p) (auto) abbreviation successors :: "('label, 'state) nba \ 'state \ 'state set" where "successors \ nba.successors TYPE('label)" - lemma successors_alt_def: "successors A p = (\ a \ alphabet A. succ A a p)" by auto + lemma successors_alt_def: "successors A p = (\ a \ alphabet A. transition A a p)" by auto - lemma reachable_succ[intro]: - assumes "a \ alphabet A" "q \ reachable A p" "r \ succ A a q" + lemma reachable_step[intro]: + assumes "a \ alphabet A" "q \ reachable A p" "r \ transition A a q" shows "r \ reachable A p" using nba.reachable.execute assms by force - lemma nodes_succ[intro]: - assumes "a \ alphabet A" "p \ nodes A" "q \ succ A a p" + lemma nodes_step[intro]: + assumes "a \ alphabet A" "p \ nodes A" "q \ transition A a p" shows "q \ nodes A" using nba.nodes.execute assms by force definition language :: "('label, 'state) nba \ 'label stream set" where "language A \ {w |w r p. p \ initial A \ run A (w ||| r) p \ infs (accepting A) (trace (w ||| r) p)}" lemma language[intro]: assumes "p \ initial A" "run A (w ||| r) p" "infs (accepting A) (trace (w ||| r) p)" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains r p where "p \ initial A" "run A (w ||| r) p" "infs (accepting A) (trace (w ||| r) p)" using assms unfolding language_def by auto lemma run_alphabet: assumes "run A (w ||| r) p" shows "w \ streams (alphabet A)" using assms by (coinduction arbitrary: w r p) (metis nba.run.cases stream.map szip_smap szip_smap_fst) lemma language_alphabet: "language A \ streams (alphabet A)" unfolding language_def by (auto dest: run_alphabet) instantiation nba :: (type, type) order begin definition less_eq_nba :: "('a, 'b) nba \ ('a, 'b) nba \ bool" where "A \ B \ alphabet A \ alphabet B \ initial A \ initial B \ - succ A \ succ B \ accepting A \ accepting B" + transition A \ transition B \ accepting A \ accepting B" definition less_nba :: "('a, 'b) nba \ ('a, 'b) nba \ bool" where "less_nba A B \ A \ B \ A \ B" instance by (intro_classes) (auto simp: less_eq_nba_def less_nba_def nba.expand) end lemma nodes_mono: "mono nodes" proof fix A B :: "('label, 'state) nba" assume 1: "A \ B" have 2: "alphabet A \ alphabet B" using 1 unfolding less_eq_nba_def by auto have 3: "initial A \ initial B" using 1 unfolding less_eq_nba_def by auto - have 4: "succ A a p \ succ B a p" for a p using 1 unfolding less_eq_nba_def le_fun_def by auto + have 4: "transition A a p \ transition B a p" for a p using 1 unfolding less_eq_nba_def le_fun_def by auto have 5: "p \ nodes B" if "p \ nodes A" for p using that 2 3 4 by induct fastforce+ show "nodes A \ nodes B" using 5 by auto qed lemma language_mono: "mono language" proof fix A B :: "('label, 'state) nba" assume 1: "A \ B" have 2: "alphabet A \ alphabet B" using 1 unfolding less_eq_nba_def by auto have 3: "initial A \ initial B" using 1 unfolding less_eq_nba_def by auto - have 4: "succ A a p \ succ B a p" for a p using 1 unfolding less_eq_nba_def le_fun_def by auto + have 4: "transition A a p \ transition B a p" for a p using 1 unfolding less_eq_nba_def le_fun_def by auto have 5: "accepting A p \ accepting B p" for p using 1 unfolding less_eq_nba_def by auto have 6: "run B wr p" if "run A wr p" for wr p using that 2 4 by coinduct auto have 7: "infs (accepting B) w" if "infs (accepting A) w" for w using infs_mono that 5 by metis show "language A \ language B" using 3 6 7 by blast qed lemma simulation_language: assumes "alphabet A \ alphabet B" assumes "\ p. p \ initial A \ \ q \ initial B. (p, q) \ R" - assumes "\ a p p' q. p' \ succ A a p \ (p, q) \ R \ \ q' \ succ B a q. (p', q') \ R" + assumes "\ a p p' q. p' \ transition A a p \ (p, q) \ R \ \ q' \ transition B a q. (p', q') \ R" assumes "\ p q. (p, q) \ R \ accepting A p \ accepting B q" shows "language A \ language B" proof fix w assume 1: "w \ language A" obtain r p where 2: "p \ initial A" "run A (w ||| r) p" "infs (accepting A) (trace (w ||| r) p)" using 1 by rule define P where "P n q \ (target (stake n (w ||| r)) p, q) \ R" for n q obtain q where 3: "q \ initial B" "(p, q) \ R" using assms(2) 2(1) by auto obtain ws where 4: "run B ws q" "\ i. P (0 + i) (target (stake i ws) q)" "\ i. fst (ws !! i) = w !! (0 + i)" proof (rule nba.invariant_run_index) have "stake k (w ||| r) @- (w !! k, target (stake (Suc k) (w ||| r)) p) ## sdrop (Suc k) (w ||| r) = w ||| r" for k by (metis id_stake_snth_sdrop snth_szip sscan_snth szip_smap_snd trace_alt_def) also have "run A \ p" using 2(2) by this - finally show "\ a. (fst a \ alphabet B \ snd a \ succ B (fst a) q) \ + finally show "\ a. (fst a \ alphabet B \ snd a \ transition B (fst a) q) \ P (Suc n) (snd a) \ fst a = w !! n" if "P n q" for n q using assms(1, 3) that unfolding P_def by fastforce show "P 0 q" unfolding P_def using 3(2) by auto qed rule obtain s where 5: "ws = w ||| s" using 4(3) by (metis add.left_neutral eqI_snth snth_smap szip_smap) show "w \ language B" proof show "q \ initial B" using 3(1) by this show "run B (w ||| s) q" using 4(1) unfolding 5 by this have 6: "(\ a b. (a, b) \ R) \ (\ a b. accepting A a \ accepting B b)" using assms(4) by auto have 7: "stream_all2 (\ p q. (p, q) \ R) (trace (w ||| r) p) (trace (w ||| s) q)" using 4(2) unfolding P_def 5 by (simp add: stream_rel_snth del: stake.simps(2)) have 8: "stream_all2 (\ a b. accepting A a \ accepting B b) (trace (w ||| r) p) (trace (w ||| s) q)" using stream.rel_mono 6 7 by auto show "infs (accepting B) (trace (w ||| s) q)" using infs_mono_strong 8 2(3) by this qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Algorithms.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Algorithms.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Algorithms.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Algorithms.thy @@ -1,135 +1,136 @@ section \Algorithms on Nondeterministic Büchi Automata\ theory NBA_Algorithms imports NBA_Graphs NBA_Implement DFS_Framework.Reachable_Nodes Gabow_SCC.Gabow_GBG_Code begin subsection \Miscellaneous Amendments\ lemma (in igb_fr_graph) acc_run_lasso_prpl: "Ex is_acc_run \ Ex is_lasso_prpl" using accepted_lasso is_lasso_prpl_of_lasso by blast lemma (in igb_fr_graph) lasso_prpl_acc_run_iff: "Ex is_lasso_prpl \ Ex is_acc_run" using acc_run_lasso_prpl lasso_prpl_acc_run by auto lemma [autoref_rel_intf]: "REL_INTF igbg_impl_rel_ext i_igbg" by (rule REL_INTFI) subsection \Operations\ definition [simp]: "op_empty A \ language A = {}" lemmas [autoref_op_pat] = op_empty_def[symmetric] subsection \Implementations\ context begin interpretation autoref_syn by this lemma nba_g_ahs: "nba_g A = \ g_V = UNIV, g_E = E_of_succ (\ p. CAST - ((\ a \ alphabet A. succ A a p ::: \S\ list_set_rel) ::: \S\ ahs_rel bhc)), g_V0 = initial A \" + ((\ a \ alphabet A. transition A a p ::: \S\ list_set_rel) ::: \S\ ahs_rel bhc)), + g_V0 = initial A \" unfolding nba_g_def successors_alt_def CAST_def id_apply autoref_tag_defs by rule schematic_goal nbai_gi: notes [autoref_ga_rules] = map2set_to_list fixes S :: "('statei \ 'state) set" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(?f :: ?'a, RETURN (nba_g A)) \ ?A" unfolding nba_g_ahs[where S = S and bhc = bhc] by (autoref_monadic (plain)) concrete_definition nbai_gi uses nbai_gi lemma nbai_gi_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" shows "(NBA_Algorithms.nbai_gi seq bhc hms, nba_g) \ \L, S\ nbai_nba_rel \ \unit_rel, S\ g_impl_rel_ext" using nbai_gi.refine[THEN RETURN_nres_relD] assms unfolding autoref_tag_defs by blast schematic_goal nba_nodes: fixes S :: "('statei \ 'state) set" assumes [simp]: "finite ((g_E (nba_g A))\<^sup>* `` g_V0 (nba_g A))" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(?f :: ?'a, op_reachable (nba_g A)) \ ?R" by autoref concrete_definition nba_nodes uses nba_nodes lemma nba_nodes_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_PRECOND (finite (nodes A))" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" assumes "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(NBA_Algorithms.nba_nodes seq bhc hms Ai, (OP nodes ::: \L, S\ nbai_nba_rel \ \S\ ahs_rel bhc) $ A) \ \S\ ahs_rel bhc" proof - have 1: "nodes A = op_reachable (nba_g A)" by (auto simp: nba_g_V0 nba_g_E_rtrancl) have 2: "finite ((g_E (nba_g A))\<^sup>* `` g_V0 (nba_g A))" using assms(1) unfolding 1 by simp show ?thesis using nba_nodes.refine assms 2 unfolding autoref_tag_defs 1 by blast qed lemma nba_igbg_ahs: "nba_igbg A = \ g_V = UNIV, g_E = E_of_succ (\ p. CAST - ((\ a \ alphabet A. succ A a p ::: \S\ list_set_rel) ::: \S\ ahs_rel bhc)), g_V0 = initial A, + ((\ a \ alphabet A. transition A a p ::: \S\ list_set_rel) ::: \S\ ahs_rel bhc)), g_V0 = initial A, igbg_num_acc = 1, igbg_acc = \ p. if accepting A p then {0} else {} \" unfolding nba_g_def nba_igbg_def successors_alt_def CAST_def id_apply autoref_tag_defs unfolding graph_rec.defs by simp schematic_goal nbai_igbgi: notes [autoref_ga_rules] = map2set_to_list fixes S :: "('statei \ 'state) set" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(?f :: ?'a, RETURN (nba_igbg A)) \ ?A" unfolding nba_igbg_ahs[where S = S and bhc = bhc] by (autoref_monadic (plain)) concrete_definition nbai_igbgi uses nbai_igbgi lemma nbai_igbgi_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" shows "(NBA_Algorithms.nbai_igbgi seq bhc hms, nba_igbg) \ \L, S\ nbai_nba_rel \ igbg_impl_rel_ext unit_rel S" using nbai_igbgi.refine[THEN RETURN_nres_relD] assms unfolding autoref_tag_defs by blast schematic_goal nba_empty: fixes A :: "('label, 'state :: hashable) nba" assumes [simp]: "igb_fr_graph (nba_igbg A)" assumes [autoref_rules]: "(Ai, A) \ \L, Id\ nbai_nba_rel" shows "(?f :: ?'a, do { r \ op_find_lasso_spec (nba_igbg A); RETURN (r = None)}) \ ?A" by (autoref_monadic (plain)) concrete_definition nba_empty uses nba_empty lemma nba_empty_refine[autoref_rules]: assumes "SIDE_PRECOND (igb_fr_graph (nba_igbg A))" assumes "(Ai, A) \ \L, Id\ nbai_nba_rel" shows "(NBA_Algorithms.nba_empty Ai, (OP op_empty ::: \L, S\ nbai_nba_rel \ bool_rel) $ A) \ bool_rel" proof - interpret igb_fr_graph "nba_igbg A" using assms(1) by simp have "(RETURN (NBA_Algorithms.nba_empty Ai), do { r \ find_lasso_spec; RETURN (r = None) }) \ \bool_rel\ nres_rel" using nba_empty.refine assms by simp also have "(do { r \ find_lasso_spec; RETURN (r = None) }, RETURN (\ Ex is_lasso_prpl)) \ \bool_rel\ nres_rel" unfolding find_lasso_spec_def by (refine_vcg) (auto split: option.splits) finally have "NBA_Algorithms.nba_empty Ai \ \ Ex is_lasso_prpl" unfolding nres_rel_comp using RETURN_nres_relD by force also have "\ \ \ Ex is_acc_run" using lasso_prpl_acc_run_iff by auto also have "\ \ language A = {}" using acc_run_language is_igb_graph by auto finally show ?thesis by simp qed end end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Explicit.thy @@ -1,162 +1,162 @@ section \Explicit Nondeterministic Büchi Automata\ theory NBA_Explicit imports NBA_Algorithms begin datatype ('label, 'state) nbae = nbae (alphabete: "'label set") (initiale: "'state set") - (transe: "('state \ 'label \ 'state) set") + (transitione: "('state \ 'label \ 'state) set") (acceptinge: "'state set") definition nbae_rel where [to_relAPP]: "nbae_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabete A\<^sub>1, alphabete A\<^sub>2) \ \L\ set_rel \ (initiale A\<^sub>1, initiale A\<^sub>2) \ \S\ set_rel \ - (transe A\<^sub>1, transe A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ + (transitione A\<^sub>1, transitione A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ (acceptinge A\<^sub>1, acceptinge A\<^sub>2) \ \S\ set_rel}" lemma nbae_param[param, autoref_rules]: "(nbae, nbae) \ \L\ set_rel \ \S\ set_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel \ \S\ set_rel \ \L, S\ nbae_rel" "(alphabete, alphabete) \ \L, S\ nbae_rel \ \L\ set_rel" "(initiale, initiale) \ \L, S\ nbae_rel \ \S\ set_rel" - "(transe, transe) \ \L, S\ nbae_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel" + "(transitione, transitione) \ \L, S\ nbae_rel \ \S \\<^sub>r L \\<^sub>r S\ set_rel" "(acceptinge, acceptinge) \ \L, S\ nbae_rel \ \S\ set_rel" unfolding nbae_rel_def by auto lemma nbae_rel_id[simp]: "\Id, Id\ nbae_rel = Id" unfolding nbae_rel_def using nbae.expand by auto lemma nbae_rel_comp[simp]: "\L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ nbae_rel = \L\<^sub>1, S\<^sub>1\ nbae_rel O \L\<^sub>2, S\<^sub>2\ nbae_rel" proof safe fix A B assume 1: "(A, B) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ nbae_rel" obtain a b c d where 2: "(alphabete A, a) \ \L\<^sub>1\ set_rel" "(a, alphabete B) \ \L\<^sub>2\ set_rel" "(initiale A, b) \ \S\<^sub>1\ set_rel" "(b, initiale B) \ \S\<^sub>2\ set_rel" - "(transe A, c) \ \S\<^sub>1 \\<^sub>r L\<^sub>1 \\<^sub>r S\<^sub>1\ set_rel" "(c, transe B) \ \S\<^sub>2 \\<^sub>r L\<^sub>2 \\<^sub>r S\<^sub>2\ set_rel" + "(transitione A, c) \ \S\<^sub>1 \\<^sub>r L\<^sub>1 \\<^sub>r S\<^sub>1\ set_rel" "(c, transitione B) \ \S\<^sub>2 \\<^sub>r L\<^sub>2 \\<^sub>r S\<^sub>2\ set_rel" "(acceptinge A, d) \ \S\<^sub>1\ set_rel" "(d, acceptinge B) \ \S\<^sub>2\ set_rel" using 1 unfolding nbae_rel_def prod_rel_compp set_rel_compp by auto show "(A, B) \ \L\<^sub>1, S\<^sub>1\ nbae_rel O \L\<^sub>2, S\<^sub>2\ nbae_rel" proof show "(A, nbae a b c d) \ \L\<^sub>1, S\<^sub>1\ nbae_rel" using 2 unfolding nbae_rel_def by auto show "(nbae a b c d, B) \ \L\<^sub>2, S\<^sub>2\ nbae_rel" using 2 unfolding nbae_rel_def by auto qed next show "(A, C) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ nbae_rel" if "(A, B) \ \L\<^sub>1, S\<^sub>1\ nbae_rel" "(B, C) \ \L\<^sub>2, S\<^sub>2\ nbae_rel" for A B C using that unfolding nbae_rel_def prod_rel_compp set_rel_compp by auto qed (* TODO: why do we need all this setup? can't i_of_rel do the trick? *) consts i_nbae_scheme :: "interface \ interface \ interface" context begin interpretation autoref_syn by this lemma nbae_scheme_itype[autoref_itype]: "nbae ::\<^sub>i \L\\<^sub>i i_set \\<^sub>i \S\\<^sub>i i_set \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set \\<^sub>i \S\\<^sub>i i_set \\<^sub>i \L, S\\<^sub>i i_nbae_scheme" "alphabete ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \L\\<^sub>i i_set" "initiale ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \S\\<^sub>i i_set" - "transe ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set" + "transitione ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \\S, \L, S\\<^sub>i i_prod\\<^sub>i i_prod\\<^sub>i i_set" "acceptinge ::\<^sub>i \L, S\\<^sub>i i_nbae_scheme \\<^sub>i \S\\<^sub>i i_set" by auto end datatype ('label, 'state) nbaei = nbaei (alphabetei: "'label list") (initialei: "'state list") - (transei: "('state \ 'label \ 'state) list") + (transitionei: "('state \ 'label \ 'state) list") (acceptingei: "'state list") definition nbaei_rel where [to_relAPP]: "nbaei_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabetei A\<^sub>1, alphabetei A\<^sub>2) \ \L\ list_rel \ (initialei A\<^sub>1, initialei A\<^sub>2) \ \S\ list_rel \ - (transei A\<^sub>1, transei A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ + (transitionei A\<^sub>1, transitionei A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ (acceptingei A\<^sub>1, acceptingei A\<^sub>2) \ \S\ list_rel}" lemma nbaei_param[param, autoref_rules]: "(nbaei, nbaei) \ \L\ list_rel \ \S\ list_rel \ \S \\<^sub>r L \\<^sub>r S\ list_rel \ \S\ list_rel \ \L, S\ nbaei_rel" "(alphabetei, alphabetei) \ \L, S\ nbaei_rel \ \L\ list_rel" "(initialei, initialei) \ \L, S\ nbaei_rel \ \S\ list_rel" - "(transei, transei) \ \L, S\ nbaei_rel \ \S \\<^sub>r L \\<^sub>r S\ list_rel" + "(transitionei, transitionei) \ \L, S\ nbaei_rel \ \S \\<^sub>r L \\<^sub>r S\ list_rel" "(acceptingei, acceptingei) \ \L, S\ nbaei_rel \ \S\ list_rel" unfolding nbaei_rel_def by auto definition nbaei_nbae_rel where [to_relAPP]: "nbaei_nbae_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabetei A\<^sub>1, alphabete A\<^sub>2) \ \L\ list_set_rel \ (initialei A\<^sub>1, initiale A\<^sub>2) \ \S\ list_set_rel \ - (transei A\<^sub>1, transe A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ + (transitionei A\<^sub>1, transitione A\<^sub>2) \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ (acceptingei A\<^sub>1, acceptinge A\<^sub>2) \ \S\ list_set_rel}" lemmas [autoref_rel_intf] = REL_INTFI[of nbaei_nbae_rel i_nbae_scheme] lemma nbaei_nbae_param[param, autoref_rules]: "(nbaei, nbae) \ \L\ list_set_rel \ \S\ list_set_rel \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel \ \S\ list_set_rel \ \L, S\ nbaei_nbae_rel" "(alphabetei, alphabete) \ \L, S\ nbaei_nbae_rel \ \L\ list_set_rel" "(initialei, initiale) \ \L, S\ nbaei_nbae_rel \ \S\ list_set_rel" - "(transei, transe) \ \L, S\ nbaei_nbae_rel \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel" + "(transitionei, transitione) \ \L, S\ nbaei_nbae_rel \ \S \\<^sub>r L \\<^sub>r S\ list_set_rel" "(acceptingei, acceptinge) \ \L, S\ nbaei_nbae_rel \ \S\ list_set_rel" unfolding nbaei_nbae_rel_def by auto definition nbaei_nbae where "nbaei_nbae A \ nbae (set (alphabetei A)) (set (initialei A)) - (set (transei A)) (set (acceptingei A))" + (set (transitionei A)) (set (acceptingei A))" lemma nbaei_nbae_id_param[param]: "(nbaei_nbae, id) \ \L, S\ nbaei_nbae_rel \ \L, S\ nbae_rel" proof fix Ai A assume 1: "(Ai, A) \ \L, S\ nbaei_nbae_rel" have 2: "nbaei_nbae Ai = nbae (set (alphabetei Ai)) (set (initialei Ai)) - (set (transei Ai)) (set (acceptingei Ai))" unfolding nbaei_nbae_def by rule + (set (transitionei Ai)) (set (acceptingei Ai))" unfolding nbaei_nbae_def by rule have 3: "id A = nbae (id (alphabete A)) (id (initiale A)) - (id (transe A)) (id (acceptinge A))" by simp + (id (transitione A)) (id (acceptinge A))" by simp show "(nbaei_nbae Ai, id A) \ \L, S\ nbae_rel" unfolding 2 3 using 1 by parametricity qed abbreviation "transitions L S s \ \ a \ L. \ p \ S. {p} \ {a} \ s a p" abbreviation "succs T a p \ (T `` {p}) `` {a}" definition nba_nbae where "nba_nbae A \ nbae (alphabet A) (initial A) - (transitions (alphabet A) (nodes A) (succ A)) (Set.filter (accepting A) (nodes A))" + (transitions (alphabet A) (nodes A) (transition A)) (Set.filter (accepting A) (nodes A))" definition nbae_nba where "nbae_nba A \ nba (alphabete A) (initiale A) - (succs (transe A)) (\ p. p \ acceptinge A)" + (succs (transitione A)) (\ p. p \ acceptinge A)" lemma nba_nbae_param[param]: "(nba_nbae, nba_nbae) \ \L, S\ nba_rel \ \L, S\ nbae_rel" unfolding nba_nbae_def by parametricity lemma nbae_nba_param[param]: assumes "bijective L" "bijective S" shows "(nbae_nba, nbae_nba) \ \L, S\ nbae_rel \ \L, S\ nba_rel" using assms assms(2)[unfolded bijective_alt] unfolding nbae_nba_def by parametricity auto lemma nbae_nba_nba_nbae_param[param]: "((nbae_nba \ nba_nbae) A, id A) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" proof - have "(nbae_nba \ nba_nbae) A = nba (alphabet A) (initial A) - (succs (transitions (alphabet A) (nodes A) (succ A))) (\ p. p \ Set.filter (accepting A) (nodes A))" + (succs (transitions (alphabet A) (nodes A) (transition A))) (\ p. p \ Set.filter (accepting A) (nodes A))" unfolding nbae_nba_def nba_nbae_def by simp - also have "(\, nba (alphabet A) (initial A) (succ A) (accepting A)) \ + also have "(\, nba (alphabet A) (initial A) (transition A) (accepting A)) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" using nba_rel_eq by parametricity auto - also have "nba (alphabet A) (initial A) (succ A) (accepting A) = id A" by simp + also have "nba (alphabet A) (initial A) (transition A) (accepting A) = id A" by simp finally show ?thesis by this qed definition nbaei_nba_rel where [to_relAPP]: "nbaei_nba_rel L S \ {(Ae, A). (nbae_nba (nbaei_nbae Ae), A) \ \L, S\ nba_rel}" lemma nbaei_nba_id[param]: "(nbae_nba \ nbaei_nbae, id) \ \L, S\ nbaei_nba_rel \ \L, S\ nba_rel" unfolding nbaei_nba_rel_def by auto schematic_goal nbae_nba_impl: "(?f, nbae_nba) \ \Id, Id\ nbaei_nbae_rel \ \Id, Id\ nbai_nba_rel" unfolding nbae_nba_def by autoref concrete_definition nbae_nba_impl uses nbae_nba_impl end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy @@ -1,150 +1,150 @@ section \Connecting Nondeterministic Büchi Automata to CAVA Automata Structures\ theory NBA_Graphs imports NBA CAVA_Automata.Automata_Impl begin no_notation build (infixr "##" 65) subsection \Regular Graphs\ definition nba_g :: "('label, 'state) nba \ 'state graph_rec" where "nba_g A \ \ g_V = UNIV, g_E = E_of_succ (successors A), g_V0 = initial A \" lemma nba_g_graph[simp]: "graph (nba_g A)" unfolding nba_g_def graph_def by simp lemma nba_g_V0: "g_V0 (nba_g A) = initial A" unfolding nba_g_def by simp lemma nba_g_E_rtrancl: "(g_E (nba_g A))\<^sup>* = {(p, q). q \ reachable A p}" unfolding nba_g_def graph_rec.simps E_of_succ_def proof safe show "(p, q) \ {(p, q). q \ successors A p}\<^sup>*" if "q \ reachable A p" for p q using that by (induct) (auto intro: rtrancl_into_rtrancl) show "q \ reachable A p" if "(p, q) \ {(p, q). q \ successors A p}\<^sup>*" for p q using that by induct auto qed lemma nba_g_rtrancl_path: "(g_E (nba_g A))\<^sup>* = {(p, target r p) |r p. NBA.path A r p}" unfolding nba_g_E_rtrancl by blast lemma nba_g_trancl_path: "(g_E (nba_g A))\<^sup>+ = {(p, target r p) |r p. NBA.path A r p \ r \ []}" unfolding nba_g_def graph_rec.simps E_of_succ_def proof safe show "\ r p. (x, y) = (p, target r p) \ NBA.path A r p \ r \ []" if "(x, y) \ {(p, q). q \ successors A p}\<^sup>+" for x y using that proof induct case (base y) - obtain a where 1: "a \ alphabet A" "y \ succ A a x" using base by auto + obtain a where 1: "a \ alphabet A" "y \ transition A a x" using base by auto show ?case proof (intro exI conjI) show "(x, y) = (x, target [(a, y)] x)" by simp show "NBA.path A [(a, y)] x" using 1 by auto show "[(a, y)] \ []" by simp qed next case (step y z) obtain r where 1: "y = target r x" "NBA.path A r x" "r \ []" using step(3) by auto - obtain a where 2: "a \ alphabet A" "z \ succ A a y" using step(2) by auto + obtain a where 2: "a \ alphabet A" "z \ transition A a y" using step(2) by auto show ?case proof (intro exI conjI) show "(x, z) = (x, target (r @ [(a, z)]) x)" by simp show "NBA.path A (r @ [(a, z)]) x" using 1 2 by auto show "r @ [(a, z)] \ []" by simp qed qed show "(p, target r p) \ {(u, v). v \ successors A u}\<^sup>+" if "NBA.path A r p" "r \ []" for r p using that by (induct) (fastforce intro: trancl_into_trancl2)+ qed lemma nba_g_ipath_run: assumes "ipath (g_E (nba_g A)) r" obtains w where "run A (w ||| smap (r \ Suc) nats) (r 0)" proof - - have 1: "\ a \ alphabet A. r (Suc i) \ succ A a (r i)" for i + have 1: "\ a \ alphabet A. r (Suc i) \ transition A a (r i)" for i using assms unfolding ipath_def nba_g_def E_of_succ_def by auto obtain wr where 2: "run A wr (r 0)" "\ i. target (stake i wr) (r 0) = r i" proof (rule nba.invariant_run_index) - show "\ aq. (fst aq \ alphabet A \ snd aq \ succ A (fst aq) p) \ snd aq = r (Suc i) \ True" + show "\ aq. (fst aq \ alphabet A \ snd aq \ transition A (fst aq) p) \ snd aq = r (Suc i) \ True" if "p = r i" for i p using that 1 by auto show "r 0 = r 0" by rule qed auto have 3: "smap (r \ Suc) nats = smap snd wr" proof (rule eqI_snth) fix i have "smap (r \ Suc) nats !! i = r (Suc i)" by simp also have "\ = target (stake (Suc i) wr) (r 0)" unfolding 2(2) by rule also have "\ = (r 0 ## trace wr (r 0)) !! Suc i" by simp also have "\ = smap snd wr !! i" unfolding trace_alt_def by simp finally show "smap (r \ Suc) nats !! i = smap snd wr !! i" by this qed show ?thesis proof show "run A (smap fst wr ||| smap (r \ Suc) nats) (r 0)" using 2(1) unfolding 3 by auto qed qed lemma nba_g_run_ipath: assumes "run A (w ||| r) p" shows "ipath (g_E (nba_g A)) (snth (p ## r))" proof fix i - have 1: "w !! i \ alphabet A" "r !! i \ succ A (w !! i) (target (stake i (w ||| r)) p)" + have 1: "w !! i \ alphabet A" "r !! i \ transition A (w !! i) (target (stake i (w ||| r)) p)" using assms by (auto dest: nba.run_snth) have 2: "r !! i \ successors A ((p ## r) !! i)" using 1 unfolding sscan_scons_snth[symmetric] trace_alt_def by auto show "((p ## r) !! i, (p ## r) !! Suc i) \ g_E (nba_g A)" using 2 unfolding nba_g_def graph_rec.simps E_of_succ_def by simp qed subsection \Indexed Generalized Büchi Graphs\ definition nba_igbg :: "('label, 'state) nba \ 'state igb_graph_rec" where "nba_igbg A \ graph_rec.extend (nba_g A) \ igbg_num_acc = 1, igbg_acc = \ p. if accepting A p then {0} else {} \" lemma acc_run_language: assumes "igb_graph (nba_igbg A)" shows "Ex (igb_graph.is_acc_run (nba_igbg A)) \ language A \ {}" proof interpret igb_graph "nba_igbg A" using assms by this have [simp]: "V0 = g_V0 (nba_g A)" "E = g_E (nba_g A)" "num_acc = 1" "0 \ acc p \ accepting A p" for p unfolding nba_igbg_def graph_rec.defs by simp+ show "language A \ {}" if run: "Ex is_acc_run" proof - obtain r where 1: "is_acc_run r" using run by rule have 2: "r 0 \ V0" "ipath E r" "is_acc r" using 1 unfolding is_acc_run_def graph_defs.is_run_def by auto obtain w where 3: "run A (w ||| smap (r \ Suc) nats) (r 0)" using nba_g_ipath_run 2(2) by auto have 4: "r 0 ## trace (w ||| smap (r \ Suc) nats) (r 0) = smap r nats" unfolding trace_alt_def by (simp) (metis stream.map_comp smap_siterate) have 5: "infs (accepting A) (r 0 ## trace (w ||| smap (r \ Suc) nats) (r 0))" using 2(3) unfolding infs_infm is_acc_def 4 by simp have "w \ language A" proof show "r 0 \ initial A" using nba_g_V0 2(1) by force show "run A (w ||| smap (r \ Suc) nats) (r 0)" using 3 by this show "infs (accepting A) (trace (w ||| smap (r \ Suc) nats) (r 0))" using 5 by simp qed then show ?thesis by auto qed show "Ex is_acc_run" if language: "language A \ {}" proof - obtain w where 1: "w \ language A" using language by auto obtain r p where 2: "p \ initial A" "run A (w ||| r) p" "infs (accepting A) (trace (w ||| r) p)" using 1 by rule have 3: "infs (accepting A) (p ## trace (w ||| r) p)" using 2(3) by simp have "is_acc_run (snth (p ## r))" unfolding is_acc_run_def graph_defs.is_run_def proof safe show "(p ## r) !! 0 \ V0" using nba_g_V0 2(1) by force show "ipath E (snth (p ## r))" using nba_g_run_ipath 2(2) by force show "is_acc (snth (p ## r))" using 3 unfolding infs_infm is_acc_def trace_alt_def by simp qed then show ?thesis by auto qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Implement.thy @@ -1,82 +1,82 @@ section \Implementation of Nondeterministic Büchi Automata\ theory NBA_Implement imports "NBA_Refine" "../../Basic/Implement" begin datatype ('label, 'state) nbai = nbai (alphabeti: "'label list") (initiali: "'state list") - (succi: "'label \ 'state \ 'state list") + (transitioni: "'label \ 'state \ 'state list") (acceptingi: "'state \ bool") definition nbai_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) nbai \ ('label\<^sub>2, 'state\<^sub>2) nbai) set" where [to_relAPP]: "nbai_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabeti A\<^sub>1, alphabeti A\<^sub>2) \ \L\ list_rel \ (initiali A\<^sub>1, initiali A\<^sub>2) \ \S\ list_rel \ - (succi A\<^sub>1, succi A\<^sub>2) \ L \ S \ \S\ list_rel \ + (transitioni A\<^sub>1, transitioni A\<^sub>2) \ L \ S \ \S\ list_rel \ (acceptingi A\<^sub>1, acceptingi A\<^sub>2) \ S \ bool_rel}" lemma nbai_param[param]: "(nbai, nbai) \ \L\ list_rel \ \S\ list_rel \ (L \ S \ \S\ list_rel) \ (S \ bool_rel) \ \L, S\ nbai_rel" "(alphabeti, alphabeti) \ \L, S\ nbai_rel \ \L\ list_rel" "(initiali, initiali) \ \L, S\ nbai_rel \ \S\ list_rel" - "(succi, succi) \ \L, S\ nbai_rel \ L \ S \ \S\ list_rel" + "(transitioni, transitioni) \ \L, S\ nbai_rel \ L \ S \ \S\ list_rel" "(acceptingi, acceptingi) \ \L, S\ nbai_rel \ (S \ bool_rel)" unfolding nbai_rel_def fun_rel_def by auto definition nbai_nba_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) nbai \ ('label\<^sub>2, 'state\<^sub>2) nba) set" where [to_relAPP]: "nbai_nba_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabeti A\<^sub>1, alphabet A\<^sub>2) \ \L\ list_set_rel \ (initiali A\<^sub>1, initial A\<^sub>2) \ \S\ list_set_rel \ - (succi A\<^sub>1, succ A\<^sub>2) \ L \ S \ \S\ list_set_rel \ + (transitioni A\<^sub>1, transition A\<^sub>2) \ L \ S \ \S\ list_set_rel \ (acceptingi A\<^sub>1, accepting A\<^sub>2) \ S \ bool_rel}" lemma nbai_nba_param[param, autoref_rules]: "(nbai, nba) \ \L\ list_set_rel \ \S\ list_set_rel \ (L \ S \ \S\ list_set_rel) \ (S \ bool_rel) \ \L, S\ nbai_nba_rel" "(alphabeti, alphabet) \ \L, S\ nbai_nba_rel \ \L\ list_set_rel" "(initiali, initial) \ \L, S\ nbai_nba_rel \ \S\ list_set_rel" - "(succi, succ) \ \L, S\ nbai_nba_rel \ L \ S \ \S\ list_set_rel" + "(transitioni, transition) \ \L, S\ nbai_nba_rel \ L \ S \ \S\ list_set_rel" "(acceptingi, accepting) \ \L, S\ nbai_nba_rel \ S \ bool_rel" unfolding nbai_nba_rel_def fun_rel_def by auto definition nbai_nba :: "('label, 'state) nbai \ ('label, 'state) nba" where - "nbai_nba A \ nba (set (alphabeti A)) (set (initiali A)) (\ a p. set (succi A a p)) (acceptingi A)" + "nbai_nba A \ nba (set (alphabeti A)) (set (initiali A)) (\ a p. set (transitioni A a p)) (acceptingi A)" definition nbai_invar :: "('label, 'state) nbai \ bool" where - "nbai_invar A \ distinct (alphabeti A) \ distinct (initiali A) \ (\ a p. distinct (succi A a p))" + "nbai_invar A \ distinct (alphabeti A) \ distinct (initiali A) \ (\ a p. distinct (transitioni A a p))" lemma nbai_nba_id_param[param]: "(nbai_nba, id) \ \L, S\ nbai_nba_rel \ \L, S\ nba_rel" proof fix Ai A assume 1: "(Ai, A) \ \L, S\ nbai_nba_rel" have 2: "nbai_nba Ai = nba (set (alphabeti Ai)) (set (initiali Ai)) - (\ a p. set (succi Ai a p)) (acceptingi Ai)" unfolding nbai_nba_def by rule + (\ a p. set (transitioni Ai a p)) (acceptingi Ai)" unfolding nbai_nba_def by rule have 3: "id A = nba (id (alphabet A)) (id (initial A)) - (\ a p. id (succ A a p)) (accepting A)" by simp + (\ a p. id (transition A a p)) (accepting A)" by simp show "(nbai_nba Ai, id A) \ \L, S\ nba_rel" unfolding 2 3 using 1 by parametricity qed lemma nbai_nba_br: "\Id, Id\ nbai_nba_rel = br nbai_nba nbai_invar" proof safe show "(A, B) \ \Id, Id\ nbai_nba_rel" if "(A, B) \ br nbai_nba nbai_invar" for A and B :: "('a, 'b) nba" using that unfolding nbai_nba_rel_def nbai_nba_def nbai_invar_def by (auto simp: in_br_conv list_set_rel_def) show "(A, B) \ br nbai_nba nbai_invar" if "(A, B) \ \Id, Id\ nbai_nba_rel" for A and B :: "('a, 'b) nba" proof - have 1: "(nbai_nba A, id B) \ \Id, Id\ nba_rel" using that by parametricity have 2: "nbai_invar A" using nbai_nba_param(2 - 5)[param_fo, OF that] by (auto simp: in_br_conv list_set_rel_def nbai_invar_def) show ?thesis using 1 2 unfolding in_br_conv by auto qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Refine.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Refine.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Refine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Refine.thy @@ -1,81 +1,81 @@ section \Relations on Nondeterministic Büchi Automata\ theory NBA_Refine imports "NBA" "../../Transition_Systems/Transition_System_Refine" begin definition nba_rel :: "('label\<^sub>1 \ 'label\<^sub>2) set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ (('label\<^sub>1, 'state\<^sub>1) nba \ ('label\<^sub>2, 'state\<^sub>2) nba) set" where [to_relAPP]: "nba_rel L S \ {(A\<^sub>1, A\<^sub>2). (alphabet A\<^sub>1, alphabet A\<^sub>2) \ \L\ set_rel \ (initial A\<^sub>1, initial A\<^sub>2) \ \S\ set_rel \ - (succ A\<^sub>1, succ A\<^sub>2) \ L \ S \ \S\ set_rel \ + (transition A\<^sub>1, transition A\<^sub>2) \ L \ S \ \S\ set_rel \ (accepting A\<^sub>1, accepting A\<^sub>2) \ S \ bool_rel}" lemma nba_param[param]: "(nba, nba) \ \L\ set_rel \ \S\ set_rel \ (L \ S \ \S\ set_rel) \ (S \ bool_rel) \ \L, S\ nba_rel" "(alphabet, alphabet) \ \L, S\ nba_rel \ \L\ set_rel" "(initial, initial) \ \L, S\ nba_rel \ \S\ set_rel" - "(succ, succ) \ \L, S\ nba_rel \ L \ S \ \S\ set_rel" + "(transition, transition) \ \L, S\ nba_rel \ L \ S \ \S\ set_rel" "(accepting, accepting) \ \L, S\ nba_rel \ S \ bool_rel" unfolding nba_rel_def fun_rel_def by auto lemma nba_rel_id[simp]: "\Id, Id\ nba_rel = Id" unfolding nba_rel_def using nba.expand by auto lemma nba_rel_comp[trans]: assumes [param]: "(A, B) \ \L\<^sub>1, S\<^sub>1\ nba_rel" "(B, C) \ \L\<^sub>2, S\<^sub>2\ nba_rel" shows "(A, C) \ \L\<^sub>1 O L\<^sub>2, S\<^sub>1 O S\<^sub>2\ nba_rel" proof - have "(accepting A, accepting B) \ S\<^sub>1 \ bool_rel" by parametricity also have "(accepting B, accepting C) \ S\<^sub>2 \ bool_rel" by parametricity finally have 1: "(accepting A, accepting C) \ S\<^sub>1 O S\<^sub>2 \ bool_rel" by simp - have "(succ A, succ B) \ L\<^sub>1 \ S\<^sub>1 \ \S\<^sub>1\ set_rel" by parametricity - also have "(succ B, succ C) \ L\<^sub>2 \ S\<^sub>2 \ \S\<^sub>2\ set_rel" by parametricity - finally have 2: "(succ A, succ C) \ L\<^sub>1 O L\<^sub>2 \ S\<^sub>1 O S\<^sub>2 \ \S\<^sub>1\ set_rel O \S\<^sub>2\ set_rel" by simp + have "(transition A, transition B) \ L\<^sub>1 \ S\<^sub>1 \ \S\<^sub>1\ set_rel" by parametricity + also have "(transition B, transition C) \ L\<^sub>2 \ S\<^sub>2 \ \S\<^sub>2\ set_rel" by parametricity + finally have 2: "(transition A, transition C) \ L\<^sub>1 O L\<^sub>2 \ S\<^sub>1 O S\<^sub>2 \ \S\<^sub>1\ set_rel O \S\<^sub>2\ set_rel" by simp show ?thesis unfolding nba_rel_def mem_Collect_eq prod.case set_rel_compp using 1 2 using nba_param(2 - 5)[THEN fun_relD, OF assms(1)] using nba_param(2 - 5)[THEN fun_relD, OF assms(2)] by auto qed lemma nba_rel_converse[simp]: "(\L, S\ nba_rel)\ = \L\, S\\ nba_rel" proof - have 1: "\L\ set_rel = (\L\\ set_rel)\" by simp have 2: "\S\ set_rel = (\S\\ set_rel)\" by simp have 3: "L \ S \ \S\ set_rel = (L\ \ S\ \ \S\\ set_rel)\" by simp have 4: "S \ bool_rel = (S\ \ bool_rel)\" by simp show ?thesis unfolding nba_rel_def unfolding 3 unfolding 1 2 4 by fastforce qed lemma nba_rel_eq: "(A, A) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" unfolding nba_rel_def by auto lemma enableds_param[param]: "(enableds, enableds) \ \L, S\ nba_rel \ S \ \L \\<^sub>r S\ set_rel" using nba_param(2, 4) unfolding nba.enableds_def fun_rel_def set_rel_def by fastforce lemma paths_param[param]: "(paths, paths) \ \L, S\ nba_rel \ S \ \\L \\<^sub>r S\ list_rel\ set_rel" unfolding paths_def by (intro fun_relI paths_param, fold enableds_def) (parametricity+) lemma runs_param[param]: "(runs, runs) \ \L, S\ nba_rel \ S \ \\L \\<^sub>r S\ stream_rel\ set_rel" unfolding runs_def by (intro fun_relI runs_param, fold enableds_def) (parametricity+) lemma reachable_param[param]: "(reachable, reachable) \ \L, S\ nba_rel \ S \ \S\ set_rel" proof - have 1: "reachable A p = (\ wr. target wr p) ` paths A p" for A :: "('label, 'state) nba" and p unfolding nba.reachable_alt_def nba.paths_def by auto show ?thesis unfolding 1 by parametricity qed lemma nodes_param[param]: "(nodes, nodes) \ \L, S\ nba_rel \ \S\ set_rel" unfolding nba.nodes_alt_def Collect_mem_eq by parametricity lemma language_param[param]: "(language, language) \ \L, S\ nba_rel \ \\L\ stream_rel\ set_rel" proof - have 1: "language A = (\ p \ initial A. \ wr \ runs A p. if infs (accepting A) (trace wr p) then {smap fst wr} else {})" for A :: "('label, 'state) nba" unfolding language_def nba.runs_def image_def by (auto iff: split_szip_ex) show ?thesis unfolding 1 by parametricity qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Translate.thy @@ -1,253 +1,253 @@ section \Explore and Enumerate Nodes of Nondeterministic Büchi Automata\ theory NBA_Translate imports NBA_Explicit begin subsection \Syntax\ (* TODO: this syntax has unnecessarily high inner binding strength, requiring extra parentheses the regular let syntax correctly uses inner binding strength 0: ("(2_ =/ _)" 10) *) no_syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" [1000, 13] 13) syntax "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" 13) section \Image on Explicit Automata\ definition nbae_image where "nbae_image f A \ nbae (alphabete A) (f ` initiale A) - ((\ (p, a, q). (f p, a, f q)) ` transe A) (f ` acceptinge A)" + ((\ (p, a, q). (f p, a, f q)) ` transitione A) (f ` acceptinge A)" lemma nbae_image_param[param]: "(nbae_image, nbae_image) \ (S \ T) \ \L, S\ nbae_rel \ \L, T\ nbae_rel" unfolding nbae_image_def by parametricity lemma nbae_image_id[simp]: "nbae_image id = id" unfolding nbae_image_def by auto lemma nbae_image_nba_nbae: "nbae_image f (nba_nbae A) = nbae (alphabet A) (f ` initial A) - (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) + (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) (f ` {p \ nodes A. accepting A p})" unfolding nba_nbae_def nbae_image_def nbae.simps Set.filter_def by force section \Exploration and Translation\ definition trans_spec where - "trans_spec A f \ \ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p" + "trans_spec A f \ \ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p" definition trans_algo where "trans_algo N L S f \ FOREACH N (\ p T. do { ASSERT (p \ N); FOREACH L (\ a T. do { ASSERT (a \ L); FOREACH (S a p) (\ q T. do { ASSERT (q \ S a p); ASSERT ((f p, a, f q) \ T); RETURN (insert (f p, a, f q) T) } ) T } ) T } ) {}" lemma trans_algo_refine: assumes "finite (nodes A)" "finite (alphabet A)" "inj_on f (nodes A)" - assumes "N = nodes A" "L = alphabet A" "S = succ A" + assumes "N = nodes A" "L = alphabet A" "S = transition A" shows "(trans_algo N L S f, SPEC (HOL.eq (trans_spec A f))) \ \Id\ nres_rel" unfolding trans_algo_def trans_spec_def assms(4-6) proof (refine_vcg FOREACH_rule_insert_eq) show "finite (nodes A)" using assms(1) by this - show "(\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) = - (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p)" by rule - show "(\ p \ {}. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) = {}" by simp + show "(\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = + (\ p \ nodes A. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by rule + show "(\ p \ {}. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = {}" by simp fix T x assume 1: "T \ nodes A" "x \ nodes A" "x \ T" show "finite (alphabet A)" using assms(2) by this - show "(\ a \ {}. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) = - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p)" - "(\ a \ alphabet A. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) = - (\ p \ insert x T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p)" by auto + show "(\ a \ {}. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" + "(\ a \ alphabet A. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = + (\ p \ insert x T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by auto fix Ta xa assume 2: "Ta \ alphabet A" "xa \ alphabet A" "xa \ Ta" - show "finite (succ A xa x)" using 1 2 assms(1) by (meson infinite_subset nodes_succ subsetI) - show "(f ` {x} \ {xa} \ f ` succ A xa x) \ - (\ a \ Ta. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) = - (\ a \ insert xa Ta. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p)" + show "finite (transition A xa x)" using 1 2 assms(1) by (meson infinite_subset nodes_succ subsetI) + show "(f ` {x} \ {xa} \ f ` transition A xa x) \ + (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = + (\ a \ insert xa Ta. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by auto show "(f ` {x} \ {xa} \ f ` {}) \ - (\ a \ Ta. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) = - (\ a \ Ta. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p)" + (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = + (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" by auto fix Tb xb - assume 3: "Tb \ succ A xa x" "xb \ succ A xa x" "xb \ Tb" + assume 3: "Tb \ transition A xa x" "xb \ transition A xa x" "xb \ Tb" show "(f x, xa, f xb) \ f ` {x} \ {xa} \ f ` Tb \ - (\ a \ Ta. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p)" + (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p)" using 1 2 3 assms(3) by (blast dest: inj_onD) show "f ` {x} \ {xa} \ f ` insert xb Tb \ - (\ a \ Ta. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p) = + (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p) = insert (f x, xa, f xb) (f ` {x} \ {xa} \ f ` Tb \ - (\ a \ Ta. f ` {x} \ {a} \ f ` succ A a x) \ - (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` succ A a p))" + (\ a \ Ta. f ` {x} \ {a} \ f ` transition A a x) \ + (\ p \ T. \ a \ alphabet A. f ` {p} \ {a} \ f ` transition A a p))" by auto qed definition to_nbaei :: "('state, 'label) nba \ ('state, 'label) nba" where "to_nbaei \ id" (* TODO: make separate implementations for "nba_nbae" and "op_set_enumerate \ nbae_image" *) schematic_goal to_nbaei_impl: fixes S :: "('statei \ 'state) set" assumes [simp]: "finite (nodes A)" assumes [autoref_ga_rules]: "is_bounded_hashcode S seq bhc" assumes [autoref_ga_rules]: "is_valid_def_hm_size TYPE('statei) hms" assumes [autoref_rules]: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes [autoref_rules]: "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(?f :: ?'a, do { let N = nodes A; f \ op_set_enumerate N; ASSERT (dom f = N); ASSERT (\ p \ initial A. f p \ None); - ASSERT (\ a \ alphabet A. \ p \ dom f. \ q \ succ A a p. f q \ None); - T \ trans_algo N (alphabet A) (succ A) (\ x. the (f x)); + ASSERT (\ a \ alphabet A. \ p \ dom f. \ q \ transition A a p. f q \ None); + T \ trans_algo N (alphabet A) (transition A) (\ x. the (f x)); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ N. accepting A p})) }) \ ?R" unfolding trans_algo_def by (autoref_monadic (plain)) concrete_definition to_nbaei_impl uses to_nbaei_impl lemma to_nbaei_impl_refine'': fixes S :: "('statei \ 'state) set" assumes "finite (nodes A)" assumes "is_bounded_hashcode S seq bhc" assumes "is_valid_def_hm_size TYPE('statei) hms" assumes "(seq, HOL.eq) \ S \ S \ bool_rel" assumes "(Ai, A) \ \L, S\ nbai_nba_rel" shows "(RETURN (to_nbaei_impl seq bhc hms Ai), do { f \ op_set_enumerate (nodes A); RETURN (nbae_image (the \ f) (nba_nbae A)) }) \ \\L, nat_rel\ nbaei_nbae_rel\ nres_rel" proof - have 1: "finite (alphabet A)" using nbai_nba_param(2)[param_fo, OF assms(5)] list_set_rel_finite unfolding finite_set_rel_def by auto note to_nbaei_impl.refine[OF assms] also have "(do { let N = nodes A; f \ op_set_enumerate N; ASSERT (dom f = N); ASSERT (\ p \ initial A. f p \ None); - ASSERT (\ a \ alphabet A. \ p \ dom f. \ q \ succ A a p. f q \ None); - T \ trans_algo N (alphabet A) (succ A) (\ x. the (f x)); + ASSERT (\ a \ alphabet A. \ p \ dom f. \ q \ transition A a p. f q \ None); + T \ trans_algo N (alphabet A) (transition A) (\ x. the (f x)); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ N. accepting A p})) }, do { f \ op_set_enumerate (nodes A); T \ SPEC (HOL.eq (trans_spec A (\ x. the (f x)))); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ nodes A. accepting A p})) }) \ \Id\ nres_rel" unfolding Let_def comp_apply op_set_enumerate_def using assms(1) 1 by (refine_vcg vcg0[OF trans_algo_refine]) (auto intro!: inj_on_map_the[unfolded comp_apply]) also have "(do { f \ op_set_enumerate (nodes A); T \ SPEC (HOL.eq (trans_spec A (\ x. the (f x)))); RETURN (nbae (alphabet A) ((\ x. the (f x)) ` initial A) T ((\ x. the (f x)) ` {p \ nodes A. accepting A p})) }, do { f \ op_set_enumerate (nodes A); RETURN (nbae_image (the \ f) (nba_nbae A)) }) \ \Id\ nres_rel" unfolding trans_spec_def nbae_image_nba_nbae by refine_vcg force finally show ?thesis unfolding nres_rel_comp by simp qed (* TODO: generalize L *) context fixes Ai A fixes seq bhc hms fixes S :: "('statei \ 'state) set" assumes a: "finite (nodes A)" assumes b: "is_bounded_hashcode S seq bhc" assumes c: "is_valid_def_hm_size TYPE('statei) hms" assumes d: "(seq, HOL.eq) \ S \ S \ bool_rel" assumes e: "(Ai, A) \ \Id, S\ nbai_nba_rel" begin definition f' where "f' \ SOME f'. (to_nbaei_impl seq bhc hms Ai, nbae_image (the \ f') (nba_nbae A)) \ \Id, nat_rel\ nbaei_nbae_rel \ dom f' = nodes A \ inj_on f' (nodes A)" lemma 1: "\ f'. (to_nbaei_impl seq bhc hms Ai, nbae_image (the \ f') (nba_nbae A)) \ \Id, nat_rel\ nbaei_nbae_rel \ dom f' = nodes A \ inj_on f' (nodes A)" using to_nbaei_impl_refine''[ OF a b c d e, unfolded op_set_enumerate_def bind_RES_RETURN_eq, THEN nres_relD, THEN RETURN_ref_SPECD] by force lemma f'_refine: "(to_nbaei_impl seq bhc hms Ai, nbae_image (the \ f') (nba_nbae A)) \ \Id, nat_rel\ nbaei_nbae_rel" using someI_ex[OF 1, folded f'_def] by auto lemma f'_dom: "dom f' = nodes A" using someI_ex[OF 1, folded f'_def] by auto lemma f'_inj: "inj_on f' (nodes A)" using someI_ex[OF 1, folded f'_def] by auto definition f where "f \ the \ f'" definition g where "g = inv_into (nodes A) f" lemma inj_f[intro!, simp]: "inj_on f (nodes A)" using f'_inj f'_dom unfolding f_def by (simp add: inj_on_map_the) lemma inj_g[intro!, simp]: "inj_on g (f ` nodes A)" unfolding g_def by (simp add: inj_on_inv_into) definition rel where "rel \ {(f p, p) |p. p \ nodes A}" lemma rel_alt_def: "rel = (br f (\ p. p \ nodes A))\" unfolding rel_def by (auto simp: in_br_conv) lemma rel_inv_def: "rel = br g (\ k. k \ f ` nodes A)" unfolding rel_alt_def g_def by (auto simp: in_br_conv) lemma rel_domain[simp]: "Domain rel = f ` nodes A" unfolding rel_def by force lemma rel_range[simp]: "Range rel = nodes A" unfolding rel_def by auto lemma [intro!, simp]: "bijective rel" unfolding rel_inv_def by (simp add: bijective_alt) lemma [simp]: "Id_on (f ` nodes A) O rel = rel" unfolding rel_def by auto lemma [simp]: "rel O Id_on (nodes A) = rel" unfolding rel_def by auto lemma [param]: "(f, f) \ Id_on (Range rel) \ Id_on (Domain rel)" unfolding rel_alt_def by auto lemma [param]: "(g, g) \ Id_on (Domain rel) \ Id_on (Range rel)" unfolding rel_inv_def by auto lemma [param]: "(id, f) \ rel \ Id_on (Domain rel)" unfolding rel_alt_def by (auto simp: in_br_conv) lemma [param]: "(f, id) \ Id_on (Range rel) \ rel" unfolding rel_alt_def by (auto simp: in_br_conv) lemma [param]: "(id, g) \ Id_on (Domain rel) \ rel" unfolding rel_inv_def by (auto simp: in_br_conv) lemma [param]: "(g, id) \ rel \ Id_on (Range rel)" unfolding rel_inv_def by (auto simp: in_br_conv) lemma to_nbaei_impl_refine': "(to_nbaei_impl seq bhc hms Ai, to_nbaei A) \ \Id_on (alphabet A), rel\ nbaei_nba_rel" proof - have "(nbae_nba (nbaei_nbae (to_nbaei_impl seq bhc hms Ai)), nbae_nba (id (nbae_image f (nba_nbae A)))) \ \Id, nat_rel\ nba_rel" using f'_refine[folded f_def] by parametricity auto also have "(nbae_nba (id (nbae_image f (nba_nbae A))), nbae_nba (id (nbae_image id (nba_nbae A)))) \ \Id_on (alphabet A), rel\ nba_rel" using nba_rel_eq by parametricity auto also have "nbae_nba (id (nbae_image id (nba_nbae A))) = (nbae_nba \ nba_nbae) A" by simp also have "(\, id A) \ \Id_on (alphabet A), Id_on (nodes A)\ nba_rel" by parametricity also have "id A = to_nbaei A" unfolding to_nbaei_def by simp finally show ?thesis unfolding nbaei_nba_rel_def by simp qed end context begin interpretation autoref_syn by this lemma to_nbaei_impl_refine[autoref_rules]: fixes S :: "('statei \ 'state) set" assumes "SIDE_PRECOND (finite (nodes A))" assumes "SIDE_GEN_ALGO (is_bounded_hashcode S seq bhc)" assumes "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('statei) hms)" assumes "GEN_OP seq HOL.eq (S \ S \ bool_rel)" assumes "(Ai, A) \ \Id, S\ nbai_nba_rel" shows "(to_nbaei_impl seq bhc hms Ai, (OP to_nbaei ::: \Id, S\ nbai_nba_rel \ \Id_on (alphabet A), rel Ai A seq bhc hms\ nbaei_nba_rel) $ A) \ \Id_on (alphabet A), rel Ai A seq bhc hms\ nbaei_nba_rel" using to_nbaei_impl_refine' assms unfolding autoref_tag_defs by this end end \ No newline at end of file