diff --git a/thys/Buchi_Complementation/Ranking.thy b/thys/Buchi_Complementation/Ranking.thy --- a/thys/Buchi_Complementation/Ranking.thy +++ b/thys/Buchi_Complementation/Ranking.thy @@ -1,481 +1,481 @@ section \Rankings\ theory Ranking imports "Alternate" "Graph" begin subsection \Rankings\ type_synonym 'state ranking = "'state node \ nat" definition ranking :: "('label, 'state) nba \ 'label stream \ 'state ranking \ bool" where "ranking A w f \ (\ v \ gunodes A w. f v \ 2 * card (nodes A)) \ (\ v \ gunodes A w. \ u \ gusuccessors A w v. f u \ f v) \ (\ v \ gunodes A w. gaccepting A v \ even (f v)) \ (\ v \ gunodes A w. \ r k. gurun A w r v \ smap f (gtrace r v) = sconst k \ odd k)" subsection \Ranking Implies Word not in Language\ lemma ranking_stuck: assumes "ranking A w f" assumes "v \ gunodes A w" "gurun A w r v" obtains n k where "smap f (gtrace (sdrop n r) (gtarget (stake n r) v)) = sconst k" proof - have 0: "f u \ f v" if "v \ gunodes A w" "u \ gusuccessors A w v" for v u using assms(1) that unfolding ranking_def by auto have 1: "shd (v ## gtrace r v) \ gunodes A w" using assms(2) by auto have 2: "sdescending (smap f (v ## gtrace r v))" using 1 assms(3) proof (coinduction arbitrary: r v rule: sdescending.coinduct) case sdescending obtain u s where 1: "r = u ## s" using stream.exhaust by blast have 2: "v \ gunodes A w" using sdescending(1) by simp have 3: "gurun A w (u ## s) v" using sdescending(2) 1 by auto have 4: "u \ gusuccessors A w v" using 3 by auto have 5: "u \ gureachable A w v" using graph.reachable_successors 4 by blast show ?case unfolding 1 proof (intro exI conjI disjI1) show "f u \ f v" using 0 2 4 by this show "shd (u ## gtrace s u) \ gunodes A w" using 2 5 by auto show "gurun A w s u" using 3 by auto qed auto qed obtain n k where 3: "sdrop n (smap f (v ## gtrace r v)) = sconst k" using sdescending_stuck[OF 2] by metis have "gtrace (sdrop (Suc n) r) (gtarget (stake (Suc n) r) v) = sdrop (Suc n) (gtrace r v)" using sscan_sdrop by rule also have "smap f \ = sdrop n (smap f (v ## gtrace r v))" by (auto, metis 3 id_apply sdrop_smap sdrop_stl siterate.simps(2) sscan_const stream.map stream.map_sel(2) stream.sel(2)) also have "\ = sconst k" using 3 by this finally show ?thesis using that by blast qed lemma ranking_stuck_odd: assumes "ranking A w f" assumes "v \ gunodes A w" "gurun A w r v" obtains n where "sset (smap f (gtrace (sdrop n r) (gtarget (stake n r) v))) \ Collect odd" proof - obtain n k where 1: "smap f (gtrace (sdrop n r) (gtarget (stake n r) v)) = sconst k" using ranking_stuck assms by this have 2: "gtarget (stake n r) v \ gunodes A w" using assms(2, 3) by (simp add: graph.nodes_target graph.run_stake) have 3: "gurun A w (sdrop n r) (gtarget (stake n r) v)" using assms(2, 3) by (simp add: graph.run_sdrop) have 4: "odd k" using 1 2 3 assms(1) unfolding ranking_def by meson have "sset (smap f (gtrace (sdrop n r) (gtarget (stake n r) v))) = sset (sconst k)" unfolding 1 by rule also have "\ \ Collect odd" using 4 by simp finally show ?thesis using that by simp qed lemma ranking_language: assumes "ranking A w f" shows "w \ language A" proof assume 1: "w \ language A" obtain r p where 2: "run A (w ||| r) p" "p \ initial A" "infs (accepting A) (trace (w ||| r) p)" using 1 by rule let ?r = "fromN 1 ||| r" let ?v = "(0, p)" have 3: "?v \ gunodes A w" "gurun A w ?r ?v" using 2(1, 2) by (auto intro: run_grun) obtain n where 4: "sset (smap f (gtrace (sdrop n ?r) (gtarget (stake n ?r) ?v))) \ {a. odd a}" using ranking_stuck_odd assms 3 by this let ?s = "stake n ?r" let ?t = "sdrop n ?r" let ?u = "gtarget ?s ?v" have "sset (gtrace ?t ?u) \ gureachable A w ?v" proof (intro graph.reachable_trace graph.reachable_target graph.reachable.reflexive) show "gupath A w ?s ?v" using graph.run_stake 3(2) by this show "gurun A w ?t ?u" using graph.run_sdrop 3(2) by this qed also have "\ \ gunodes A w" using 3(1) by blast finally have 7: "sset (gtrace ?t ?u) \ gunodes A w" by this have 8: "\ p. p \ gunodes A w \ gaccepting A p \ even (f p)" using assms unfolding ranking_def by auto have 9: "\ p. p \ sset (gtrace ?t ?u) \ gaccepting A p \ even (f p)" using 7 8 by auto have 19: "infs (accepting A) (smap snd ?r)" using 2(3) unfolding trace_alt_def by simp - have 18: "infs (gaccepting A) ?r" using 19 by (force simp: comp_def) + have 18: "infs (gaccepting A) ?r" using 19 by simp have 17: "infs (gaccepting A) (gtrace ?r ?v)" using 18 unfolding gtrace_alt_def by this have 16: "infs (gaccepting A) (gtrace (?s @- ?t) ?v)" using 17 unfolding stake_sdrop by this have 15: "infs (gaccepting A) (gtrace ?t ?u)" using 16 by simp have 13: "infs (even \ f) (gtrace ?t ?u)" using infs_mono[OF _ 15] 9 by simp have 12: "infs even (smap f (gtrace ?t ?u))" using 13 by (simp add: comp_def) have 11: "Bex (sset (smap f (gtrace ?t ?u))) even" using 12 infs_any by metis show False using 4 11 by auto qed subsection \Word not in Language implies Ranking\ subsubsection \Removal of Endangered Nodes\ definition clean :: "('label, 'state) nba \ 'label stream \ 'state node set \ 'state node set" where "clean A w V \ {v \ V. infinite (greachable A w V v)}" lemma clean_decreasing: "clean A w V \ V" unfolding clean_def by auto lemma clean_successors: assumes "v \ V" "u \ gusuccessors A w v" shows "u \ clean A w V \ v \ clean A w V" proof - assume 1: "u \ clean A w V" have 2: "u \ V" "infinite (greachable A w V u)" using 1 unfolding clean_def by auto have 3: "u \ greachable A w V v" using graph.reachable.execute assms(2) 2(1) by blast have 4: "greachable A w V u \ greachable A w V v" using 3 by blast have 5: "infinite (greachable A w V v)" using 2(2) 4 by (simp add: infinite_super) show "v \ clean A w V" unfolding clean_def using assms(1) 5 by simp qed subsubsection \Removal of Safe Nodes\ definition prune :: "('label, 'state) nba \ 'label stream \ 'state node set \ 'state node set" where "prune A w V \ {v \ V. \ u \ greachable A w V v. gaccepting A u}" lemma prune_decreasing: "prune A w V \ V" unfolding prune_def by auto lemma prune_successors: assumes "v \ V" "u \ gusuccessors A w v" shows "u \ prune A w V \ v \ prune A w V" proof - assume 1: "u \ prune A w V" have 2: "u \ V" "\ x \ greachable A w V u. gaccepting A x" using 1 unfolding prune_def by auto have 3: "u \ greachable A w V v" using graph.reachable.execute assms(2) 2(1) by blast have 4: "greachable A w V u \ greachable A w V v" using 3 by blast show "v \ prune A w V" unfolding prune_def using assms(1) 2(2) 4 by auto qed subsubsection \Run Graph Interation\ definition graph :: "('label, 'state) nba \ 'label stream \ nat \ 'state node set" where "graph A w k \ alternate (clean A w) (prune A w) k (gunodes A w)" abbreviation "level A w k l \ {v \ graph A w k. fst v = l}" lemma graph_0[simp]: "graph A w 0 = gunodes A w" unfolding graph_def by simp lemma graph_Suc[simp]: "graph A w (Suc k) = (if even k then clean A w else prune A w) (graph A w k)" unfolding graph_def by simp lemma graph_antimono: "antimono (graph A w)" using alternate_antimono clean_decreasing prune_decreasing unfolding antimono_def le_fun_def graph_def by metis lemma graph_nodes: "graph A w k \ gunodes A w" using graph_0 graph_antimono le0 antimonoD by metis lemma graph_successors: assumes "v \ gunodes A w" "u \ gusuccessors A w v" shows "u \ graph A w k \ v \ graph A w k" using assms proof (induct k arbitrary: u v) case 0 show ?case using 0(2) by simp next case (Suc k) have 1: "v \ graph A w k" using Suc using antimono_iff_le_Suc graph_antimono rev_subsetD by blast show ?case using Suc(2) clean_successors[OF 1 Suc(4)] prune_successors[OF 1 Suc(4)] by auto qed lemma graph_level_finite: assumes "finite (nodes A)" shows "finite (level A w k l)" proof - have "level A w k l \ {v \ gunodes A w. fst v = l}" by (simp add: graph_nodes subset_CollectI) also have "{v \ gunodes A w. fst v = l} \ {l} \ nodes A" using gnodes_nodes by force also have "finite ({l} \ nodes A)" using assms(1) by simp finally show ?thesis by this qed lemma find_safe: assumes "w \ language A" assumes "V \ {}" "V \ gunodes A w" assumes "\ v. v \ V \ gsuccessors A w V v \ {}" obtains v where "v \ V" "\ u \ greachable A w V v. \ gaccepting A u" proof (rule ccontr) assume 1: "\ thesis" have 2: "\ v. v \ V \ \ u \ greachable A w V v. gaccepting A u" using that 1 by auto have 3: "\ r v. v \ initial A \ run A (w ||| r) v \ fins (accepting A) (trace (w ||| r) v)" using assms(1) by auto obtain v where 4: "v \ V" using assms(2) by force obtain x where 5: "x \ greachable A w V v" "gaccepting A x" using 2 4 by blast obtain y where 50: "gpath A w V y v" "x = gtarget y v" using 5(1) by rule obtain r where 6: "grun A w V r x" "infs (\ x. x \ V \ gaccepting A x) r" proof (rule graph.recurring_condition) show "x \ V \ gaccepting A x" using greachable_subset 4 5 by blast next fix v assume 1: "v \ V \ gaccepting A v" obtain v' where 20: "v' \ gsuccessors A w V v" using assms(4) 1 by (meson IntE equals0I) have 21: "v' \ V" using 20 by auto have 22: "\ u \ greachable A w V v'. u \ V \ gaccepting A u" using greachable_subset 2 21 by blast obtain r where 30: "gpath A w V r v'" "gtarget r v' \ V \ gaccepting A (gtarget r v')" using 22 by blast show "\ r. r \ [] \ gpath A w V r v \ gtarget r v \ V \ gaccepting A (gtarget r v)" proof (intro exI conjI) show "v' # r \ []" by simp show "gpath A w V (v' # r) v" using 20 30 by auto show "gtarget (v' # r) v \ V" using 30 by simp show "gaccepting A (gtarget (v' # r) v)" using 30 by simp qed qed auto obtain u where 100: "u \ ginitial A" "v \ gureachable A w u" using 4 assms(3) by blast have 101: "gupath A w y v" using gpath_subset 50(1) subset_UNIV by this have 102: "gurun A w r x" using grun_subset 6(1) subset_UNIV by this obtain t where 103: "gupath A w t u" "v = gtarget t u" using 100(2) by rule have 104: "gurun A w (t @- y @- r) u" using 101 102 103 50(2) by auto obtain s q where 7: "t @- y @- r = fromN (Suc 0) ||| s" "u = (0, q)" using grun_elim[OF 104] 100(1) by blast have 8: "run A (w ||| s) q" using grun_run[OF 104[unfolded 7]] by simp have 9: "q \ initial A" using 100(1) 7(2) by auto have 91: "sset (trace (w ||| s) q) \ reachable A q" using nba.reachable_trace nba.reachable.reflexive 8 by this have 10: "fins (accepting A) (trace (w ||| s) q)" using 3 9 8 by this have 11: "fins (accepting A) s" using 10 unfolding trace_alt_def by simp have 12: "infs (gaccepting A) r" using infs_mono[OF _ 6(2)] by simp have "s = smap snd (t @- y @- r)" unfolding 7(1) by simp also have "infs (accepting A) \" using 12 by (simp add: comp_def) finally have 13: "infs (accepting A) s" by this show False using 11 13 by simp qed lemma remove_run: assumes "finite (nodes A)" "w \ language A" assumes "V \ gunodes A w" "clean A w V \ {}" obtains v r where "grun A w V r v" "sset (gtrace r v) \ clean A w V" "sset (gtrace r v) \ - prune A w (clean A w V)" proof - obtain u where 1: "u \ clean A w V" "\ x \ greachable A w (clean A w V) u. \ gaccepting A x" proof (rule find_safe) show "w \ language A" using assms(2) by this show "clean A w V \ {}" using assms(4) by this show "clean A w V \ gunodes A w" using assms(3) by (meson clean_decreasing subset_iff) next fix v assume 1: "v \ clean A w V" have 2: "v \ V" using 1 clean_decreasing by blast have 3: "infinite (greachable A w V v)" using 1 clean_def by auto have "gsuccessors A w V v \ gusuccessors A w v" by auto also have "finite \" using 2 assms(1, 3) finite_nodes_gsuccessors by blast finally have 4: "finite (gsuccessors A w V v)" by this have 5: "infinite (insert v (\((greachable A w V) ` (gsuccessors A w V v))))" using graph.reachable_step 3 by metis obtain u where 6: "u \ gsuccessors A w V v" "infinite (greachable A w V u)" using 4 5 by auto have 7: "u \ clean A w V" using 6 unfolding clean_def by auto show "gsuccessors A w (clean A w V) v \ {}" using 6(1) 7 by auto qed auto have 2: "u \ V" using 1(1) unfolding clean_def by auto have 3: "infinite (greachable A w V u)" using 1(1) unfolding clean_def by simp have 4: "finite (gsuccessors A w V v)" if "v \ greachable A w V u" for v proof - have 1: "v \ V" using that greachable_subset 2 by blast have "gsuccessors A w V v \ gusuccessors A w v" by auto also have "finite \" using 1 assms(1, 3) finite_nodes_gsuccessors by blast finally show ?thesis by this qed obtain r where 5: "grun A w V r u" using graph.koenig[OF 3 4] by this have 6: "greachable A w V u \ V" using 2 greachable_subset by blast have 7: "sset (gtrace r u) \ V" using graph.reachable_trace[OF graph.reachable.reflexive 5(1)] 6 by blast have 8: "sset (gtrace r u) \ clean A w V" unfolding clean_def using 7 infinite_greachable_gtrace[OF 5(1)] by auto have 9: "sset (gtrace r u) \ greachable A w (clean A w V) u" using 5 8 by (metis graph.reachable.reflexive graph.reachable_trace grun_subset) show ?thesis proof show "grun A w V r u" using 5(1) by this show "sset (gtrace r u) \ clean A w V" using 8 by this show "sset (gtrace r u) \ - prune A w (clean A w V)" proof (intro subsetI ComplI) fix p assume 10: "p \ sset (gtrace r u)" "p \ prune A w (clean A w V)" have 20: "\ x \ greachable A w (clean A w V) p. gaccepting A x" using 10(2) unfolding prune_def by auto have 30: "greachable A w (clean A w V) p \ greachable A w (clean A w V) u" using 10(1) 9 by blast show "False" using 1(2) 20 30 by force qed qed qed lemma level_bounded: assumes "finite (nodes A)" "w \ language A" obtains n where "\ l. l \ n \ card (level A w (2 * k) l) \ card (nodes A) - k" proof (induct k arbitrary: thesis) case (0) show ?case proof (rule 0) fix l :: nat have "finite ({l} \ nodes A)" using assms(1) by simp also have "level A w 0 l \ {l} \ nodes A" using gnodes_nodes by force also (card_mono) have "card \ = card (nodes A)" using assms(1) by simp finally show "card (level A w (2 * 0) l) \ card (nodes A) - 0" by simp qed next case (Suc k) show ?case proof (cases "graph A w (Suc (2 * k)) = {}") case True have 3: "graph A w (2 * Suc k) = {}" using True prune_decreasing by simp blast show ?thesis using Suc(2) 3 by simp next case False obtain v r where 1: "grun A w (graph A w (2 * k)) r v" "sset (gtrace r v) \ graph A w (Suc (2 * k))" "sset (gtrace r v) \ - graph A w (Suc (Suc (2 * k)))" proof (rule remove_run) show "finite (nodes A)" "w \ language A" using assms by this show "clean A w (graph A w (2 * k)) \ {}" using False by simp show "graph A w (2 * k) \ gunodes A w" using graph_nodes by this qed auto obtain l q where 2: "v = (l, q)" by force obtain n where 90: "\ l. n \ l \ card (level A w (2 * k) l) \ card (nodes A) - k" using Suc(1) by blast show ?thesis proof (rule Suc(2)) fix j assume 100: "n + Suc l \ j" have 6: "graph A w (Suc (Suc (2 * k))) \ graph A w (Suc (2 * k))" using graph_antimono antimono_iff_le_Suc by blast have 101: "gtrace r v !! (j - Suc l) \ graph A w (Suc (2 * k))" using 1(2) snth_sset by auto have 102: "gtrace r v !! (j - Suc l) \ graph A w (Suc (Suc (2 * k)))" using 1(3) snth_sset by blast have 103: "gtrace r v !! (j - Suc l) \ level A w (Suc (2 * k)) j" using 1(1) 100 101 2 by (auto elim: grun_elim) have 104: "gtrace r v !! (j - Suc l) \ level A w (Suc (Suc (2 * k))) j" using 100 102 by simp have "level A w (2 * Suc k) j = level A w (Suc (Suc (2 * k))) j" by simp also have "\ \ level A w (Suc (2 * k)) j" using 103 104 6 by blast also have "\ \ level A w (2 * k) j" by (simp add: Collect_mono clean_def) finally have 105: "level A w (2 * Suc k) j \ level A w (2 * k) j" by this have "card (level A w (2 * Suc k) j) < card (level A w (2 * k) j)" using assms(1) 105 by (simp add: graph_level_finite psubset_card_mono) also have "\ \ card (nodes A) - k" using 90 100 by simp finally show "card (level A w (2 * Suc k) j) \ card (nodes A) - Suc k" by simp qed qed qed lemma graph_empty: assumes "finite (nodes A)" "w \ language A" shows "graph A w (Suc (2 * card (nodes A))) = {}" proof - obtain n where 1: "\ l. l \ n \ card (level A w (2 * card (nodes A)) l) = 0" using level_bounded[OF assms(1, 2), of "card (nodes A)"] by auto have "graph A w (2 * card (nodes A)) = (\ l \ {..< n}. level A w (2 * card (nodes A)) l) \ (\ l \ {n ..}. level A w (2 * card (nodes A)) l)" by auto also have "(\ l \ {n ..}. level A w (2 * card (nodes A)) l) = {}" using graph_level_finite assms(1) 1 by fastforce also have "finite ((\ l \ {..< n}. level A w (2 * card (nodes A)) l) \ {})" using graph_level_finite assms(1) by auto finally have 100: "finite (graph A w (2 * card (nodes A)))" by this have 101: "finite (greachable A w (graph A w (2 * card (nodes A))) v)" for v using 100 greachable_subset[of A w "graph A w (2 * card (nodes A))" v] using finite_insert infinite_super by auto show ?thesis using 101 by (simp add: clean_def) qed lemma graph_le: assumes "finite (nodes A)" "w \ language A" assumes "v \ graph A w k" shows "k \ 2 * card (nodes A)" using graph_empty graph_antimono assms by (metis (no_types, lifting) Suc_leI antimono_def basic_trans_rules(30) empty_iff not_le_imp_less) subsection \Node Ranks\ definition rank :: "('label, 'state) nba \ 'label stream \ 'state node \ nat" where "rank A w v \ GREATEST k. v \ graph A w k" lemma rank_member: assumes "finite (nodes A)" "w \ language A" "v \ gunodes A w" shows "v \ graph A w (rank A w v)" unfolding rank_def proof (rule GreatestI_nat) show "v \ graph A w 0" using assms(3) by simp show "\ k. v \ graph A w k \ k \ 2 * card (nodes A)" using graph_le assms(1, 2) by blast qed lemma rank_removed: assumes "finite (nodes A)" "w \ language A" shows "v \ graph A w (Suc (rank A w v))" proof assume "v \ graph A w (Suc (rank A w v))" then have 2: "Suc (rank A w v) \ rank A w v" unfolding rank_def using Greatest_le_nat graph_le assms by metis then show "False" by auto qed lemma rank_le: assumes "finite (nodes A)" "w \ language A" assumes "v \ gunodes A w" "u \ gusuccessors A w v" shows "rank A w u \ rank A w v" unfolding rank_def proof (rule Greatest_le_nat) have 1: "u \ gureachable A w v" using graph.reachable_successors assms(4) by blast have 2: "u \ gunodes A w" using assms(3) 1 by auto show "v \ graph A w (GREATEST k. u \ graph A w k)" unfolding rank_def[symmetric] proof (rule graph_successors) show "v \ gunodes A w" using assms(3) by this show "u \ gusuccessors A w v" using assms(4) by this show "u \ graph A w (rank A w u)" using rank_member assms(1, 2) 2 by this qed show "\ k. v \ graph A w k \ k \ 2 * card (nodes A)" using graph_le assms(1, 2) by blast qed lemma language_ranking: assumes "finite (nodes A)" "w \ language A" shows "ranking A w (rank A w)" unfolding ranking_def proof (intro conjI ballI allI impI) fix v assume 1: "v \ gunodes A w" have 2: "v \ graph A w (rank A w v)" using rank_member assms 1 by this show "rank A w v \ 2 * card (nodes A)" using graph_le assms 2 by this next fix v u assume 1: "v \ gunodes A w" "u \ gusuccessors A w v" show "rank A w u \ rank A w v" using rank_le assms 1 by this next fix v assume 1: "v \ gunodes A w" "gaccepting A v" have 2: "v \ graph A w (rank A w v)" using rank_member assms 1(1) by this have 3: "v \ graph A w (Suc (rank A w v))" using rank_removed assms by this have 4: "v \ prune A w (graph A w (rank A w v))" using 2 1(2) unfolding prune_def by auto have 5: "graph A w (Suc (rank A w v)) \ prune A w (graph A w (rank A w v))" using 3 4 by blast show "even (rank A w v)" using 5 by auto next fix v r k assume 1: "v \ gunodes A w" "gurun A w r v" "smap (rank A w) (gtrace r v) = sconst k" have "sset (gtrace r v) \ gureachable A w v" using 1(2) by (metis graph.reachable.reflexive graph.reachable_trace) then have 6: "sset (gtrace r v) \ gunodes A w" using 1(1) by blast have 60: "rank A w ` sset (gtrace r v) \ {k}" using 1(3) by (metis equalityD1 sset_sconst stream.set_map) have 50: "sset (gtrace r v) \ graph A w k" using rank_member[OF assms] subsetD[OF 6] 60 unfolding image_subset_iff by auto have 70: "grun A w (graph A w k) r v" using grun_subset 1(2) 50 by this have 7: "sset (gtrace r v) \ clean A w (graph A w k)" unfolding clean_def using 50 infinite_greachable_gtrace[OF 70] by auto have 8: "sset (gtrace r v) \ graph A w (Suc k) = {}" using rank_removed[OF assms] 60 by blast have 9: "sset (gtrace r v) \ {}" using stream.set_sel(1) by auto have 10: "graph A w (Suc k) \ clean A w (graph A w k)" using 7 8 9 by blast show "odd k" using 10 unfolding graph_Suc by auto qed subsection \Correctness Theorem\ theorem language_ranking_iff: assumes "finite (nodes A)" shows "w \ language A \ (\ f. ranking A w f)" using ranking_language language_ranking assms by blast end 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 transitionei acceptingei + alphabetei initialei transitionei conditionei 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,443 +1,436 @@ (* 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 + "DBA.run (dba UNIV p \ \) (to_stream w) p" unfolding dba.run_alt_def by simp lemma dca_run: - "DCA.run (dca UNIV p \ \) (to_stream w) p" - unfolding DCA.run_def by (rule transition_system.snth_run) fastforce + "DCA.run (dca UNIV p \ \) (to_stream w) p" unfolding dca.run_alt_def by simp 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.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)) + unfolding \\<^sub>\_def dba.initial_def dba.accepting_def by (auto simp: dba_run) 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.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)) + unfolding \\<^sub>\_GF_def dba.initial_def dba.accepting_def by (auto simp: dba_run) 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.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)) + unfolding \\<^sub>\_def dca.initial_def dca.rejecting_def by (auto simp: dca_run) 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.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)) + unfolding \\<^sub>\_FG_def dca.initial_def dca.rejecting_def by (auto simp: dca_run) 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 + unfolding \\<^sub>\_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 + unfolding \\<^sub>\_GF_def dba.nodes_alt_def dba.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 + unfolding \\<^sub>\_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 + unfolding \\<^sub>\_FG_def dca.nodes_alt_def dca.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.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) + 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 + unfolding \_def dca.nodes_alt_def dca.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) + 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) + 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)" + "\\<^sub>2 xs ys = DBA_Combine.intersect_list (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) + by (simp add: \\<^sub>2_def \\<^sub>\_GF_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)" + "\\<^sub>3 xs ys = DCA_Combine.intersect_list (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) + by (simp add: \\<^sub>3_def \\<^sub>\_FG_def) -definition "\' \ xs ys = dbcrai (\\<^sub>2 xs ys) (dcai (\\<^sub>1 \ xs) (\\<^sub>3 xs ys))" +definition "\' \ xs ys = intersect_bc (\\<^sub>2 xs ys) (DCA_Combine.intersect (\\<^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) + by (simp add: \'_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 \))" +definition "ltl_to_dra \ = DRA_Combine.union_list (map (\(xs, ys). \' \ xs ys) (advice_sets \))" lemma ltl_to_dra_language: "to_omega ` DRA.language (ltl_to_dra \) = language_ltln \" proof - have "(\(a, b)\set (advice_sets \). dra.alphabet (\' \ a b)) = (\(a, b)\set (advice_sets \). dra.alphabet (\' \ a b))" using advice_sets_not_empty by (simp add: \'_alphabet) - then have *: "DRA.language (draul (map (\(x, y). \' \ x y) (advice_sets \))) = + then have *: "DRA.language (DRA_Combine.union_list (map (\(x, y). \' \ x y) (advice_sets \))) = \ (DRA.language ` set (map (\(x, y). \' \ x y) (advice_sets \)))" by (simp add: split_def) 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 using * by (auto simp add: ltl_to_dra_def) 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) + by (auto simp: ltl_to_dra_def \'_alphabet) 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 \) (transition \) (accepting \)" + "dra_set_alphabet \ \ = dra \ (initial \) (transition \) (condition \)" 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) + by (auto simp add: dra_set_alphabet_def dra.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) + unfolding dra_set_alphabet_def dra.nodes_alt_def dra.reachable_alt_def dra.path_alt_def + by auto 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 \) (transition \) (accepting \)" + "dra_to_drai \ \ = drai \ (initial \) (transition \) (condition \)" 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) (transition A) (accepting A)) \ \Id, Id\drai_dra_rel" + then have "(dra_to_drai A \, dra (alphabet A) (initial A) (transition A) (condition 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/LTL_to_DRA/DRA_Instantiation.thy b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Instantiation.thy --- a/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Instantiation.thy +++ b/thys/LTL_Master_Theorem/LTL_to_DRA/DRA_Instantiation.thy @@ -1,419 +1,417 @@ (* Author: Benedikt Seidl License: BSD *) section \Instantiation of the LTL to DRA construction\ theory DRA_Instantiation imports DRA_Construction LTL.Equivalence_Relations "HOL-Library.Log_Nat" begin text \We instantiate the construction locale with propositional equivalence and obtain a function converting a formula into an abstract automaton.\ global_interpretation ltl_to_dra: dra_construction "(\\<^sub>P)" rep_ltln\<^sub>P abs_ltln\<^sub>P defines ltl_to_dra = ltl_to_dra.ltl_to_dra and ltl_to_dra_alphabet = ltl_to_dra.ltl_to_dra_alphabet and \' = ltl_to_dra.\' and \\<^sub>1 = ltl_to_dra.\\<^sub>1 and \\<^sub>2 = ltl_to_dra.\\<^sub>2 and \\<^sub>3 = ltl_to_dra.\\<^sub>3 and \\<^sub>\_FG = ltl_to_dra.\\<^sub>\_FG and \\<^sub>\_GF = ltl_to_dra.\\<^sub>\_GF and af_letter\<^sub>G = ltl_to_dra.af_letter\<^sub>G and af_letter\<^sub>F = ltl_to_dra.af_letter\<^sub>F and af_letter\<^sub>G_lifted = ltl_to_dra.af_letter\<^sub>G_lifted and af_letter\<^sub>F_lifted = ltl_to_dra.af_letter\<^sub>F_lifted and af_letter\<^sub>\_lifted = ltl_to_dra.af_letter\<^sub>\_lifted and \ = ltl_to_dra.\ and af_letter\<^sub>\ = ltl_to_dra.af_letter\<^sub>\ by unfold_locales (meson Quotient_abs_rep Quotient_ltln\<^sub>P, simp add: Quotient_abs_rep Quotient_ltln\<^sub>P ltln\<^sub>P.abs_eq_iff) text \We obtain the following theorem:\ thm ltl_to_dra.ltl_to_dra_language text \Furthermore, we verify the size bound of the automaton to be double-exponential.\ lemma prop_equiv_nested_prop_atoms_finite: "finite {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" using prop_equiv_finite'[OF nested_prop_atoms_finite] . lemma prop_equiv_nested_prop_atoms_card: "card {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms \} \ 2 ^ 2 ^ card (nested_prop_atoms \)" using prop_equiv_card'[OF nested_prop_atoms_finite] . lemma prop_equiv_nested_prop_atoms\<^sub>\_finite: "finite {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X}" using prop_equiv_finite'[OF nested_prop_atoms\<^sub>\_finite] by fast lemma prop_equiv_nested_prop_atoms\<^sub>\_card: "card {abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X} \ 2 ^ 2 ^ card (nested_prop_atoms \)" (is "?lhs \ ?rhs") proof - have "finite {abs_ltln\<^sub>P \ | \. prop_atoms \ \ nested_prop_atoms\<^sub>\ \ X}" by (simp add: prop_equiv_nested_prop_atoms\<^sub>\_finite nested_prop_atoms\<^sub>\_finite prop_equiv_finite) then have "?lhs \ card {abs_ltln\<^sub>P \ | \. prop_atoms \ \ (nested_prop_atoms\<^sub>\ \ X)}" using card_mono prop_equiv_subset by blast also have "\ \ 2 ^ 2 ^ card (nested_prop_atoms\<^sub>\ \ X)" using prop_equiv_card[OF nested_prop_atoms\<^sub>\_finite] by fast also have "\ \ ?rhs" using nested_prop_atoms\<^sub>\_card by auto finally show ?thesis . qed lemma \\<^sub>\_GF_nodes_finite: "finite (DBA.nodes (\\<^sub>\_GF \))" using finite_subset[OF ltl_to_dra.\\<^sub>\_GF_nodes prop_equiv_nested_prop_atoms_finite] . lemma \\<^sub>\_FG_nodes_finite: "finite (DCA.nodes (\\<^sub>\_FG \))" using finite_subset[OF ltl_to_dra.\\<^sub>\_FG_nodes prop_equiv_nested_prop_atoms_finite] . lemma \\<^sub>\_GF_nodes_card: "card (DBA.nodes (\\<^sub>\_GF \)) \ 2 ^ 2 ^ card (nested_prop_atoms (F\<^sub>n \))" using le_trans[OF card_mono[OF prop_equiv_nested_prop_atoms_finite ltl_to_dra.\\<^sub>\_GF_nodes] prop_equiv_nested_prop_atoms_card] . lemma \\<^sub>\_FG_nodes_card: "card (DCA.nodes (\\<^sub>\_FG \)) \ 2 ^ 2 ^ card (nested_prop_atoms (G\<^sub>n \))" using le_trans[OF card_mono[OF prop_equiv_nested_prop_atoms_finite ltl_to_dra.\\<^sub>\_FG_nodes] prop_equiv_nested_prop_atoms_card] . lemma \\<^sub>2_nodes_finite_helper: "list_all (finite \ DBA.nodes) (map (\\. \\<^sub>\_GF (\[set ys]\<^sub>\)) xs)" by (auto simp: list.pred_map list_all_iff \\<^sub>\_GF_nodes_finite) lemma \\<^sub>2_nodes_finite: "finite (DBA.nodes (\\<^sub>2 xs ys))" - unfolding ltl_to_dra.\\<^sub>2_def using dbail_nodes_finite \\<^sub>2_nodes_finite_helper . + unfolding ltl_to_dra.\\<^sub>2_def using DBA_Combine.intersect_list_nodes_finite \\<^sub>2_nodes_finite_helper . lemma \\<^sub>3_nodes_finite_helper: "list_all (finite \ DCA.nodes) (map (\\. \\<^sub>\_FG (\[set xs]\<^sub>\)) ys)" by (auto simp: list.pred_map list_all_iff \\<^sub>\_FG_nodes_finite) lemma \\<^sub>3_nodes_finite: "finite (DCA.nodes (\\<^sub>3 xs ys))" - unfolding ltl_to_dra.\\<^sub>3_def using dcail_finite \\<^sub>3_nodes_finite_helper . + unfolding ltl_to_dra.\\<^sub>3_def using DCA_Combine.intersect_list_nodes_finite \\<^sub>3_nodes_finite_helper . (* TODO add to HOL/Groups_List.thy *) lemma list_prod_mono: "f \ g \ (\x\xs. f x) \ (\x\xs. g x)" for f g :: "'a \ nat" by (induction xs) (auto simp: le_funD mult_le_mono) (* TODO add to HOL/Groups_List.thy *) lemma list_prod_const: "(\x. x \ set xs \ f x \ c) \ (\x\xs. f x) \ c ^ length xs" for f :: "'a \ nat" by (induction xs) (auto simp: mult_le_mono) (* TODO add to HOL/Finite_Set.thy *) lemma card_insert_Suc: "card (insert x S) \ Suc (card S)" by (metis Suc_n_not_le_n card.infinite card_insert_if finite_insert linear) (* TODO add to HOL/Power.thy *) lemma nat_power_le_imp_le: "0 < a \ a \ b \ x ^ a \ x ^ b" for x :: nat by (metis leD linorder_le_less_linear nat_power_less_imp_less neq0_conv power_eq_0_iff) (* TODO add to HOL/Power.thy *) lemma const_less_power: "n < x ^ n" if "x > 1" using that by (induction n) (auto simp: less_trans_Suc) (* TODO add to HOL-Library/Log_Nat.thy *) lemma floorlog_le_const: "floorlog x n \ n" by (induction n) (simp add: floorlog_eq_zero_iff, metis Suc_lessI floorlog_le_iff le_SucI power_inject_exp) lemma \\<^sub>2_nodes_card: assumes "length xs \ n" and "\\. \ \ set xs \ card (nested_prop_atoms \) \ n" shows "card (DBA.nodes (\\<^sub>2 xs ys)) \ 2 ^ 2 ^ (n + floorlog 2 n + 2)" proof - have 1: "\\. \ \ set xs \ card (nested_prop_atoms (F\<^sub>n \[set ys]\<^sub>\)) \ Suc n" proof - fix \ assume "\ \ set xs" have "card (nested_prop_atoms (F\<^sub>n (\[set ys]\<^sub>\))) \ Suc (card (nested_prop_atoms (\[set ys]\<^sub>\)))" by (simp add: card_insert_Suc) also have "\ \ Suc (card (nested_prop_atoms \))" by (simp add: FG_advice_nested_prop_atoms_card) also have "\ \ Suc n" by (simp add: assms(2) \\ \ set xs\) finally show "card (nested_prop_atoms (F\<^sub>n (\[set ys]\<^sub>\))) \ Suc n" . qed have "(\\\xs. card (DBA.nodes (\\<^sub>\_GF (\[set ys]\<^sub>\)))) \ (\\\xs. 2 ^ 2 ^ card (nested_prop_atoms (F\<^sub>n (\[set ys]\<^sub>\))))" by (rule list_prod_mono) (insert \\<^sub>\_GF_nodes_card le_fun_def, blast) also have "\ \ (2 ^ 2 ^ Suc n) ^ length xs" by (rule list_prod_const) (metis 1 Suc_leI nat_power_le_imp_le nat_power_eq_Suc_0_iff neq0_conv pos2 zero_less_power) also have "\ \ (2 ^ 2 ^ Suc n) ^ n" using assms(1) nat_power_le_imp_le by fastforce also have "\ = 2 ^ (n * 2 ^ Suc n)" by (metis Groups.mult_ac(2) power_mult) also have "\ \ 2 ^ (2 ^ floorlog 2 n * 2 ^ Suc n)" by (cases "n = 0") (auto simp: floorlog_bounds less_imp_le_nat) also have "\ = 2 ^ 2 ^ (Suc n + floorlog 2 n)" by (simp add: power_add) finally have 2: "(\\\xs. card (DBA.nodes (\\<^sub>\_GF (\[set ys]\<^sub>\)))) \ 2 ^ 2 ^ (Suc n + floorlog 2 n)" . have "card (DBA.nodes (\\<^sub>2 xs ys)) \ max 1 (length xs) * (\\\xs. card (DBA.nodes (\\<^sub>\_GF (\[set ys]\<^sub>\))))" - using dbail_nodes_card[OF \\<^sub>2_nodes_finite_helper] + using DBA_Combine.intersect_list_nodes_card[OF \\<^sub>2_nodes_finite_helper] by (auto simp: ltl_to_dra.\\<^sub>2_def comp_def) also have "\ \ max 1 n * 2 ^ 2 ^ (Suc n + floorlog 2 n)" using assms(1) 2 by (simp add: mult_le_mono) also have "\ \ 2 ^ (floorlog 2 n) * 2 ^ 2 ^ (Suc n + floorlog 2 n)" by (cases "n = 0") (auto simp: floorlog_bounds less_imp_le_nat) also have "\ = 2 ^ (floorlog 2 n + 2 ^ (Suc n + floorlog 2 n))" by (simp add: power_add) also have "\ \ 2 ^ (n + 2 ^ (Suc n + floorlog 2 n))" by (simp add: floorlog_le_const) also have "\ \ 2 ^ 2 ^ (n + floorlog 2 n + 2)" by simp (metis const_less_power Suc_1 add_Suc_right add_leE lessI less_imp_le_nat power_Suc) finally show ?thesis . qed lemma \\<^sub>3_nodes_card: assumes "length ys \ n" and "\\. \ \ set ys \ card (nested_prop_atoms \) \ n" shows "card (DCA.nodes (\\<^sub>3 xs ys)) \ 2 ^ 2 ^ (n + floorlog 2 n + 1)" proof - have 1: "\\. \ \ set ys \ card (DCA.nodes (\\<^sub>\_FG (\[set xs]\<^sub>\))) \ 2 ^ 2 ^ Suc n" proof - fix \ assume "\ \ set ys" have "card (nested_prop_atoms (G\<^sub>n \[set xs]\<^sub>\)) \ Suc (card (nested_prop_atoms (\[set xs]\<^sub>\)))" by (simp add: card_insert_Suc) also have "\ \ Suc (card (nested_prop_atoms \))" by (simp add: GF_advice_nested_prop_atoms_card) also have "\ \ Suc n" by (simp add: assms(2) \\ \ set ys\) finally have 2: "card (nested_prop_atoms (G\<^sub>n \[set xs]\<^sub>\)) \ Suc n" . then show "?thesis \" by (intro le_trans[OF \\<^sub>\_FG_nodes_card]) (meson one_le_numeral power_increasing) qed have "card (DCA.nodes (\\<^sub>3 xs ys)) \ (\\\ys. card (DCA.nodes (\\<^sub>\_FG (\[set xs]\<^sub>\))))" - unfolding ltl_to_dra.\\<^sub>3_def using dcail_nodes_card[OF \\<^sub>3_nodes_finite_helper] + unfolding ltl_to_dra.\\<^sub>3_def using DCA_Combine.intersect_list_nodes_card[OF \\<^sub>3_nodes_finite_helper] by (auto simp: comp_def) also have "\ \ (2 ^ 2 ^ Suc n) ^ length ys" by (rule list_prod_const) (rule 1) also have "\ \ (2 ^ 2 ^ Suc n) ^ n" by (simp add: assms(1) power_increasing) also have "\ \ 2 ^ (n * 2 ^ Suc n)" by (metis le_refl mult.commute power_mult) also have "\ \ 2 ^ (2 ^ floorlog 2 n * 2 ^ Suc n)" by (cases \n > 0\) (simp_all add: floorlog_bounds less_imp_le_nat) also have "\ = 2 ^ 2 ^ (n + floorlog 2 n + 1)" by (simp add: power_add) finally show ?thesis . qed lemma \\<^sub>1_nodes_finite: "finite (DCA.nodes (\\<^sub>1 \ xs))" unfolding ltl_to_dra.\\<^sub>1_def by (metis (no_types, lifting) finite_subset ltl_to_dra.\_nodes finite_SigmaI prop_equiv_nested_prop_atoms\<^sub>\_finite prop_equiv_nested_prop_atoms_finite) lemma \\<^sub>1_nodes_card: assumes "card (subfrmlsn \) \ n" shows "card (DCA.nodes (\\<^sub>1 \ xs)) \ 2 ^ 2 ^ (n + 1)" proof - let ?fst = "{abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms \}" let ?snd = "{abs_ltln\<^sub>P \ | \. nested_prop_atoms \ \ nested_prop_atoms\<^sub>\ \ (set xs)}" have 1: "card (nested_prop_atoms \) \ n" by (meson card_mono[OF subfrmlsn_finite nested_prop_atoms_subfrmlsn] assms le_trans) have "card (DCA.nodes (\\<^sub>1 \ xs)) \ card (?fst \ ?snd)" unfolding ltl_to_dra.\\<^sub>1_def by (rule card_mono) (simp_all add: ltl_to_dra.\_nodes prop_equiv_nested_prop_atoms\<^sub>\_finite prop_equiv_nested_prop_atoms_finite) also have "\ = card ?fst * card ?snd" using prop_equiv_nested_prop_atoms\<^sub>\_finite card_cartesian_product by blast also have "\ \ 2 ^ 2 ^ card (nested_prop_atoms \) * 2 ^ 2 ^ card (nested_prop_atoms \)" using prop_equiv_nested_prop_atoms\<^sub>\_card prop_equiv_nested_prop_atoms_card mult_le_mono by blast also have "\ = 2 ^ 2 ^ (card (nested_prop_atoms \) + 1)" by (simp add: semiring_normalization_rules(36)) also have "\ \ 2 ^ 2 ^ (n + 1)" using assms 1 by simp finally show ?thesis . qed lemma \'_nodes_finite: "finite (DRA.nodes (\' \ xs ys))" unfolding ltl_to_dra.\'_def - using dcai_nodes_finite dbcrai_nodes_finite + using intersect_nodes_finite intersect_bc_nodes_finite using \\<^sub>1_nodes_finite \\<^sub>2_nodes_finite \\<^sub>3_nodes_finite by fast lemma \'_nodes_card: assumes "length xs \ n" and "\\. \ \ set xs \ card (nested_prop_atoms \) \ n" and "length ys \ n" and "\\. \ \ set ys \ card (nested_prop_atoms \) \ n" and "card (subfrmlsn \) \ n" shows "card (DRA.nodes (\' \ xs ys)) \ 2 ^ 2 ^ (n + floorlog 2 n + 4)" proof - have "n + 1 \ n + floorlog 2 n + 2" by auto then have 1: "(2::nat) ^ (n + 1) \ 2 ^ (n + floorlog 2 n + 2)" using one_le_numeral power_increasing by blast have "card (DRA.nodes (\' \ xs ys)) \ card (DCA.nodes (\\<^sub>1 \ xs)) * card (DBA.nodes (\\<^sub>2 xs ys)) * card (DCA.nodes (\\<^sub>3 xs ys))" (is "?lhs \ ?rhs") proof (unfold ltl_to_dra.\'_def) - have "card (DBA.nodes (\\<^sub>2 xs ys)) * card (DCA.nodes (dcai (\\<^sub>1 \ xs) (\\<^sub>3 xs ys))) \ ?rhs" - by (simp add: dcai_nodes_card[OF \\<^sub>1_nodes_finite \\<^sub>3_nodes_finite]) - then show "card (DRA.nodes (dbcrai (\\<^sub>2 xs ys) (dcai (\\<^sub>1 \ xs) (\\<^sub>3 xs ys)))) \ ?rhs" - by (meson dbcrai_nodes_card[OF \\<^sub>2_nodes_finite dcai_nodes_finite[OF \\<^sub>1_nodes_finite \\<^sub>3_nodes_finite]] basic_trans_rules(23)) + have "card (DBA.nodes (\\<^sub>2 xs ys)) * card (DCA.nodes (DCA_Combine.intersect (\\<^sub>1 \ xs) (\\<^sub>3 xs ys))) \ ?rhs" + by (simp add: intersect_nodes_card[OF \\<^sub>1_nodes_finite \\<^sub>3_nodes_finite]) + then show "card (DRA.nodes (intersect_bc (\\<^sub>2 xs ys) (DCA_Combine.intersect (\\<^sub>1 \ xs) (\\<^sub>3 xs ys)))) \ ?rhs" + by (meson intersect_bc_nodes_card[OF \\<^sub>2_nodes_finite intersect_nodes_finite[OF \\<^sub>1_nodes_finite \\<^sub>3_nodes_finite]] basic_trans_rules(23)) qed also have "\ \ 2 ^ 2 ^ (n + 1) * 2 ^ 2 ^ (n + floorlog 2 n + 2) * 2 ^ 2 ^ (n + floorlog 2 n + 1)" using \\<^sub>1_nodes_card[OF assms(5)] \\<^sub>2_nodes_card[OF assms(1,2)] \\<^sub>3_nodes_card[OF assms(3,4)] by (metis mult_le_mono) also have "\ = 2 ^ (2 ^ (n + 1) + 2 ^ (n + floorlog 2 n + 2) + 2 ^ (n + floorlog 2 n + 1))" by (metis power_add) also have "\ \ 2 ^ (4 * 2 ^ (n + floorlog 2 n + 2))" using 1 by auto finally show ?thesis by (simp add: numeral.simps(2) power_add) qed lemma subformula_nested_prop_atoms_subfrmlsn: "\ \ subfrmlsn \ \ nested_prop_atoms \ \ subfrmlsn \" using nested_prop_atoms_subfrmlsn subfrmlsn_subset by blast lemma ltl_to_dra_nodes_finite: "finite (DRA.nodes (ltl_to_dra \))" unfolding ltl_to_dra.ltl_to_dra_def - apply (rule draul_nodes_finite) + apply (rule DRA_Combine.union_list_nodes_finite) apply (simp add: split_def ltl_to_dra.\'_alphabet advice_sets_not_empty) apply (simp add: list.pred_set split_def \'_nodes_finite) done lemma ltl_to_dra_nodes_card: assumes "card (subfrmlsn \) \ n" shows "card (DRA.nodes (ltl_to_dra \)) \ 2 ^ 2 ^ (2 * n + floorlog 2 n + 4)" proof - let ?map = "map (\(x, y). \' \ x y) (advice_sets \)" have 1: "\x::nat. x > 0 \ x ^ length (advice_sets \) \ x ^ 2 ^ card (subfrmlsn \)" by (metis advice_sets_length linorder_not_less nat_power_less_imp_less) have "card (DRA.nodes (ltl_to_dra \)) \ prod_list (map (card \ DRA.nodes) ?map)" unfolding ltl_to_dra.ltl_to_dra_def - apply (rule draul_nodes_card) - unfolding set_map image_image split_def ltl_to_dra.\'_alphabet - apply (simp add: advice_sets_not_empty) - unfolding split_def list.pred_set using \'_nodes_finite by auto + apply (rule DRA_Combine.union_list_nodes_card) + unfolding list.pred_set using \'_nodes_finite by auto also have "\ = (\(x, y)\advice_sets \. card (DRA.nodes (\' \ x y)))" by (induction "advice_sets \") (auto, metis (no_types, lifting) comp_apply split_def) also have "\ \ (2 ^ 2 ^ (n + floorlog 2 n + 4)) ^ length (advice_sets \)" proof (rule list_prod_const, unfold split_def, rule \'_nodes_card) show "\x. x \ set (advice_sets \) \ length (fst x) \ n" using advice_sets_element_length assms by fastforce show "\x \. \x \ set (advice_sets \); \ \ set (fst x)\ \ card (nested_prop_atoms \) \ n" using advice_sets_element_subfrmlsn(1) assms subformula_nested_prop_atoms_subfrmlsn subformulas\<^sub>\_subfrmlsn by (metis (no_types, lifting) card_mono subfrmlsn_finite subset_iff sup.absorb_iff2 sup.coboundedI1 surjective_pairing) show "\x. x \ set (advice_sets \) \ length (snd x) \ n" using advice_sets_element_length assms by fastforce show "\x \. \x \ set (advice_sets \); \ \ set (snd x)\ \ card (nested_prop_atoms \) \ n" using advice_sets_element_subfrmlsn(2) assms subformula_nested_prop_atoms_subfrmlsn subformulas\<^sub>\_subfrmlsn by (metis (no_types, lifting) card_mono subfrmlsn_finite subset_iff sup.absorb_iff2 sup.coboundedI1 surjective_pairing) qed (insert assms, blast) also have "\ \ (2 ^ 2 ^ (n + floorlog 2 n + 4)) ^ (2 ^ card (subfrmlsn \))" by (simp add: 1) also have "\ \ (2 ^ 2 ^ (n + floorlog 2 n + 4)) ^ (2 ^ n)" by (simp add: assms power_increasing) also have "\ = 2 ^ (2 ^ n * 2 ^ (n + floorlog 2 n + 4))" by (simp add: ac_simps power_mult [symmetric]) also have "\ = 2 ^ 2 ^ (2 * n + floorlog 2 n + 4)" by (simp add: power_add) (simp add: mult_2 power_add) finally show ?thesis . qed theorem ltl_to_dra_size: "card (DRA.nodes (ltl_to_dra \)) \ 2 ^ 2 ^ (2 * size \ + floorlog 2 (size \) + 4)" using ltl_to_dra_nodes_card subfrmlsn_card by blast 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" + ^ "Acceptance: " ^ serialize_acceptance (conditionei 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" + ^ serialize_state_labels (conditionei 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 (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,29 @@ 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" +imports "../Deterministic" begin datatype ('label, 'state) dba = dba (alphabet: "'label set") (initial: "'state") (transition: "'label \ 'state \ 'state") (accepting: "'state pred") - global_interpretation dba: transition_system_initial - "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 + (* TODO: if we interpret everything at once, some abbreviations don't get folded back + see DBA_Combine.intersection_list.combine_finite *) + global_interpretation dba: automaton dba alphabet initial transition accepting + defines path = dba.path and run = dba.run and reachable = dba.reachable and nodes = dba.nodes + by unfold_locales auto + global_interpretation dba: automaton_trace dba alphabet initial transition accepting infs + defines language = dba.language + by standard 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 + (* TODO: why does this happen? if I can fix it here, why can't the interpretation do it? + same happens with reachable, requiring the use of defines directives *) + abbreviation successors where "successors \ dba.successors TYPE('label)" 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,304 +1,107 @@ section \Deterministic Büchi Automata Combinations\ theory DBA_Combine -imports "DBA" "DGBA" +imports DBA DGBA begin - definition dbgail :: "('label, 'state) dba list \ ('label, 'state list) dgba" where - "dbgail AA \ dgba - (\ (dba.alphabet ` set AA)) - (map dba.initial AA) - (\ 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 + global_interpretation degeneralization: automaton_degeneralization_trace + dgba dgba.alphabet dgba.initial dgba.transition dgba.accepting "gen infs" + dba dba.alphabet dba.initial dba.transition dba.accepting infs + defines degeneralize = degeneralization.degeneralize + by unfold_locales auto - lemma dbgail_language[simp]: "DGBA.language (dbgail AA) = \ (DBA.language ` set AA)" - 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 \ \ (DBA.language ` set AA)" - 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 + lemmas degeneralize_language[simp] = degeneralization.degeneralize_language[folded DBA.language_def] + lemmas degeneralize_nodes_finite[iff] = degeneralization.degeneralize_nodes_finite[folded DBA.nodes_def] + lemmas degeneralize_nodes_card = degeneralization.degeneralize_nodes_card[folded DBA.nodes_def] - definition dbail :: "('label, 'state) dba list \ ('label, 'state list degen) dba" where - "dbail = dgbad \ dbgail" + global_interpretation intersection: automaton_intersection_trace + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + dgba.dgba dgba.alphabet dgba.initial dgba.transition dgba.accepting "gen infs" + "\ c\<^sub>1 c\<^sub>2. [c\<^sub>1 \ fst, c\<^sub>2 \ snd]" + defines intersect' = intersection.combine + by unfold_locales auto - 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)" + lemmas intersect'_language[simp] = intersection.combine_language[folded DGBA.language_def] + lemmas intersect'_nodes_finite = intersection.combine_nodes_finite[folded DGBA.nodes_def] + lemmas intersect'_nodes_card = intersection.combine_nodes_card[folded DGBA.nodes_def] + + global_interpretation union: automaton_union_trace + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + "\ c\<^sub>1 c\<^sub>2 pq. (c\<^sub>1 \ fst) pq \ (c\<^sub>2 \ snd) pq" + defines union = union.combine + by (unfold_locales) (simp del: comp_apply) + + lemmas union_language = union.combine_language + lemmas union_nodes_finite = union.combine_nodes_finite + lemmas union_nodes_card = union.combine_nodes_card + + global_interpretation intersection_list: automaton_intersection_list_trace + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + dgba.dgba dgba.alphabet dgba.initial dgba.transition dgba.accepting "gen infs" + "\ cs. map (\ k pp. (cs ! k) (pp ! k)) [0 ..< length cs]" + defines intersect_list' = intersection_list.combine + by (unfold_locales) (auto simp: gen_def comp_def) + + lemmas intersect_list'_language[simp] = intersection_list.combine_language[folded DGBA.language_def] + lemmas intersect_list'_nodes_finite = intersection_list.combine_nodes_finite[folded DGBA.nodes_def] + lemmas intersect_list'_nodes_card = intersection_list.combine_nodes_card[folded DGBA.nodes_def] + + global_interpretation union_list: automaton_union_list_trace + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + "\ cs pp. \ k < length cs. (cs ! k) (pp ! k)" + defines union_list = union_list.combine + by (unfold_locales) (simp add: comp_def) + + lemmas union_list_language = union_list.combine_language + lemmas union_list_nodes_finite = union_list.combine_nodes_finite + lemmas union_list_nodes_card = union_list.combine_nodes_card + + (* TODO: these compound definitions are annoying, can we move those into Deterministic theory *) + + abbreviation intersect where "intersect A B \ degeneralize (intersect' A B)" + + lemma intersect_language[simp]: "DBA.language (intersect A B) = DBA.language A \ DBA.language B" + by simp + lemma intersect_nodes_finite: + assumes "finite (DBA.nodes A)" "finite (DBA.nodes B)" + shows "finite (DBA.nodes (intersect A B))" + using intersect'_nodes_finite assms by simp + lemma intersect_nodes_card: + assumes "finite (DBA.nodes A)" "finite (DBA.nodes B)" + shows "card (DBA.nodes (intersect A B)) \ 2 * card (DBA.nodes A) * card (DBA.nodes B)" 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 + have "card (DBA.nodes (intersect A B)) \ + max 1 (length (dgba.accepting (intersect' A B))) * card (DGBA.nodes (intersect' A B))" + using degeneralize_nodes_card by this + also have "length (dgba.accepting (intersect' A B)) = 2" by simp + also have "card (DGBA.nodes (intersect' A B)) \ card (DBA.nodes A) * card (DBA.nodes B)" + using intersect'_nodes_card assms by this finally show ?thesis by simp qed - lemma dbail_language [simp]: - "DBA.language (dbail AA) = \ (DBA.language ` set AA)" - by (simp add: dbail_def) - - 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.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 - (\ (dba.alphabet ` set AA)) - (map dba.initial AA) - (\ a pp. map2 (\ A p. dba.transition A a p) AA pp) - (\ pp. \ k < length AA. dba.accepting (AA ! k) (pp ! k))" + abbreviation intersect_list where "intersect_list AA \ degeneralize (intersect_list' AA)" - 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 "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" - 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 "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" + lemma intersect_list_language[simp]: "DBA.language (intersect_list AA) = \ (DBA.language ` set AA)" + by simp + lemma intersect_list_nodes_finite: 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 "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" + shows "finite (DBA.nodes (intersect_list AA))" + using intersect_list'_nodes_finite assms by simp + lemma intersect_list_nodes_card: assumes "list_all (finite \ DBA.nodes) AA" - shows "card (DBA.nodes (dbaul AA)) \ prod_list (map (card \ DBA.nodes) AA)" + shows "card (DBA.nodes (intersect_list AA)) \ max 1 (length 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 "\ (dba.alphabet ` set AA) = \ (dba.alphabet ` set AA)" - shows "DBA.language (dbaul AA) = \ (DBA.language ` set AA)" - 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 \ \ (DBA.language ` set AA)" - 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 + have "card (DBA.nodes (intersect_list AA)) \ + max 1 (length (dgba.accepting (intersect_list' AA))) * card (DGBA.nodes (intersect_list' AA))" + using degeneralize_nodes_card by this + also have "length (dgba.accepting (intersect_list' AA)) = length AA" by simp + also have "card (DGBA.nodes (intersect_list' AA)) \ prod_list (map (card \ DBA.nodes) AA)" + using intersect_list'_nodes_card assms by this + finally show ?thesis by simp 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,25 @@ section \Deterministic Generalized Büchi Automata\ theory DGBA -imports - "DBA" - "../../Basic/Degeneralization" +imports "../Deterministic" begin datatype ('label, 'state) dgba = dgba (alphabet: "'label set") (initial: "'state") (transition: "'label \ 'state \ 'state") (accepting: "'state pred gen") - global_interpretation dgba: transition_system_initial - "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 + global_interpretation dgba: automaton dgba alphabet initial transition accepting + defines path = dgba.path and run = dgba.run and reachable = dgba.reachable and nodes = dgba.nodes + by unfold_locales auto + global_interpretation dgba: automaton_trace dgba alphabet initial transition accepting "gen infs" + defines language = dgba.language + by standard 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). (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.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.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.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 + abbreviation successors where "successors \ dgba.successors TYPE('label)" 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,25 @@ 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" +imports "../Deterministic" begin datatype ('label, 'state) dca = dca (alphabet: "'label set") (initial: "'state") (transition: "'label \ 'state \ 'state") (rejecting: "'state \ bool") - global_interpretation dca: transition_system_initial - "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 + global_interpretation dca: automaton dca alphabet initial transition rejecting + defines path = dca.path and run = dca.run and reachable = dca.reachable and nodes = dca.nodes + by unfold_locales auto + global_interpretation dca: automaton_trace dca alphabet initial transition rejecting fins + defines language = dca.language + by standard 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 + abbreviation successors where "successors \ dca.successors TYPE('label)" 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,109 @@ section \Deterministic Co-Büchi Automata Combinations\ theory DCA_Combine -imports "DCA" "DGCA" +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.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 - (\ (dca.alphabet ` set AA)) - (map dca.initial AA) - (\ 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 + global_interpretation degeneralization: automaton_degeneralization_trace + dgca dgca.alphabet dgca.initial dgca.transition dgca.rejecting "cogen fins" + dca dca.alphabet dca.initial dca.transition dca.rejecting fins + defines degeneralize = degeneralization.degeneralize + by unfold_locales auto - lemma dcail_language[simp]: "DCA.language (dcail AA) = \ (DCA.language ` set AA)" - 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 \ \ (DCA.language `set AA)" - 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 - (\ (dca.alphabet ` set AA)) - (map dca.initial AA) - (\ 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 "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" - 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+ + lemmas degeneralize_language[simp] = degeneralization.degeneralize_language[folded DCA.language_def] + lemmas degeneralize_nodes_finite[iff] = degeneralization.degeneralize_nodes_finite[folded DCA.nodes_def] + lemmas degeneralize_nodes_card = degeneralization.degeneralize_nodes_card[folded DCA.nodes_def] - lemma dcgaul_nodes_finite[intro]: - assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" - 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 "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" - 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 + global_interpretation intersection: automaton_intersection_trace + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + "\ c\<^sub>1 c\<^sub>2 pq. (c\<^sub>1 \ fst) pq \ (c\<^sub>2 \ snd) pq" + defines intersect = intersection.combine + by (unfold_locales) (simp del: comp_apply) - lemma dcgaul_language[simp]: - assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" - shows "DGCA.language (dcgaul AA) = \ (DCA.language ` set AA)" - 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 \ \ (DCA.language ` set AA)" - 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 + lemmas intersect_language = intersection.combine_language + lemmas intersect_nodes_finite = intersection.combine_nodes_finite + lemmas intersect_nodes_card = intersection.combine_nodes_card - definition dcaul :: "('label, 'state) dca list \ ('label, 'state list degen) dca" where - "dcaul = dgcad \ dcgaul" + global_interpretation union: automaton_union_trace + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + dgca.dgca dgca.alphabet dgca.initial dgca.transition dgca.rejecting "cogen fins" + "\ c\<^sub>1 c\<^sub>2. [c\<^sub>1 \ fst, c\<^sub>2 \ snd]" + defines union' = union.combine + by unfold_locales auto - lemma dcaul_nodes_finite[intro]: - assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" - 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 "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" - assumes "list_all (finite \ DCA.nodes) AA" - shows "card (DCA.nodes (dcaul AA)) \ max 1 (length AA) * prod_list (map (card \ DCA.nodes) AA)" + lemmas union'_language[simp] = union.combine_language[folded DGCA.language_def] + lemmas union'_nodes_finite = union.combine_nodes_finite[folded DGCA.nodes_def] + lemmas union'_nodes_card = union.combine_nodes_card[folded DGCA.nodes_def] + + global_interpretation intersection_list: automaton_intersection_list_trace + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + "\ cs pp. \ k < length cs. (cs ! k) (pp ! k)" + defines intersect_list = intersection_list.combine + by (unfold_locales) (simp add: comp_def) + + lemmas intersect_list_language = intersection_list.combine_language + lemmas intersect_list_nodes_finite = intersection_list.combine_nodes_finite + lemmas intersect_list_nodes_card = intersection_list.combine_nodes_card + + global_interpretation union_list: automaton_union_list_trace + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + dgca.dgca dgca.alphabet dgca.initial dgca.transition dgca.rejecting "cogen fins" + "\ cs. map (\ k pp. (cs ! k) (pp ! k)) [0 ..< length cs]" + defines union_list' = union_list.combine + by (unfold_locales) (auto simp: cogen_def comp_def) + + lemmas union_list'_language[simp] = union_list.combine_language[folded DGCA.language_def] + lemmas union_list'_nodes_finite = union_list.combine_nodes_finite[folded DGCA.nodes_def] + lemmas union_list'_nodes_card = union_list.combine_nodes_card[folded DGCA.nodes_def] + + abbreviation union where "union A B \ degeneralize (union' A B)" + + lemma union_language[simp]: + assumes "dca.alphabet A = dca.alphabet B" + shows "DCA.language (union A B) = DCA.language A \ DCA.language B" + using assms by simp + lemma union_nodes_finite: + assumes "finite (DCA.nodes A)" "finite (DCA.nodes B)" + shows "finite (DCA.nodes (union A B))" + using union'_nodes_finite assms by simp + lemma union_nodes_card: + assumes "finite (DCA.nodes A)" "finite (DCA.nodes B)" + shows "card (DCA.nodes (union A B)) \ 2 * card (DCA.nodes A) * card (DCA.nodes B)" 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 + have "card (DCA.nodes (union A B)) \ + max 1 (length (dgca.rejecting (union' A B))) * card (DGCA.nodes (union' A B))" + using degeneralize_nodes_card by this + also have "length (dgca.rejecting (union' A B)) = 2" by simp + also have "card (DGCA.nodes (union' A B)) \ card (DCA.nodes A) * card (DCA.nodes B)" + using union'_nodes_card assms by this finally show ?thesis by simp qed - lemma dcaul_language[simp]: + abbreviation union_list where "union_list AA \ degeneralize (union_list' AA)" + + lemma union_list_language[simp]: assumes "\ (dca.alphabet ` set AA) = \ (dca.alphabet ` set AA)" - shows "DCA.language (dcaul AA) = \ (DCA.language ` set AA)" - unfolding dcaul_def using dgcad_language dcgaul_language[OF assms] by auto + shows "DCA.language (union_list AA) = \ (DCA.language ` set AA)" + using assms by simp + lemma union_list_nodes_finite: + assumes "list_all (finite \ DCA.nodes) AA" + shows "finite (DCA.nodes (union_list AA))" + using union_list'_nodes_finite assms by simp + lemma union_list_nodes_card: + assumes "list_all (finite \ DCA.nodes) AA" + shows "card (DCA.nodes (union_list AA)) \ max 1 (length AA) * prod_list (map (card \ DCA.nodes) AA)" + proof - + have "card (DCA.nodes (union_list AA)) \ + max 1 (length (dgca.rejecting (union_list' AA))) * card (DGCA.nodes (union_list' AA))" + using degeneralize_nodes_card by this + also have "length (dgca.rejecting (union_list' AA)) = length AA" by simp + also have "card (DGCA.nodes (union_list' AA)) \ prod_list (map (card \ DCA.nodes) AA)" + using union_list'_nodes_card assms by this + finally show ?thesis by simp + qed 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,25 @@ section \Deterministic Co-Generalized Co-Büchi Automata\ theory DGCA -imports - "DCA" - "../../Basic/Degeneralization" +imports "../Deterministic" begin datatype ('label, 'state) dgca = dgca (alphabet: "'label set") (initial: "'state") (transition: "'label \ 'state \ 'state") (rejecting: "'state pred gen") - global_interpretation dgca: transition_system_initial - "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 + global_interpretation dgca: automaton dgca alphabet initial transition rejecting + defines path = dgca.path and run = dgca.run and reachable = dgca.reachable and nodes = dgca.nodes + by unfold_locales auto + global_interpretation dgca: automaton_trace dgca alphabet initial transition rejecting "cogen fins" + defines language = dgca.language + by standard 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). (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.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.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.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 + abbreviation successors where "successors \ dgca.successors TYPE('label)" 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,25 @@ 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" +imports "../Deterministic" begin datatype ('label, 'state) dra = dra (alphabet: "'label set") (initial: "'state") (transition: "'label \ 'state \ 'state") - (accepting: "'state rabin gen") + (condition: "'state rabin gen") - global_interpretation dra: transition_system_initial - "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 + global_interpretation dra: automaton dra alphabet initial transition condition + defines path = dra.path and run = dra.run and reachable = dra.reachable and nodes = dra.nodes + by unfold_locales auto + global_interpretation dra: automaton_trace dra alphabet initial transition condition "cogen rabin" + defines language = dra.language + by standard 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 + abbreviation successors where "successors \ dra.successors TYPE('label)" 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,33 @@ 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.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.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.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 - (\ (dra.alphabet ` set AA)) - (map dra.initial AA) - (\ 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)] })" + global_interpretation intersection_bc: automaton_intersection_trace + dba.dba dba.alphabet dba.initial dba.transition dba.accepting infs + dca.dca dca.alphabet dca.initial dca.transition dca.rejecting fins + dra.dra dra.alphabet dra.initial dra.transition dra.condition "cogen rabin" + "\ c\<^sub>1 c\<^sub>2. [(c\<^sub>1 \ fst, c\<^sub>2 \ snd)]" + defines intersect_bc = intersection_bc.combine + by (unfold_locales) (simp add: cogen_def rabin_def) - 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 "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" - 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 "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" - 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 "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" - 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 + (* TODO: why are only the constants of the first interpretation folded back correctly? *) + lemmas intersect_bc_language[simp] = intersection_bc.combine_language[folded DCA.language_def DRA.language_def] + lemmas intersect_bc_nodes_finite = intersection_bc.combine_nodes_finite[folded DCA.nodes_def DRA.nodes_def] + lemmas intersect_bc_nodes_card = intersection_bc.combine_nodes_card[folded DCA.nodes_def DRA.nodes_def] - lemma draul_language[simp]: - assumes "\ (dra.alphabet ` set AA) = \ (dra.alphabet ` set AA)" - shows "DRA.language (draul AA) = \ (DRA.language ` set AA)" - 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 \ \ (DRA.language ` set AA)" - 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 + (* TODO: are there some statements about the rabin constant hidden in here? + same for gen/cogen, also in other combinations, shouldn't have to unfold those *) + global_interpretation union_list: automaton_union_list_trace + dra.dra dra.alphabet dra.initial dra.transition dra.condition "cogen rabin" + dra.dra dra.alphabet dra.initial dra.transition dra.condition "cogen rabin" + "\ cs. do { k \ [0 ..< length cs]; (f, g) \ cs ! k; [(\ pp. f (pp ! k), \ pp. g (pp ! k))] }" + defines union_list = union_list.combine + by (unfold_locales) (auto simp: cogen_def rabin_def comp_def split_beta) + + lemmas union_list_language = union_list.combine_language + lemmas union_list_nodes_finite = union_list.combine_nodes_finite + lemmas union_list_nodes_card = union_list.combine_nodes_card 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") (transitione: "('state \ 'label \ 'state) set") - (acceptinge: "('state set \ 'state set) list") + (conditione: "('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 \ (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}" + (conditione A\<^sub>1, conditione 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" "(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" + "(conditione, conditione) \ \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" "(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" + "(conditione A, d) \ \\S\<^sub>1\ set_rel \\<^sub>r \S\<^sub>1\ set_rel\ list_rel" + "(d, conditione 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" "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" + "conditione ::\<^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") (transitionei: "('state \ 'label \ 'state) list") - (acceptingei: "('state list \ 'state list) list") + (conditionei: "('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 \ (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}" + (conditionei A\<^sub>1, conditionei 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" "(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" + "(conditionei, conditionei) \ \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 \ (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}" + (conditionei A\<^sub>1, conditione 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" "(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" + "(conditionei, conditione) \ \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 (transitionei A)) (map (map_prod set set) (acceptingei A))" + (set (transitionei A)) (map (map_prod set set) (conditionei 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 (transitionei Ai)) (map (map_prod set set) (acceptingei Ai))" unfolding draei_drae_def by rule + (set (transitionei Ai)) (map (map_prod set set) (conditionei Ai))" unfolding draei_drae_def by rule have 3: "id A = drae (id (alphabete A)) (initiale A) - (id (transitione A)) (map (map_prod id id) (acceptinge A))" by simp + (id (transitione A)) (map (map_prod id id) (conditione 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) (transition A)) - (map (\ (P, Q). (Set.filter P (nodes A), Set.filter Q (nodes A))) (accepting A))" + (map (\ (P, Q). (Set.filter P (nodes A), Set.filter Q (nodes A))) (condition A))" definition drae_dra where "drae_dra A \ dra (alphabete A) (initiale A) - (succs (transitione A)) (map (\ (I, F). (\ p. p \ I, \ p. p \ F)) (acceptinge A))" + (succs (transitione A)) (map (\ (I, F). (\ p. p \ I, \ p. p \ F)) (conditione 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) (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) (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) (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)) (transition A)) - (map (\ (P, Q). (\ p. p \ Set.filter P (nodes A), \ p. p \ Set.filter Q (nodes A))) (accepting A))" + (map (\ (P, Q). (\ p. p \ Set.filter P (nodes A), \ p. p \ Set.filter Q (nodes A))) (condition A))" unfolding drae_dra_def dra_drae_def by auto - also have "(\, dra (alphabet A) (initial A) (id (transition A)) (map id (accepting A))) \ + also have "(\, dra (alphabet A) (initial A) (id (transition A)) (map id (condition 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 (transition A)) (map id (accepting A)) = id A" by simp + also have "dra (alphabet A) (initial A) (id (transition A)) (map id (condition 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") (transitioni: "'label \ 'state \ 'state") - (acceptingi: "'state rabin gen") + (conditioni: "'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 \ (transitioni A\<^sub>1, transitioni A\<^sub>2) \ L \ S \ S \ - (acceptingi A\<^sub>1, acceptingi A\<^sub>2) \ \rabin_rel S\ list_rel}" + (conditioni A\<^sub>1, conditioni 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" "(transitioni, transitioni) \ \L, S\ drai_rel \ L \ S \ S" - "(acceptingi, acceptingi) \ \L, S\ drai_rel \ \rabin_rel S\ list_rel" + "(conditioni, conditioni) \ \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 \ (transitioni A\<^sub>1, transition A\<^sub>2) \ L \ S \ S \ - (acceptingi A\<^sub>1, accepting A\<^sub>2) \ \rabin_rel S\ list_rel}" + (conditioni A\<^sub>1, condition 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" "(transitioni, transition) \ \L, S\ drai_dra_rel \ L \ S \ S" - "(acceptingi, accepting) \ \L, S\ drai_dra_rel \ \rabin_rel S\ list_rel" + "(conditioni, condition) \ \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) (transitioni A) (acceptingi A)" + "drai_dra A \ dra (set (alphabeti A)) (initiali A) (transitioni A) (conditioni 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) (transitioni Ai) (acceptingi Ai)" + have 2: "drai_dra Ai = dra (set (alphabeti Ai)) (initiali Ai) (transitioni Ai) (conditioni Ai)" unfolding drai_dra_def by rule - have 3: "id A = dra (id (alphabet A)) (initial A) (transition A) (accepting A)" by simp + have 3: "id A = dra (id (alphabet A)) (initial A) (transition A) (condition 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_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 \ (transition A\<^sub>1, transition A\<^sub>2) \ L \ S \ S \ - (accepting A\<^sub>1, accepting A\<^sub>2) \ \rabin_rel S\ list_rel}" + (condition A\<^sub>1, condition 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" "(transition, transition) \ \L, S\ dra_rel \ L \ S \ S" - "(accepting, accepting) \ \L, S\ dra_rel \ \rabin_rel S\ list_rel" + "(condition, condition) \ \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 "(condition A, condition B) \ \rabin_rel S\<^sub>1\ list_rel" by parametricity + also have "(condition B, condition C) \ \rabin_rel S\<^sub>2\ list_rel" by parametricity + finally have 1: "(condition A, condition 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 3: "(condition A, condition C) \ \rabin_rel (S\<^sub>1 O S\<^sub>2)\ list_rel" using 1 2 list_rel_mono by blast 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 enableds_param[param]: "(dra.enableds, dra.enableds) \ \L, S\ dra_rel \ S \ \L\ set_rel" + unfolding dra.enableds_def by simp parametricity + lemma paths_param[param]: "(dra.paths, dra.paths) \ \L, S\ dra_rel \ S \ \\L\ list_rel\ set_rel" + using enableds_param[param_fo] by parametricity + lemma runs_param[param]: "(dra.runs, dra.runs) \ \L, S\ dra_rel \ S \ \\L\ stream_rel\ set_rel" + using enableds_param[param_fo] by 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 + have 1: "reachable A p = (\ w. target A w p) ` dra.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 + show ?thesis unfolding 1 using enableds_param[param_fo] 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 {})" + have 1: "language A = (\ w \ dra.runs A (initial A). + if cogen rabin (condition 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 + unfolding dra.language_def dra.runs_def by auto + show ?thesis unfolding 1 using enableds_param[param_fo] 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)) ` transitione 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)) (conditione 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 ` {transition A a p}) - (map (\ (P, Q). (f ` {p \ nodes A. P p}, f ` {p \ nodes A. Q p})) (accepting A))" + (map (\ (P, Q). (f ` {p \ nodes A. P p}, f ` {p \ nodes A. Q p})) (condition 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 ` {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 = 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 ` {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 ` {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 (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 ` {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 (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))) + (map (\ (P, Q). ((\ x. the (f x)) ` {p \ N. P p}, (\ x. the (f x)) ` {p \ N. Q p})) (condition 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 (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))) + (map (\ (P, Q). ((\ x. the (f x)) ` {p \ N. P p}, (\ x. the (f x)) ` {p \ N. Q p})) (condition 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))) + (map (\ (P, Q). ((\ x. the (f x)) ` {p \ nodes A. P p}, (\ x. the (f x)) ` {p \ nodes A. Q p})) (condition 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))) + (map (\ (P, Q). ((\ x. the (f x)) ` {p \ nodes A. P p}, (\ x. the (f x)) ` {p \ nodes A. Q p})) (condition 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) (transitione (dra_drae A))" using wft_transitions unfolding dra_drae_def drae.sel by this 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) (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/Deterministic.thy b/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy new file mode 100644 --- /dev/null +++ b/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy @@ -0,0 +1,559 @@ +theory Deterministic +imports + "../Transition_Systems/Transition_System" + "../Transition_Systems/Transition_System_Extra" + "../Transition_Systems/Transition_System_Construction" + "../Basic/Degeneralization" +begin + + type_synonym ('label, 'state) trans = "'label \ 'state \ 'state" + + (* TODO: is there a way to be less verbose in these locales? do we need to specify all the types? *) + locale automaton = + fixes automaton :: "'label set \ 'state \ ('label, 'state) trans \ 'condition \ 'automaton" + fixes alphabet :: "'automaton \ 'label set" + fixes initial :: "'automaton \ 'state" + fixes transition :: "'automaton \ ('label, 'state) trans" + fixes condition :: "'automaton \ 'condition" + assumes automaton[simp]: "automaton (alphabet A) (initial A) (transition A) (condition A) = A" + assumes alphabet[simp]: "alphabet (automaton a i t c) = a" + assumes initial[simp]: "initial (automaton a i t c) = i" + assumes transition[simp]: "transition (automaton a i t c) = t" + assumes condition[simp]: "condition (automaton a i t c) = c" + begin + + (* TODO: is there a way to use defines without renaming the constants? *) + sublocale transition_system_initial + "transition A" "\ a p. a \ alphabet A" "\ p. p = initial A" + for A + defines path' = path and run' = run and reachable' = reachable and nodes' = nodes + by this + + 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: run.cases) + show "run A w p" if "w \ streams (alphabet A)" + using that by (coinduction arbitrary: w p) (force elim: streams.cases) + qed + + end + + (* TODO: create analogous locale for DFAs (automaton_target) *) + locale automaton_trace = + automaton automaton alphabet initial transition condition + for automaton :: "'label set \ 'state \ ('label, 'state) trans \ 'condition \ 'automaton" + and alphabet :: "'automaton \ 'label set" + and initial :: "'automaton \ 'state" + and transition :: "'automaton \ ('label, 'state) trans" + and condition :: "'automaton \ 'condition" + + + fixes test :: "'condition \ 'state stream \ bool" + begin + + definition language :: "'automaton \ 'label stream set" where + "language A \ {w. run A w (initial A) \ test (condition A) (trace A w (initial A))}" + + lemma language[intro]: + assumes "run A w (initial A)" "test (condition 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)" "test (condition 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 + + locale automaton_degeneralization = + a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'state pred gen \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'state pred gen" + and automaton\<^sub>2 :: "'label set \ 'state degen \ ('label, 'state degen) trans \ 'state degen pred \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state degen) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen pred" + begin + + definition degeneralize :: "'automaton\<^sub>1 \ 'automaton\<^sub>2" where + "degeneralize A \ automaton\<^sub>2 + (alphabet\<^sub>1 A) + (initial\<^sub>1 A, 0) + (\ a (p, k). (transition\<^sub>1 A a p, count (condition\<^sub>1 A) p k)) + (degen (condition\<^sub>1 A))" + + lemma degeneralize_simps[simp]: + "alphabet\<^sub>2 (degeneralize A) = alphabet\<^sub>1 A" + "initial\<^sub>2 (degeneralize A) = (initial\<^sub>1 A, 0)" + "transition\<^sub>2 (degeneralize A) a (p, k) = (transition\<^sub>1 A a p, count (condition\<^sub>1 A) p k)" + "condition\<^sub>2 (degeneralize A) = degen (condition\<^sub>1 A)" + unfolding degeneralize_def by auto + + lemma degeneralize_target[simp]: "b.target (degeneralize A) w (p, k) = + (a.target A w p, fold (count (condition\<^sub>1 A)) (butlast (p # a.states A w p)) k)" + by (induct w arbitrary: p k) (auto) + lemma degeneralize_states[simp]: "b.states (degeneralize A) w (p, k) = + a.states A w p || scan (count (condition\<^sub>1 A)) (p # a.states A w p) k" + by (induct w arbitrary: p k) (auto) + lemma degeneralize_trace[simp]: "b.trace (degeneralize A) w (p, k) = + a.trace A w p ||| sscan (count (condition\<^sub>1 A)) (p ## a.trace A w p) k" + by (coinduction arbitrary: w p k) (auto) + + lemma degeneralize_path[iff]: "b.path (degeneralize A) w (p, k) \ a.path A w p" + unfolding a.path_alt_def b.path_alt_def by simp + lemma degeneralize_run[iff]: "b.run (degeneralize A) w (p, k) \ a.run A w p" + unfolding a.run_alt_def b.run_alt_def by simp + + lemma degeneralize_reachable_fst[simp]: "fst ` b.reachable (degeneralize A) (p, k) = a.reachable A p" + unfolding a.reachable_alt_def b.reachable_alt_def image_def by simp + lemma degeneralize_reachable_snd_empty[simp]: + assumes "condition\<^sub>1 A = []" + shows "snd ` b.reachable (degeneralize A) (p, k) = {k}" + proof - + have "snd ql = k" if "ql \ b.reachable (degeneralize A) (p, k)" for ql + using that assms by induct auto + then show ?thesis by auto + qed + lemma degeneralize_reachable_empty[simp]: + assumes "condition\<^sub>1 A = []" + shows "b.reachable (degeneralize A) (p, k) = a.reachable A p \ {k}" + using degeneralize_reachable_fst degeneralize_reachable_snd_empty assms + by (metis prod_singleton(2)) + lemma degeneralize_reachable_snd: + "snd ` b.reachable (degeneralize A) (p, k) \ insert k {0 ..< length (condition\<^sub>1 A)}" + by (cases "condition\<^sub>1 A = []") (auto) + lemma degeneralize_reachable: + "b.reachable (degeneralize A) (p, k) \ a.reachable A p \ insert k {0 ..< length (condition\<^sub>1 A)}" + by (cases "condition\<^sub>1 A = []") (auto 0 3) + + lemma degeneralize_nodes_fst[simp]: "fst ` b.nodes (degeneralize A) = a.nodes A" + unfolding a.nodes_alt_def b.nodes_alt_def by simp + lemma degeneralize_nodes_snd_empty: + assumes "condition\<^sub>1 A = []" + shows "snd ` b.nodes (degeneralize A) = {0}" + using assms unfolding b.nodes_alt_def by auto + lemma degeneralize_nodes_empty: + assumes "condition\<^sub>1 A = []" + shows "b.nodes (degeneralize A) = a.nodes A \ {0}" + using assms unfolding b.nodes_alt_def by auto + lemma degeneralize_nodes_snd: + "snd ` b.nodes (degeneralize A) \ insert 0 {0 ..< length (condition\<^sub>1 A)}" + using degeneralize_reachable_snd unfolding b.nodes_alt_def by auto + lemma degeneralize_nodes: + "b.nodes (degeneralize A) \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" + using degeneralize_reachable unfolding a.nodes_alt_def b.nodes_alt_def by simp + + lemma degeneralize_nodes_finite[iff]: "finite (b.nodes (degeneralize A)) \ finite (a.nodes A)" + proof + show "finite (a.nodes A)" if "finite (b.nodes (degeneralize A))" + using that by (auto simp flip: degeneralize_nodes_fst) + show "finite (b.nodes (degeneralize A))" if "finite (a.nodes A)" + using degeneralize_nodes that finite_subset by fastforce + qed + lemma degeneralize_nodes_card: "card (b.nodes (degeneralize A)) \ + max 1 (length (condition\<^sub>1 A)) * card (a.nodes A)" + proof (cases "finite (a.nodes A)") + case True + have "card (b.nodes (degeneralize A)) \ card (a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)})" + using degeneralize_nodes True by (blast intro: card_mono) + also have "\ = card (insert 0 {0 ..< length (condition\<^sub>1 A)}) * card (a.nodes A)" + unfolding card_cartesian_product by simp + also have "card (insert 0 {0 ..< length (condition\<^sub>1 A)}) = max 1 (length (condition\<^sub>1 A))" + by (simp add: card_insert_if Suc_leI max_absorb2) + finally show ?thesis by this + next + case False + then have "card (a.nodes A) = 0" "card (b.nodes (degeneralize A)) = 0" by auto + then show ?thesis by simp + qed + + end + + locale automaton_degeneralization_trace = + automaton_degeneralization + automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'state pred gen \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'state pred gen" + and test\<^sub>1 :: "'state pred gen \ 'state stream \ bool" + and automaton\<^sub>2 :: "'label set \ 'state degen \ ('label, 'state degen) trans \ 'state degen pred \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state degen) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'state degen pred" + and test\<^sub>2 :: "'state degen pred \ 'state degen stream \ bool" + + + assumes test[iff]: "test\<^sub>2 (degen cs) (r ||| sscan (count cs) (p ## r) k) \ test\<^sub>1 cs r" + begin + + lemma degeneralize_language[simp]: "b.language (degeneralize A) = a.language A" + by (force simp del: sscan_scons) + + end + + locale automaton_combination = + a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + + c: automaton automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 + for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" + and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + and automaton\<^sub>3 :: "'label set \ 'state\<^sub>1 \ 'state\<^sub>2 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ + 'condition\<^sub>3 \ 'automaton\<^sub>3" + and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" + and initial\<^sub>3 :: "'automaton\<^sub>3 \ 'state\<^sub>1 \ 'state\<^sub>2" + and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" + and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" + + + fixes condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + begin + + definition combine :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where + "combine A B \ automaton\<^sub>3 + (alphabet\<^sub>1 A \ alphabet\<^sub>2 B) + (initial\<^sub>1 A, initial\<^sub>2 B) + (\ a (p, q). (transition\<^sub>1 A a p, transition\<^sub>2 B a q)) + (condition (condition\<^sub>1 A) (condition\<^sub>2 B))" + + lemma combine_simps[simp]: + "alphabet\<^sub>3 (combine A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" + "initial\<^sub>3 (combine A B) = (initial\<^sub>1 A, initial\<^sub>2 B)" + "transition\<^sub>3 (combine A B) a (p, q) = (transition\<^sub>1 A a p, transition\<^sub>2 B a q)" + "condition\<^sub>3 (combine A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" + unfolding combine_def by auto + + lemma combine_target[simp]: "c.target (combine A B) w (p, q) = (a.target A w p, b.target B w q)" + by (induct w arbitrary: p q) (auto) + lemma combine_states[simp]: "c.states (combine A B) w (p, q) = a.states A w p || b.states B w q" + by (induct w arbitrary: p q) (auto) + lemma combine_trace[simp]: "c.trace (combine A B) w (p, q) = a.trace A w p ||| b.trace B w q" + by (coinduction arbitrary: w p q) (auto) + + lemma combine_path[iff]: "c.path (combine A B) w (p, q) \ a.path A w p \ b.path B w q" + unfolding a.path_alt_def b.path_alt_def c.path_alt_def by simp + lemma combine_run[iff]: "c.run (combine A B) w (p, q) \ a.run A w p \ b.run B w q" + unfolding a.run_alt_def b.run_alt_def c.run_alt_def by simp + + lemma combine_reachable[simp]: "c.reachable (combine A B) (p, q) \ a.reachable A p \ b.reachable B q" + unfolding c.reachable_alt_def by auto + lemma combine_nodes[simp]: "c.nodes (combine A B) \ a.nodes A \ b.nodes B" + unfolding a.nodes_alt_def b.nodes_alt_def c.nodes_alt_def by auto + lemma combine_reachable_fst[simp]: + assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" + shows "fst ` c.reachable (combine A B) (p, q) = a.reachable A p" + using assms + unfolding a.reachable_alt_def a.path_alt_def + unfolding b.reachable_alt_def b.path_alt_def + unfolding c.reachable_alt_def c.path_alt_def + by auto force + lemma combine_reachable_snd[simp]: + assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" + shows "snd ` c.reachable (combine A B) (p, q) = b.reachable B q" + using assms + unfolding a.reachable_alt_def a.path_alt_def + unfolding b.reachable_alt_def b.path_alt_def + unfolding c.reachable_alt_def c.path_alt_def + by auto force + lemma combine_nodes_fst[simp]: + assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" + shows "fst ` c.nodes (combine A B) = a.nodes A" + using assms combine_reachable_fst + unfolding a.nodes_alt_def b.nodes_alt_def c.nodes_alt_def + by fastforce + lemma combine_nodes_snd[simp]: + assumes "alphabet\<^sub>1 A \ alphabet\<^sub>2 B" + shows "snd ` c.nodes (combine A B) = b.nodes B" + using assms combine_reachable_snd + unfolding a.nodes_alt_def b.nodes_alt_def c.nodes_alt_def + by fastforce + + lemma combine_nodes_finite[intro]: + assumes "finite (a.nodes A)" "finite (b.nodes B)" + shows "finite (c.nodes (combine A B))" + proof (rule finite_subset) + show "c.nodes (combine A B) \ a.nodes A \ b.nodes B" using combine_nodes by this + show "finite (a.nodes A \ b.nodes B)" using assms by simp + qed + lemma combine_nodes_finite_strong[iff]: + assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" + shows "finite (c.nodes (combine A B)) \ finite (a.nodes A) \ finite (b.nodes B)" + proof safe + show "finite (a.nodes A)" if "finite (c.nodes (combine A B))" + using combine_nodes_fst assms that by (metis finite_imageI equalityD1) + show "finite (b.nodes B)" if "finite (c.nodes (combine A B))" + using combine_nodes_snd assms that by (metis finite_imageI equalityD2) + show "finite (c.nodes (combine A B))" if "finite (a.nodes A)" "finite (b.nodes B)" + using that by rule + qed + lemma combine_nodes_card[intro]: + assumes "finite (a.nodes A)" "finite (b.nodes B)" + shows "card (c.nodes (combine A B)) \ card (a.nodes A) * card (b.nodes B)" + proof - + have "card (c.nodes (combine A B)) \ card (a.nodes A \ b.nodes B)" + proof (rule card_mono) + show "finite (a.nodes A \ b.nodes B)" using assms by simp + show "c.nodes (combine A B) \ a.nodes A \ b.nodes B" using combine_nodes by this + qed + also have "\ = card (a.nodes A) * card (b.nodes B)" using card_cartesian_product by this + finally show ?thesis by this + qed + lemma combine_nodes_card_strong[intro]: + assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" + shows "card (c.nodes (combine A B)) \ card (a.nodes A) * card (b.nodes B)" + proof (cases "finite (a.nodes A) \ finite (b.nodes B)") + case True + show ?thesis using True by auto + next + case False + have 1: "card (c.nodes (combine A B)) = 0" using False assms by simp + have 2: "card (a.nodes A) * card (b.nodes B) = 0" using False by auto + show ?thesis using 1 2 by simp + qed + + end + + locale automaton_intersection_trace = + automaton_combination + automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + + c: automaton_trace automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 + for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" + and test\<^sub>1 :: "'condition\<^sub>1 \ 'state\<^sub>1 stream \ bool" + and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + and test\<^sub>2 :: "'condition\<^sub>2 \ 'state\<^sub>2 stream \ bool" + and automaton\<^sub>3 :: "'label set \ 'state\<^sub>1 \ 'state\<^sub>2 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ + 'condition\<^sub>3 \ 'automaton\<^sub>3" + and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" + and initial\<^sub>3 :: "'automaton\<^sub>3 \ 'state\<^sub>1 \ 'state\<^sub>2" + and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" + and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" + and test\<^sub>3 :: "'condition\<^sub>3 \ ('state\<^sub>1 \ 'state\<^sub>2) stream \ bool" + and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + + + assumes test[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) (u ||| v) \ test\<^sub>1 c\<^sub>1 u \ test\<^sub>2 c\<^sub>2 v" + begin + + lemma combine_language[simp]: "c.language (combine A B) = a.language A \ b.language B" by force + + end + + locale automaton_union_trace = + automaton_combination + automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + + c: automaton_trace automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 + for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 \ ('label, 'state\<^sub>1) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state\<^sub>1" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state\<^sub>1) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" + and test\<^sub>1 :: "'condition\<^sub>1 \ 'state\<^sub>1 stream \ bool" + and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 \ ('label, 'state\<^sub>2) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state\<^sub>2" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state\<^sub>2) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + and test\<^sub>2 :: "'condition\<^sub>2 \ 'state\<^sub>2 stream \ bool" + and automaton\<^sub>3 :: "'label set \ 'state\<^sub>1 \ 'state\<^sub>2 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans \ + 'condition\<^sub>3 \ 'automaton\<^sub>3" + and alphabet\<^sub>3 :: "'automaton\<^sub>3 \ 'label set" + and initial\<^sub>3 :: "'automaton\<^sub>3 \ 'state\<^sub>1 \ 'state\<^sub>2" + and transition\<^sub>3 :: "'automaton\<^sub>3 \ ('label, 'state\<^sub>1 \ 'state\<^sub>2) trans" + and condition\<^sub>3 :: "'automaton\<^sub>3 \ 'condition\<^sub>3" + and test\<^sub>3 :: "'condition\<^sub>3 \ ('state\<^sub>1 \ 'state\<^sub>2) stream \ bool" + and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + + + assumes test[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) (u ||| v) \ test\<^sub>1 c\<^sub>1 u \ test\<^sub>2 c\<^sub>2 v" + begin + + lemma combine_language[simp]: + assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" + shows "c.language (combine A B) = a.language A \ b.language B" + using assms by (force simp: a.run_alt_def b.run_alt_def) + + end + + (* TODO: complete this in analogy to the pair case *) + locale automaton_combination_list = + a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" + and automaton\<^sub>2 :: "'label set \ 'state list \ ('label, 'state list) trans \ + 'condition\<^sub>2 \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state list" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state list) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + + + fixes condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + begin + + definition combine :: "'automaton\<^sub>1 list \ 'automaton\<^sub>2" where + "combine AA \ automaton\<^sub>2 + (\ (alphabet\<^sub>1 ` set AA)) + (map initial\<^sub>1 AA) + (\ a pp. map2 (\ A p. transition\<^sub>1 A a p) AA pp) + (condition (map condition\<^sub>1 AA))" + + lemma combine_simps[simp]: + "alphabet\<^sub>2 (combine AA) = \ (alphabet\<^sub>1 ` set AA)" + "initial\<^sub>2 (combine AA) = map initial\<^sub>1 AA" + "transition\<^sub>2 (combine AA) a pp = map2 (\ A p. transition\<^sub>1 A a p) AA pp" + "condition\<^sub>2 (combine AA) = condition (map condition\<^sub>1 AA)" + unfolding combine_def by auto + + (* TODO: get rid of indices, express this using stranspose *) + lemma combine_trace_smap: + assumes "length pp = length AA" "k < length AA" + shows "smap (\ pp. pp ! k) (b.trace (combine AA) w pp) = a.trace (AA ! k) w (pp ! k)" + using assms by (coinduction arbitrary: w pp) (force) + lemma combine_nodes_length: + assumes "pp \ b.nodes (combine AA)" + shows "length pp = length AA" + using assms by induct auto + lemma combine_nodes[intro]: + assumes "pp \ b.nodes (combine AA)" "k < length pp" + shows "pp ! k \ a.nodes (AA ! k)" + using assms by induct auto + + lemma combine_nodes_finite[intro]: + assumes "list_all (finite \ a.nodes) AA" + shows "finite (b.nodes (combine AA))" + proof (rule finite_subset) + show "b.nodes (combine AA) \ listset (map a.nodes AA)" + by (force simp: listset_member list_all2_conv_all_nth combine_nodes_length) + have "finite (listset (map a.nodes AA)) \ list_all finite (map a.nodes AA)" + by (rule listset_finite) (auto simp: list_all_iff) + then show "finite (listset (map a.nodes AA))" using assms by (simp add: list.pred_map) + qed + lemma combine_nodes_card: + assumes "list_all (finite \ a.nodes) AA" + shows "card (b.nodes (combine AA)) \ prod_list (map (card \ a.nodes) AA)" + proof - + have "card (b.nodes (combine AA)) \ card (listset (map a.nodes AA))" + proof (rule card_mono) + have "finite (listset (map a.nodes AA)) \ list_all finite (map a.nodes AA)" + by (rule listset_finite) (auto simp: list_all_iff) + then show "finite (listset (map a.nodes AA))" using assms by (simp add: list.pred_map) + show "b.nodes (combine AA) \ listset (map a.nodes AA)" + by (force simp: listset_member list_all2_conv_all_nth combine_nodes_length) + qed + also have "\ = prod_list (map (card \ a.nodes) AA)" by simp + finally show ?thesis by this + qed + + end + + locale automaton_intersection_list_trace = + automaton_combination_list + automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 condition + + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" + and test\<^sub>1 :: "'condition\<^sub>1 \ 'state stream \ bool" + and automaton\<^sub>2 :: "'label set \ 'state list \ ('label, 'state list) trans \ + 'condition\<^sub>2 \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state list" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state list) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + and test\<^sub>2 :: "'condition\<^sub>2 \ 'state list stream \ bool" + and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + + + assumes test[iff]: "test\<^sub>2 (condition cs) w \ + (\ k < length cs. test\<^sub>1 (cs ! k) (smap (\ pp. pp ! k) w))" + begin + + lemma combine_language[simp]: "b.language (combine AA) = \ (a.language ` set AA)" + unfolding a.language_def b.language_def + unfolding a.run_alt_def b.run_alt_def + by (auto simp: combine_trace_smap iff: in_set_conv_nth) + + end + + locale automaton_union_list_trace = + automaton_combination_list + automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 condition + + a: automaton_trace automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + + b: automaton_trace automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + for automaton\<^sub>1 :: "'label set \ 'state \ ('label, 'state) trans \ 'condition\<^sub>1 \ 'automaton\<^sub>1" + and alphabet\<^sub>1 :: "'automaton\<^sub>1 \ 'label set" + and initial\<^sub>1 :: "'automaton\<^sub>1 \ 'state" + and transition\<^sub>1 :: "'automaton\<^sub>1 \ ('label, 'state) trans" + and condition\<^sub>1 :: "'automaton\<^sub>1 \ 'condition\<^sub>1" + and test\<^sub>1 :: "'condition\<^sub>1 \ 'state stream \ bool" + and automaton\<^sub>2 :: "'label set \ 'state list \ ('label, 'state list) trans \ + 'condition\<^sub>2 \ 'automaton\<^sub>2" + and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" + and initial\<^sub>2 :: "'automaton\<^sub>2 \ 'state list" + and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, 'state list) trans" + and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + and test\<^sub>2 :: "'condition\<^sub>2 \ 'state list stream \ bool" + and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + + + assumes test[iff]: "test\<^sub>2 (condition cs) w \ + (\ k < length cs. test\<^sub>1 (cs ! k) (smap (\ pp. pp ! k) w))" + begin + + lemma combine_language[simp]: + assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" + shows "b.language (combine AA) = \ (a.language ` set AA)" + using assms + unfolding a.language_def b.language_def + unfolding a.run_alt_def b.run_alt_def + by (auto simp: combine_trace_smap iff: in_set_conv_nth) (metis INT_subset_iff in_set_conv_nth) + + end + +end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Basic/Basic.thy b/thys/Transition_Systems_and_Automata/Basic/Basic.thy --- a/thys/Transition_Systems_and_Automata/Basic/Basic.thy +++ b/thys/Transition_Systems_and_Automata/Basic/Basic.thy @@ -1,31 +1,35 @@ section \Basics\ theory Basic imports Main begin subsection \Miscellaneous\ abbreviation (input) "const x \ \ _. x" lemmas [simp] = map_prod.id map_prod.comp[symmetric] lemma prod_UNIV[iff]: "A \ B = UNIV \ A = UNIV \ B = UNIV" by auto + lemma prod_singleton: + "fst ` A = {x} \ A = fst ` A \ snd ` A" + "snd ` A = {y} \ A = fst ` A \ snd ` A" + by force+ lemma infinite_subset[trans]: "infinite A \ A \ B \ infinite B" using infinite_super by this lemma finite_subset[trans]: "A \ B \ finite B \ finite A" using finite_subset by this declare infinite_coinduct[case_names infinite, coinduct pred: infinite] lemma infinite_psubset_coinduct[case_names infinite, consumes 1]: assumes "R A" assumes "\ A. R A \ \ B \ A. R B" shows "infinite A" proof show "False" if "finite A" using that assms by (induct rule: finite_psubset_induct) (auto) qed (* TODO: why are there two copies of this theorem? *) thm inj_on_subset subset_inj_on lemma inj_inj_on[dest]: "inj f \ inj_on f S" using inj_on_subset by auto end diff --git a/thys/Transition_Systems_and_Automata/Basic/Degeneralization.thy b/thys/Transition_Systems_and_Automata/Basic/Degeneralization.thy --- a/thys/Transition_Systems_and_Automata/Basic/Degeneralization.thy +++ b/thys/Transition_Systems_and_Automata/Basic/Degeneralization.thy @@ -1,224 +1,230 @@ theory Degeneralization imports Acceptance Sequence_Zip begin type_synonym 'a degen = "'a \ nat" definition degen :: "'a pred gen \ 'a degen pred" where "degen cs \ \ (a, k). k \ length cs \ (cs ! k) a" lemma degen_simps[iff]: "degen cs (a, k) \ k \ length cs \ (cs ! k) a" unfolding degen_def by simp definition count :: "'a pred gen \ 'a \ nat \ nat" where - "count cs a k \ if k < length cs then if (cs ! k) a then Suc k mod length cs else k else 0" + "count cs a k \ + if k < length cs + then if (cs ! k) a then Suc k mod length cs else k + else if cs = [] then k else 0" - lemma count_empty[simp]: "count [] a k = 0" unfolding count_def by simp - lemma count_nonempty[simp]: - assumes "cs \ []" - shows "count cs a k < length cs" - using assms unfolding count_def by simp + lemma count_empty[simp]: "count [] a k = k" unfolding count_def by simp + lemma count_nonempty[simp]: "cs \ [] \ count cs a k < length cs" unfolding count_def by simp lemma count_constant_1: assumes "k < length cs" assumes "\ a. a \ set w \ \ (cs ! k) a" shows "fold (count cs) w k = k" using assms unfolding count_def by (induct w) (auto) lemma count_constant_2: assumes "k < length cs" assumes "\ a. a \ set (w || k # scan (count cs) w k) \ \ degen cs a" shows "fold (count cs) w k = k" using assms unfolding count_def by (induct w) (auto) + lemma count_step: + assumes "k < length cs" + assumes "(cs ! k) a" + shows "count cs a k = Suc k mod length cs" + using assms unfolding count_def by simp lemma degen_skip_condition: assumes "k < length cs" assumes "infs (degen cs) (w ||| k ## sscan (count cs) w k)" obtains u a v where "w = u @- a ## v" "fold (count cs) u k = k" "(cs ! k) a" proof - have 1: "Collect (degen cs) \ sset (w ||| k ## sscan (count cs) w k) \ {}" using infs_any assms(2) by auto obtain ys x zs where 2: "w ||| k ## sscan (count cs) w k = ys @- x ## zs" "Collect (degen cs) \ set ys = {}" "x \ Collect (degen cs)" using split_stream_first 1 by this define u where "u \ stake (length ys) w" define a where "a \ w !! length ys" define v where "v \ sdrop (Suc (length ys)) w" have "ys = stake (length ys) (w ||| k ## sscan (count cs) w k)" using shift_eq 2(1) by auto also have "\ = stake (length ys) w || stake (length ys) (k ## sscan (count cs) w k)" by simp also have "\ = take (length ys) u || take (length ys) (k # scan (count cs) u k)" unfolding u_def using append_eq_conv_conj length_stake length_zip stream.sel using sscan_stake stake.simps(2) stake_Suc stake_szip take_stake by metis also have "\ = take (length ys) (u || k # scan (count cs) u k)" using take_zip by rule also have "\ = u || k # scan (count cs) u k" unfolding u_def by simp finally have 3: "ys = u || k # scan (count cs) u k" by this have "x = (w ||| k ## sscan (count cs) w k) !! length ys" unfolding 2(1) by simp also have "\ = (w !! length ys, (k ## sscan (count cs) w k) !! length ys)" by simp also have "\ = (a, fold (count cs) u k)" unfolding u_def a_def by simp finally have 4: "x = (a, fold (count cs) u k)" by this have 5: "fold (count cs) u k = k" using count_constant_2 assms(1) 2(2) unfolding 3 by blast show ?thesis proof show "w = u @- a ## v" unfolding u_def a_def v_def using id_stake_snth_sdrop by this show "fold (count cs) u k = k" using 5 by this show "(cs ! k) a" using assms(1) 2(3) unfolding 4 5 by simp qed qed lemma degen_skip_arbitrary: assumes "k < length cs" "l < length cs" assumes "infs (degen cs) (w ||| k ## sscan (count cs) w k)" obtains u v where "w = u @- v" "fold (count cs) u k = l" using assms proof (induct "nat ((int l - int k) mod length cs)" arbitrary: l thesis) case (0) have 1: "length cs > 0" using assms(1) by auto have 2: "(int l - int k) mod length cs = 0" using 0(1) 1 by (auto intro: antisym) have 3: "int l mod length cs = int k mod length cs" using mod_eq_dvd_iff 2 by force have 4: "k = l" using 0(3, 4) 3 by simp show ?case proof (rule 0(2)) show "w = [] @- w" by simp show "fold (count cs) [] k = l" using 4 by simp qed next case (Suc n) have 1: "length cs > 0" using assms(1) by auto define l' where "l' = nat ((int l - 1) mod length cs)" obtain u v where 2: "w = u @- v" "fold (count cs) u k = l'" proof (rule Suc(1)) have 2: "Suc n < length cs" using nat_less_iff Suc(2) 1 by simp have "n = nat (int n)" by simp also have "int n = (int (Suc n) - 1) mod length cs" using 2 by simp also have "\ = (int l - int k - 1) mod length cs" using Suc(2) by (simp add: mod_simps) also have "\ = (int l - 1 - int k) mod length cs" by (simp add: algebra_simps) also have "\ = (int l' - int k) mod length cs" using l'_def 1 by (simp add: mod_simps) finally show "n = nat ((int l' - int k) mod length cs)" by this show "k < length cs" using Suc(4) by this show "l' < length cs" using nat_less_iff l'_def 1 by simp show "infs (degen cs) (w ||| k ## sscan (count cs) w k)" using Suc(6) by this qed have 3: "l' < length cs" using nat_less_iff l'_def 1 by simp have 4: "v ||| l' ## sscan (count cs) v l' = sdrop (length u) (w ||| k ## sscan (count cs) w k)" using 2 eq_scons eq_shift by (metis sdrop.simps(2) sdrop_simps sdrop_szip sscan_scons_snth sscan_sdrop stream.sel(2)) have 5: "infs (degen cs) (v ||| l' ## sscan (count cs) v l')" using Suc(6) unfolding 4 by blast obtain vu a vv where 6: "v = vu @- a ## vv" "fold (count cs) vu l' = l'" "(cs ! l') a" using degen_skip_condition 3 5 by this have "l = nat (int l)" by simp also have "int l = int l mod length cs" using Suc(5) by simp also have "\ = int (Suc l') mod length cs" using l'_def 1 by (simp add: mod_simps) finally have 7: "l = Suc l' mod length cs" using nat_mod_as_int by metis show ?case proof (rule Suc(3)) show "w = (u @ vu @ [a]) @- vv" unfolding 2(1) 6(1) by simp - show "fold (count cs) (u @ vu @ [a]) k = l" using 2(2) 3 6(2, 3) 7 unfolding count_def by simp + show "fold (count cs) (u @ vu @ [a]) k = l" using 2(2) 3 6(2, 3) 7 count_step by simp qed qed lemma degen_skip_arbitrary_condition: assumes "l < length cs" assumes "infs (degen cs) (w ||| k ## sscan (count cs) w k)" obtains u a v where "w = u @- a ## v" "fold (count cs) u k = l" "(cs ! l) a" proof - - have 1: "count cs (shd w) k < length cs" using mod_Suc assms(1) unfolding count_def by auto + have 0: "cs \ []" using assms(1) by auto + have 1: "count cs (shd w) k < length cs" using 0 by simp have 2: "infs (degen cs) (stl w ||| count cs (shd w) k ## sscan (count cs) (stl w) (count cs (shd w) k))" using assms(2) by (metis alw.cases sscan.code stream.sel(2) szip.simps(2)) obtain u v where 3: "stl w = u @- v" "fold (count cs) u (count cs (shd w) k) = l" using degen_skip_arbitrary 1 assms(1) 2 by this have 4: "v ||| l ## sscan (count cs) v l = sdrop (length u) (stl w ||| count cs (shd w) k ## sscan (count cs) (stl w) (count cs (shd w) k))" using 3 eq_scons eq_shift by (metis sdrop.simps(2) sdrop_simps sdrop_szip sscan_scons_snth sscan_sdrop stream.sel(2)) have 5: "infs (degen cs) (v ||| l ## sscan (count cs) v l)" using 2 unfolding 4 by blast obtain vu a vv where 6: "v = vu @- a ## vv" "fold (count cs) vu l = l" "(cs ! l) a" using degen_skip_condition assms(1) 5 by this show ?thesis proof show "w = (shd w # u @ vu) @- a ## vv" using 3(1) 6(1) by (simp add: eq_scons) show "fold (count cs) (shd w # u @ vu) k = l" using 3(2) 6(2) by simp show "(cs ! l) a" using 6(3) by this qed qed lemma gen_degen_step: assumes "gen infs cs w" obtains u a v where "w = u @- a ## v" "degen cs (a, fold (count cs) u k)" proof (cases "k < length cs") case True have 1: "infs (cs ! k) w" using assms True by auto have 2: "{a. (cs ! k) a} \ sset w \ {}" using infs_any 1 by auto obtain u a v where 3: "w = u @- a ## v" "{a. (cs ! k) a} \ set u = {}" "a \ {a. (cs ! k) a}" using split_stream_first 2 by this have 4: "fold (count cs) u k = k" using count_constant_1 True 3(2) by auto show ?thesis using 3(1, 3) 4 that by simp next case False show ?thesis proof show "w = [] @- shd w ## stl w" by simp show "degen cs (shd w, fold (count cs) [] k)" using False by simp qed qed lemma degen_infs[iff]: "infs (degen cs) (w ||| k ## sscan (count cs) w k) \ gen infs cs w" proof show "gen infs cs w" if "infs (degen cs) (w ||| k ## sscan (count cs) w k)" proof fix c assume 1: "c \ set cs" obtain l where 2: "c = cs ! l" "l < length cs" using in_set_conv_nth 1 by metis show "infs c w" using that unfolding 2(1) proof (coinduction arbitrary: w k rule: infs_set_coinduct) case (infs_set w k) obtain u a v where 3: "w = u @- a ## v" "(cs ! l) a" using degen_skip_arbitrary_condition 2(2) infs_set by this let ?k = "fold (count cs) u k" let ?l = "fold (count cs) (u @ [a]) k" have 4: "a ## v ||| ?k ## sscan (count cs) (a ## v) ?k = sdrop (length u) (w ||| k ## sscan (count cs) w k)" using 3(1) eq_shift scons_eq by (metis sdrop_simps(1) sdrop_stl sdrop_szip sscan_scons_snth sscan_sdrop stream.sel(2)) have 5: "infs (degen cs) (a ## v ||| ?k ## sscan (count cs) (a ## v) ?k)" using infs_set unfolding 4 by blast show ?case proof (intro exI conjI bexI) show "w = (u @ [a]) @- v" "(cs ! l) a" "a \ set (u @ [a])" "v = v" using 3 by auto show "infs (degen cs) (v ||| ?l ## sscan (count cs) v ?l)" using 5 by simp qed qed qed show "infs (degen cs) (w ||| k ## sscan (count cs) w k)" if "gen infs cs w" using that proof (coinduction arbitrary: w k rule: infs_set_coinduct) case (infs_set w k) obtain u a v where 1: "w = u @- a ## v" "degen cs (a, fold (count cs) u k)" using gen_degen_step infs_set by this let ?u = "u @ [a] || k # scan (count cs) u k" let ?l = "fold (count cs) (u @ [a]) k" show ?case proof (intro exI conjI bexI) have "w ||| k ## sscan (count cs) w k = (u @ [a]) @- v ||| k ## scan (count cs) u k @- ?l ## sscan (count cs) v ?l" unfolding 1(1) by simp also have "\ = ?u @- (v ||| ?l ## sscan (count cs) v ?l)" by (metis length_Cons length_append_singleton scan_length shift.simps(2) szip_shift) finally show "w ||| k ## sscan (count cs) w k = ?u @- (v ||| ?l ## sscan (count cs) v ?l)" by this show "degen cs (a, fold (count cs) u k)" using 1(2) by this have "(a, fold (count cs) u k) = (last (u @ [a]), last (k # scan (count cs) u k))" unfolding scan_last by simp also have "\ = last ?u" by (simp add: zip_eq_Nil_iff) also have "\ \ set ?u" by (fastforce intro: last_in_set simp: zip_eq_Nil_iff) finally show "(a, fold (count cs) u k) \ set ?u" by this show "v ||| ?l ## sscan (count cs) v ?l = v ||| ?l ## sscan (count cs) v ?l" by rule show "gen infs cs v" using infs_set unfolding 1(1) by auto qed qed qed -end \ No newline at end of file +end diff --git a/thys/Transition_Systems_and_Automata/Basic/Sequence.thy b/thys/Transition_Systems_and_Automata/Basic/Sequence.thy --- a/thys/Transition_Systems_and_Automata/Basic/Sequence.thy +++ b/thys/Transition_Systems_and_Automata/Basic/Sequence.thy @@ -1,436 +1,441 @@ section \Finite and Infinite Sequences\ theory Sequence imports "Basic" "HOL-Library.Stream" "HOL-Library.Monad_Syntax" begin subsection \List Basics\ declare upt_Suc[simp del] declare last.simps[simp del] declare butlast.simps[simp del] declare Cons_nth_drop_Suc[simp] declare list.pred_True[simp] lemma list_pred_cases: assumes "list_all P xs" obtains (nil) "xs = []" | (cons) y ys where "xs = y # ys" "P y" "list_all P ys" using assms by (cases xs) (auto) lemma lists_iff_set: "w \ lists A \ set w \ A" by auto lemma fold_const: "fold const xs a = last (a # xs)" by (induct xs arbitrary: a) (auto simp: last.simps) lemma take_Suc: "take (Suc n) xs = (if xs = [] then [] else hd xs # take n (tl xs))" by (simp add: take_Suc) lemma bind_map[simp]: "map f xs \ g = xs \ g \ f" unfolding List.bind_def by simp + lemma ball_bind[iff]: "Ball (set (xs \ f)) P \ (\ x \ set xs. \ y \ set (f x). P y)" + unfolding set_list_bind by simp + lemma bex_bind[iff]: "Bex (set (xs \ f)) P \ (\ x \ set xs. \ y \ set (f x). P y)" + unfolding set_list_bind by simp + lemma listset_member: "ys \ listset XS \ list_all2 (\) ys XS" by (induct XS arbitrary: ys) (auto simp: set_Cons_def list_all2_Cons2) lemma listset_empty[iff]: "listset XS = {} \ \ list_all (\ A. A \ {}) XS" by (induct XS) (auto simp: set_Cons_def) lemma listset_finite[iff]: assumes "list_all (\ A. A \ {}) XS" shows "finite (listset XS) \ list_all finite XS" using assms proof (induct XS) case Nil show ?case by simp next case (Cons A XS) note [simp] = finite_image_iff finite_cartesian_product_iff have "listset (A # XS) = case_prod Cons ` (A \ listset XS)" by (auto simp: set_Cons_def) also have "finite \ \ finite (A \ listset XS)" by (simp add: inj_on_def) also have "\ \ finite A \ finite (listset XS)" using Cons(2) by simp also have "finite (listset XS) \ list_all finite XS" using Cons by simp also have "finite A \ \ \ list_all finite (A # XS)" by simp finally show ?case by this qed lemma listset_card[simp]: "card (listset XS) = prod_list (map card XS)" proof (induct XS) case Nil show ?case by simp next case (Cons A XS) have 1: "inj (case_prod Cons)" unfolding inj_def by simp have "listset (A # XS) = case_prod Cons ` (A \ listset XS)" by (auto simp: set_Cons_def) also have "card \ = card (A \ listset XS)" using card_image 1 by auto also have "\ = card A * card (listset XS)" using card_cartesian_product by this also have "card (listset XS) = prod_list (map card XS)" using Cons by this also have "card A * \ = prod_list (map card (A # XS))" by simp finally show ?case by this qed subsection \Stream Basics\ declare stream.map_id[simp] declare stream.set_map[simp] declare stream.set_sel(1)[intro!, simp] declare stream.pred_True[simp] declare stream.pred_map[iff] declare stream.rel_map[iff] declare shift_simps[simp del] declare stake_sdrop[simp] declare stake_siterate[simp del] declare sdrop_snth[simp] lemma stream_pred_cases: assumes "pred_stream P xs" obtains (scons) y ys where "xs = y ## ys" "P y" "pred_stream P ys" using assms by (cases xs) (auto) lemma stream_rel_coinduct[case_names stream_rel, coinduct pred: stream_all2]: assumes "R u v" assumes "\ a u b v. R (a ## u) (b ## v) \ P a b \ R u v" shows "stream_all2 P u v" using assms by (coinduct) (metis stream.collapse) lemma stream_rel_coinduct_shift[case_names stream_rel, consumes 1]: assumes "R u v" assumes "\ u v. R u v \ \ u\<^sub>1 u\<^sub>2 v\<^sub>1 v\<^sub>2. u = u\<^sub>1 @- u\<^sub>2 \ v = v\<^sub>1 @- v\<^sub>2 \ u\<^sub>1 \ [] \ v\<^sub>1 \ [] \ list_all2 P u\<^sub>1 v\<^sub>1 \ R u\<^sub>2 v\<^sub>2" shows "stream_all2 P u v" proof - have "\ u\<^sub>1 u\<^sub>2 v\<^sub>1 v\<^sub>2. u = u\<^sub>1 @- u\<^sub>2 \ v = v\<^sub>1 @- v\<^sub>2 \ list_all2 P u\<^sub>1 v\<^sub>1 \ R u\<^sub>2 v\<^sub>2" using assms(1) by force then show ?thesis using assms(2) by (coinduct) (force elim: list.rel_cases) qed lemma stream_pred_coinduct[case_names stream_pred, coinduct pred: pred_stream]: assumes "R w" assumes "\ a w. R (a ## w) \ P a \ R w" shows "pred_stream P w" using assms unfolding stream.pred_rel eq_onp_def by (coinduction arbitrary: w) (auto) lemma stream_pred_coinduct_shift[case_names stream_pred, consumes 1]: assumes "R w" assumes "\ w. R w \ \ u v. w = u @- v \ u \ [] \ list_all P u \ R v" shows "pred_stream P w" proof - have "\ u v. w = u @- v \ list_all P u \ R v" using assms(1) by (metis list_all_simps(2) shift.simps(1)) then show ?thesis using assms(2) by (coinduct) (force elim: list_pred_cases) qed lemma stream_pred_flat_coinduct[case_names stream_pred, consumes 1]: assumes "R ws" assumes "\ w ws. R (w ## ws) \ w \ [] \ list_all P w \ R ws" shows "pred_stream P (flat ws)" using assms by (coinduction arbitrary: ws rule: stream_pred_coinduct_shift) (metis stream.exhaust flat_Stream) lemmas stream_eq_coinduct[case_names stream_eq, coinduct pred: HOL.eq] = stream_rel_coinduct[where ?P = HOL.eq, unfolded stream.rel_eq] lemmas stream_eq_coinduct_shift[case_names stream_eq, consumes 1] = stream_rel_coinduct_shift[where ?P = HOL.eq, unfolded stream.rel_eq list.rel_eq] lemma stream_pred_shift[iff]: "pred_stream P (u @- v) \ list_all P u \ pred_stream P v" by (induct u) (auto) lemma stream_rel_shift[iff]: assumes "length u\<^sub>1 = length v\<^sub>1" shows "stream_all2 P (u\<^sub>1 @- u\<^sub>2) (v\<^sub>1 @- v\<^sub>2) \ list_all2 P u\<^sub>1 v\<^sub>1 \ stream_all2 P u\<^sub>2 v\<^sub>2" using assms by (induct rule: list_induct2) (auto) lemma sset_subset_stream_pred: "sset w \ A \ pred_stream (\ a. a \ A) w" unfolding stream.pred_set by auto lemma eq_scons: "w = a ## v \ a = shd w \ v = stl w" by auto lemma scons_eq: "a ## v = w \ shd w = a \ stl w = v" by auto lemma eq_shift: "w = u @- v \ stake (length u) w = u \ sdrop (length u) w = v" by (induct u arbitrary: w) (force+) lemma shift_eq: "u @- v = w \ u = stake (length u) w \ v = sdrop (length u) w" by (induct u arbitrary: w) (force+) lemma scons_eq_shift: "a ## w = u @- v \ ([] = u \ a ## w = v) \ (\ u'. a # u' = u \ w = u' @- v)" by (cases u) (auto) lemma shift_eq_scons: "u @- v = a ## w \ (u = [] \ v = a ## w) \ (\ u'. u = a # u' \ u' @- v = w)" by (cases u) (auto) lemma stream_all2_sset1: assumes "stream_all2 P xs ys" shows "\ x \ sset xs. \ y \ sset ys. P x y" proof - have "pred_stream (\ x. \ y \ S. P x y) xs" if "sset ys \ S" for S using assms that by (coinduction arbitrary: xs ys) (force elim: stream.rel_cases) then show ?thesis unfolding stream.pred_set by auto qed lemma stream_all2_sset2: assumes "stream_all2 P xs ys" shows "\ y \ sset ys. \ x \ sset xs. P x y" proof - have "pred_stream (\ y. \ x \ S. P x y) ys" if "sset xs \ S" for S using assms that by (coinduction arbitrary: xs ys) (force elim: stream.rel_cases) then show ?thesis unfolding stream.pred_set by auto qed lemma smap_eq_scons[iff]: "smap f xs = y ## ys \ f (shd xs) = y \ smap f (stl xs) = ys" using smap_ctr by metis lemma scons_eq_smap[iff]: "y ## ys = smap f xs \ y = f (shd xs) \ ys = smap f (stl xs)" using smap_ctr by metis lemma smap_eq_shift[iff]: "smap f w = u @- v \ (\ w\<^sub>1 w\<^sub>2. w = w\<^sub>1 @- w\<^sub>2 \ map f w\<^sub>1 = u \ smap f w\<^sub>2 = v)" using sdrop_smap eq_shift stake_sdrop stake_smap by metis lemma shift_eq_smap[iff]: "u @- v = smap f w \ (\ w\<^sub>1 w\<^sub>2. w = w\<^sub>1 @- w\<^sub>2 \ u = map f w\<^sub>1 \ v = smap f w\<^sub>2)" using sdrop_smap eq_shift stake_sdrop stake_smap by metis lemma szip_eq_scons[iff]: "szip xs ys = z ## zs \ (shd xs, shd ys) = z \ szip (stl xs) (stl ys) = zs" using szip.ctr stream.inject by metis lemma scons_eq_szip[iff]: "z ## zs = szip xs ys \ z = (shd xs, shd ys) \ zs = szip (stl xs) (stl ys)" using szip.ctr stream.inject by metis lemma siterate_eq_scons[iff]: "siterate f s = a ## w \ s = a \ siterate f (f s) = w" using siterate.ctr stream.inject by metis lemma scons_eq_siterate[iff]: "a ## w = siterate f s \ a = s \ w = siterate f (f s)" using siterate.ctr stream.inject by metis lemma snth_0: "(a ## w) !! 0 = a" by simp lemma eqI_snth: assumes "\ i. u !! i = v !! i" shows "u = v" using assms by (coinduction arbitrary: u v) (metis stream.sel snth.simps) lemma stream_pred_snth: "pred_stream P w \ (\ i. P (w !! i))" unfolding stream.pred_set sset_range by simp lemma stream_rel_snth: "stream_all2 P u v \ (\ i. P (u !! i) (v !! i))" proof safe show "P (u !! i) (v !! i)" if "stream_all2 P u v" for i using that by (induct i arbitrary: u v) (auto elim: stream.rel_cases) show "stream_all2 P u v" if "\ i. P (u !! i) (v !! i)" using that by (coinduct) (metis snth_0 snth_Stream) qed lemma stream_rel_pred_szip: "stream_all2 P u v \ pred_stream (case_prod P) (szip u v)" unfolding stream_pred_snth stream_rel_snth by simp lemma sconst_eq[iff]: "sconst x = sconst y \ x = y" by (auto) (metis siterate.simps(1)) lemma stream_pred__sconst[iff]: "pred_stream P (sconst x) \ P x" unfolding stream_pred_snth by simp lemma stream_rel_sconst[iff]: "stream_all2 P (sconst x) (sconst y) \ P x y" unfolding stream_rel_snth by simp lemma set_sset_stake[intro!, simp]: "set (stake n xs) \ sset xs" by (metis sset_shift stake_sdrop sup_ge1) lemma sset_sdrop[intro!, simp]: "sset (sdrop n xs) \ sset xs" by (metis sset_shift stake_sdrop sup_ge2) lemma set_stake_snth: "x \ set (stake n xs) \ (\ i < n. xs !! i = x)" unfolding in_set_conv_nth by auto (* TODO: these should be generated by the transfer option on the primcorec definitions *) (* TODO: transfer_prover cannot handle corecursive definitions due to the \(s1, s2). undefined abortion term that is never used *) lemma szip_transfer[transfer_rule]: includes lifting_syntax shows "(stream_all2 A ===> stream_all2 B ===> stream_all2 (rel_prod A B)) szip szip" by (intro rel_funI, coinduction) (force elim: stream.rel_cases) lemma siterate_transfer[transfer_rule]: includes lifting_syntax shows "((A ===> A) ===> A ===> stream_all2 A) siterate siterate" by (intro rel_funI, coinduction) (force dest: rel_funD) lemma split_stream_first: assumes "A \ sset xs \ {}" obtains ys a zs where "xs = ys @- a ## zs" "A \ set ys = {}" "a \ A" proof let ?n = "LEAST n. xs !! n \ A" have 1: "xs !! n \ A" if "n < ?n" for n using that by (metis (full_types) not_less_Least) show "xs = stake ?n xs @- (xs !! ?n) ## sdrop (Suc ?n) xs" using id_stake_snth_sdrop by blast show "A \ set (stake ?n xs) = {}" using 1 by (metis (no_types, lifting) disjoint_iff_not_equal set_stake_snth) show "xs !! ?n \ A" using assms unfolding sset_range by (auto intro: LeastI) qed lemma split_stream_first': assumes "x \ sset xs" obtains ys zs where "xs = ys @- x ## zs" "x \ set ys" proof let ?n = "LEAST n. xs !! n = x" have 1: "xs !! ?n = x" using assms unfolding sset_range by (auto intro: LeastI) have 2: "xs !! n \ x" if "n < ?n" for n using that by (metis (full_types) not_less_Least) show "xs = stake ?n xs @- x ## sdrop (Suc ?n) xs" using 1 by (metis id_stake_snth_sdrop) show "x \ set (stake ?n xs)" using 2 by (meson set_stake_snth) qed lemma streams_UNIV[iff]: "streams A = UNIV \ A = UNIV" proof show "A = UNIV \ streams A = UNIV" by simp next assume "streams A = UNIV" then have "w \ streams A" for w by simp then have "sset w \ A" for w unfolding streams_iff_sset by this then have "sset (sconst a) \ A" for a by blast then have "a \ A" for a by simp then show "A = UNIV" by auto qed lemma pred_list_listsp[pred_set_conv]: "list_all = listsp" unfolding list.pred_set by auto lemma pred_stream_streamsp[pred_set_conv]: "pred_stream = streamsp" unfolding stream.pred_set streams_iff_sset[to_pred] by auto subsection \The scan Function\ primrec (transfer) scan :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b list" where "scan f [] a = []" | "scan f (x # xs) a = f x a # scan f xs (f x a)" lemma scan_append[simp]: "scan f (xs @ ys) a = scan f xs a @ scan f ys (fold f xs a)" by (induct xs arbitrary: a) (auto) lemma scan_eq_nil[iff]: "scan f xs a = [] \ xs = []" by (cases xs) (auto) lemma scan_eq_cons[iff]: "scan f xs a = b # w \ (\ y ys. xs = y # ys \ f y a = b \ scan f ys (f y a) = w)" by (cases xs) (auto) lemma scan_eq_append[iff]: "scan f xs a = u @ v \ (\ ys zs. xs = ys @ zs \ scan f ys a = u \ scan f zs (fold f ys a) = v)" by (induct u arbitrary: xs a) (auto, metis append_Cons fold_simps(2), blast) lemma scan_length[simp]: "length (scan f xs a) = length xs" by (induct xs arbitrary: a) (auto) (* TODO: figure out how this is used, should it be a simp rule? or flipped? also target_alt_def *) lemma scan_last: "last (a # scan f xs a) = fold f xs a" by (induct xs arbitrary: a) (auto simp: last.simps) lemma scan_butlast[simp]: "scan f (butlast xs) a = butlast (scan f xs a)" by (induct xs arbitrary: a) (auto simp: butlast.simps) lemma scan_const[simp]: "scan const xs a = xs" by (induct xs arbitrary: a) (auto) lemma scan_nth[simp]: assumes "i < length (scan f xs a)" shows "scan f xs a ! i = fold f (take (Suc i) xs) a" using assms by (cases xs, simp, induct i arbitrary: xs a, auto simp: take_Suc neq_Nil_conv) lemma scan_map[simp]: "scan f (map g xs) a = scan (f \ g) xs a" by (induct xs arbitrary: a) (auto) lemma scan_take[simp]: "take k (scan f xs a) = scan f (take k xs) a" by (induct k arbitrary: xs a) (auto simp: take_Suc neq_Nil_conv) lemma scan_drop[simp]: "drop k (scan f xs a) = scan f (drop k xs) (fold f (take k xs) a)" by (induct k arbitrary: xs a) (auto simp: take_Suc neq_Nil_conv) primcorec (transfer) sscan :: "('a \ 'b \ 'b) \ 'a stream \ 'b \ 'b stream" where "sscan f xs a = f (shd xs) a ## sscan f (stl xs) (f (shd xs) a)" lemma sscan_scons[simp]: "sscan f (x ## xs) a = f x a ## sscan f xs (f x a)" by (simp add: stream.expand) lemma sscan_shift[simp]: "sscan f (xs @- ys) a = scan f xs a @- sscan f ys (fold f xs a)" by (induct xs arbitrary: a) (auto) lemma sscan_eq_scons[iff]: "sscan f xs a = b ## w \ f (shd xs) a = b \ sscan f (stl xs) (f (shd xs) a) = w" using sscan.ctr stream.inject by metis lemma scons_eq_sscan[iff]: "b ## w = sscan f xs a \ b = f (shd xs) a \ w = sscan f (stl xs) (f (shd xs) a)" using sscan.ctr stream.inject by metis lemma sscan_const[simp]: "sscan const xs a = xs" by (coinduction arbitrary: xs a) (auto) lemma sscan_snth[simp]: "sscan f xs a !! i = fold f (stake (Suc i) xs) a" by (induct i arbitrary: xs a) (auto) lemma sscan_scons_snth[simp]: "(a ## sscan f xs a) !! i = fold f (stake i xs) a" by (induct i arbitrary: xs a) (auto) lemma sscan_smap[simp]: "sscan f (smap g xs) a = sscan (f \ g) xs a" by (coinduction arbitrary: xs a) (auto) lemma sscan_stake[simp]: "stake k (sscan f xs a) = scan f (stake k xs) a" by (induct k arbitrary: a xs) (auto) lemma sscan_sdrop[simp]: "sdrop k (sscan f xs a) = sscan f (sdrop k xs) (fold f (stake k xs) a)" by (induct k arbitrary: a xs) (auto) subsection \Distinct Streams\ coinductive sdistinct :: "'a stream \ bool" where scons[intro!]: "x \ sset xs \ sdistinct xs \ sdistinct (x ## xs)" lemma sdistinct_scons_elim[elim!]: assumes "sdistinct (x ## xs)" obtains "x \ sset xs" "sdistinct xs" using assms by (auto elim: sdistinct.cases) lemma sdistinct_coinduct[case_names sdistinct, coinduct pred: sdistinct]: assumes "P xs" assumes "\ x xs. P (x ## xs) \ x \ sset xs \ P xs" shows "sdistinct xs" using stream.collapse sdistinct.coinduct assms by metis lemma sdistinct_shift[intro!]: assumes "distinct xs" "sdistinct ys" "set xs \ sset ys = {}" shows "sdistinct (xs @- ys)" using assms by (induct xs) (auto) lemma sdistinct_shift_elim[elim!]: assumes "sdistinct (xs @- ys)" obtains "distinct xs" "sdistinct ys" "set xs \ sset ys = {}" using assms by (induct xs) (auto) lemma sdistinct_infinite_sset: assumes "sdistinct w" shows "infinite (sset w)" using assms by (coinduction arbitrary: w) (force elim: sdistinct.cases) lemma not_sdistinct_decomp: assumes "\ sdistinct w" obtains u v a w' where "w = u @- a ## v @- a ## w'" proof (rule ccontr) assume 1: "\ thesis" assume 2: "w = u @- a ## v @- a ## w' \ thesis" for u a v w' have 3: "\ u v a w'. w \ u @- a ## v @- a ## w'" using 1 2 by auto have 4: "sdistinct w" using 3 by (coinduct) (metis id_stake_snth_sdrop imageE shift.simps sset_range) show False using assms 4 by auto qed subsection \Sorted Streams\ coinductive (in order) sascending :: "'a stream \ bool" where "a \ b \ sascending (b ## w) \ sascending (a ## b ## w)" coinductive (in order) sdescending :: "'a stream \ bool" where "a \ b \ sdescending (b ## w) \ sdescending (a ## b ## w)" lemma sdescending_coinduct[case_names sdescending, coinduct pred: sdescending]: assumes "P w" assumes "\ a b w. P (a ## b ## w) \ a \ b \ P (b ## w)" shows "sdescending w" using stream.collapse sdescending.coinduct assms by (metis (no_types)) lemma sdescending_sdrop: assumes "sdescending w" shows "sdescending (sdrop k w)" using assms by (induct k) (auto, metis sdescending.cases sdrop_stl stream.sel(2)) lemma sdescending_snth_antimono: assumes "sdescending w" shows "antimono (snth w)" unfolding antimono_iff_le_Suc proof fix k have "sdescending (sdrop k w)" using sdescending_sdrop assms by this then obtain a b v where 2: "sdrop k w = a ## b ## v" "a \ b" by rule then show "w !! k \ w !! Suc k" by (metis sdrop_simps stream.sel) qed lemma sdescending_stuck: fixes w :: "'a :: wellorder stream" assumes "sdescending w" obtains k where "sdrop k w = sconst (w !! k)" using assms proof (induct "w !! 0" arbitrary: w thesis rule: less_induct) case less show ?case proof (cases "w = sconst (w !! 0)") case True show ?thesis using less(2)[of 0] True by simp next case False obtain k where 1: "w !! k \ w !! 0" using False by (metis empty_iff eqI_snth insert_iff snth_sset sset_sconst) have 2: "antimono (snth w)" using sdescending_snth_antimono less(3) by this have 3: "w !! k \ w !! 0" using 1 2 by (blast dest: antimonoD) have 4: "sdrop k w !! 0 < w !! 0" using 1 3 by auto have 5: "sdescending (sdrop k w)" using sdescending_sdrop less(3) by this obtain l where 5: "sdrop l (sdrop k w) = sconst (sdrop k w !! l)" using less(1)[OF 4 _ 5] by this show ?thesis using less(2) 5 by simp qed qed end diff --git a/thys/Transition_Systems_and_Automata/Basic/Sequence_LTL.thy b/thys/Transition_Systems_and_Automata/Basic/Sequence_LTL.thy --- a/thys/Transition_Systems_and_Automata/Basic/Sequence_LTL.thy +++ b/thys/Transition_Systems_and_Automata/Basic/Sequence_LTL.thy @@ -1,114 +1,116 @@ section \Linear Temporal Logic on Streams\ theory Sequence_LTL imports "Sequence" "HOL-Library.Linear_Temporal_Logic_on_Streams" begin subsection \Basics\ text \Avoid destroying the constant @{const holds} prematurely.\ lemmas [simp del] = holds.simps holds_eq1 holds_eq2 not_holds_eq + (* TODO: these cannot be applied successively to simplify infs due to introduction of \ + avoid or add extra simplification rules for infs *) lemma ev_smap[iff]: "ev P (smap f w) \ ev (P \ smap f) w" using ev_smap unfolding comp_apply by this lemma alw_smap[iff]: "alw P (smap f w) \ alw (P \ smap f) w" using alw_smap unfolding comp_apply by this lemma holds_smap[iff]: "holds P (smap f w) \ holds (P \ f) w" unfolding holds.simps by simp lemmas [iff] = ev_sconst alw_sconst hld_smap' lemmas [iff] = alw_ev_stl lemma alw_ev_sdrop[iff]: "alw (ev P) (sdrop n w) \ alw (ev P) w" using alw_ev_sdrop alw_sdrop by blast lemma alw_ev_scons[iff]: "alw (ev P) (a ## w) \ alw (ev P) w" by (metis alw_ev_stl stream.sel(2)) lemma alw_ev_shift[iff]: "alw (ev P) (u @- v) \ alw (ev P) v" by (induct u) (auto) lemmas [simp del, iff] = ev_alw_stl lemma ev_alw_sdrop[iff]: "ev (alw P) (sdrop n w) \ ev (alw P) w" using alwD alw_alw alw_sdrop ev_alw_imp_alw_ev not_ev_iff by metis lemma ev_alw_scons[iff]: "ev (alw P) (a ## w) \ ev (alw P) w" by (metis ev_alw_stl stream.sel(2)) lemma ev_alw_shift[iff]: "ev (alw P) (u @- v) \ ev (alw P) v" by (induct u) (auto) lemma holds_sconst[iff]: "holds P (sconst a) \ P a" unfolding holds.simps by simp lemma HLD_sconst[iff]: "HLD A (sconst a) \ a \ A" unfolding HLD_def holds.simps by simp lemma ev_alt_def: "ev \ w \ (\ u v. w = u @- v \ \ v)" using ev.base ev_shift ev_imp_shift by metis lemma ev_stl_alt_def: "ev \ (stl w) \ (\ u v. w = u @- v \ u \ [] \ \ v)" unfolding ev_alt_def by (cases w) (force simp: scons_eq) lemma ev_HLD_sset: "ev (HLD A) w \ sset w \ A \ {}" unfolding HLD_def ev_holds_sset by auto lemma alw_ev_coinduct[case_names alw_ev, consumes 1]: assumes "R w" assumes "\ w. R w \ ev \ w \ ev R (stl w)" shows "alw (ev \) w" proof - have "ev R w" using assms(1) by rule then show ?thesis using assms(2) by (coinduct) (metis alw_sdrop not_ev_iff sdrop_stl sdrop_wait) qed subsection \Infinite Occurrence\ abbreviation "infs P w \ alw (ev (holds P)) w" abbreviation "fins P w \ \ infs P w" lemma infs_suffix: "infs P w \ (\ u v. w = u @- v \ Bex (sset v) P)" using alwD alw_iff_sdrop alw_shift ev_holds_sset stake_sdrop by (metis (mono_tags, hide_lams)) lemma infs_snth: "infs P w \ (\ n. \ k \ n. P (w !! k))" by (auto simp: alw_iff_sdrop ev_iff_sdrop holds.simps intro: le_add1 dest: le_Suc_ex) lemma infs_infm: "infs P w \ (\\<^sub>\ i. P (w !! i))" unfolding infs_snth INFM_nat_le by rule lemma infs_coinduct[case_names infs, coinduct pred: infs]: assumes "R w" assumes "\ w. R w \ Bex (sset w) P \ ev R (stl w)" shows "infs P w" using assms by (coinduct rule: alw_ev_coinduct) (auto simp: ev_holds_sset) lemma infs_set_coinduct[case_names infs_set, consumes 1]: assumes "R w" assumes "\ w. R w \ \ u v. w = u @- v \ Bex (set u) P \ R v" shows "infs P w" using assms by (coinduct) (force simp: ev_stl_alt_def) lemma infs_flat_coinduct[case_names infs_flat, consumes 1]: assumes "R w" assumes "\ u v. R (u ## v) \ Bex (set u) P \ R v" shows "infs P (flat w)" using assms by (coinduction arbitrary: w rule: infs_set_coinduct) (metis empty_iff flat_Stream list.set(1) stream.exhaust) lemma infs_mono: "(\ a. a \ sset w \ P a \ Q a) \ infs P w \ infs Q w" unfolding infs_snth by force lemma infs_mono_strong: "stream_all2 (\ a b. P a \ Q b) u v \ infs P u \ infs Q v" unfolding stream_rel_snth infs_snth by blast lemma infs_all: "Ball (sset w) P \ infs P w" unfolding infs_snth by auto lemma infs_any: "infs P w \ Bex (sset w) P" unfolding ev_holds_sset by auto lemma infs_bot[iff]: "infs bot w \ False" using infs_any by auto lemma infs_top[iff]: "infs top w \ True" by (simp add: infs_all) lemma infs_disj[iff]: "infs (\ a. P a \ Q a) w \ infs P w \ infs Q w" unfolding infs_snth using le_trans le_cases by metis lemma infs_bex[iff]: assumes "finite S" shows "infs (\ a. \ x \ S. P x a) w \ (\ x \ S. infs (P x) w)" using assms infs_any by induct auto lemma infs_bex_le_nat[iff]: "infs (\ a. \ k < n :: nat. P k a) w \ (\ k < n. infs (P k) w)" proof - have "infs (\ a. \ k < n. P k a) w \ infs (\ a. \ k \ {k. k < n}. P k a) w" by simp also have "\ \ (\ k \ {k. k < n}. infs (P k) w)" by blast also have "\ \ (\ k < n. infs (P k) w)" by simp finally show ?thesis by this qed lemma infs_cycle[iff]: assumes "w \ []" shows "infs P (cycle w) \ Bex (set w) P" proof show "infs P (cycle w) \ Bex (set w) P" using assms by (auto simp: ev_holds_sset dest: alwD) show "Bex (set w) P \ infs P (cycle w)" using assms by (coinduction rule: infs_set_coinduct) (blast dest: cycle_decomp) qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Basic/Sequence_Zip.thy b/thys/Transition_Systems_and_Automata/Basic/Sequence_Zip.thy --- a/thys/Transition_Systems_and_Automata/Basic/Sequence_Zip.thy +++ b/thys/Transition_Systems_and_Automata/Basic/Sequence_Zip.thy @@ -1,128 +1,143 @@ section \Zipping Sequences\ theory Sequence_Zip -imports "Sequence" +imports "Sequence_LTL" begin subsection \Zipping Lists\ notation zip (infixr "||" 51) lemmas [simp] = zip_map_fst_snd lemma split_zip[no_atp]: "(\ x. PROP P x) \ (\ y z. length y = length z \ PROP P (y || z))" proof fix y z assume 1: "\ x. PROP P x" show "PROP P (y || z)" using 1 by this next fix x :: "('a \ 'b) list" assume 1: "\ y z. length y = length z \ PROP P (y || z)" have 2: "length (map fst x) = length (map snd x)" by simp have 3: "PROP P (map fst x || map snd x)" using 1 2 by this show "PROP P x" using 3 by simp qed lemma split_zip_all[no_atp]: "(\ x. P x) \ (\ y z. length y = length z \ P (y || z))" by (fastforce iff: split_zip) lemma split_zip_ex[no_atp]: "(\ x. P x) \ (\ y z. length y = length z \ P (y || z))" by (fastforce iff: split_zip) lemma zip_eq[iff]: assumes "length u = length v" "length r = length s" shows "u || v = r || s \ u = r \ v = s" using assms zip_eq_conv by metis lemma list_rel_zip[iff]: assumes "length u = length v" "length r = length s" shows "list_all2 (rel_prod A B) (u || v) (r || s) \ list_all2 A u r \ list_all2 B v s" proof safe assume [transfer_rule]: "list_all2 (rel_prod A B) (u || v) (r || s)" have "list_all2 A (map fst (u || v)) (map fst (r || s))" by transfer_prover then show "list_all2 A u r" using assms by simp have "list_all2 B (map snd (u || v)) (map snd (r || s))" by transfer_prover then show "list_all2 B v s" using assms by simp next assume [transfer_rule]: "list_all2 A u r" "list_all2 B v s" show "list_all2 (rel_prod A B) (u || v) (r || s)" by transfer_prover qed lemma zip_last[simp]: assumes "xs || ys \ []" "length xs = length ys" shows "last (xs || ys) = (last xs, last ys)" proof - have 1: "xs \ []" "ys \ []" using assms(1) by auto have "last (xs || ys) = (xs || ys) ! (length (xs || ys) - 1)" using last_conv_nth assms by blast also have "\ = (xs ! (length (xs || ys) - 1), ys ! (length (xs || ys) - 1))" using assms 1 by simp also have "\ = (xs ! (length xs - 1), ys ! (length ys - 1))" using assms(2) by simp also have "\ = (last xs, last ys)" using last_conv_nth 1 by metis finally show ?thesis by this qed subsection \Zipping Streams\ notation szip (infixr "|||" 51) lemmas [simp] = szip_unfold lemma szip_smap[simp]: "smap fst zs ||| smap snd zs = zs" by (coinduction arbitrary: zs) (auto) lemma szip_smap_fst[simp]: "smap fst (xs ||| ys) = xs" by (coinduction arbitrary: xs ys) (auto) lemma szip_smap_snd[simp]: "smap snd (xs ||| ys) = ys" by (coinduction arbitrary: xs ys) (auto) lemma split_szip[no_atp]: "(\ x. PROP P x) \ (\ y z. PROP P (y ||| z))" proof fix y z assume 1: "\ x. PROP P x" show "PROP P (y ||| z)" using 1 by this next fix x assume 1: "\ y z. PROP P (y ||| z)" have 2: "PROP P (smap fst x ||| smap snd x)" using 1 by this show "PROP P x" using 2 by simp qed lemma split_szip_all[no_atp]: "(\ x. P x) \ (\ y z. P (y ||| z))" by (fastforce iff: split_szip) lemma split_szip_ex[no_atp]: "(\ x. P x) \ (\ y z. P (y ||| z))" by (fastforce iff: split_szip) lemma szip_eq[iff]: "u ||| v = r ||| s \ u = r \ v = s" using szip_smap_fst szip_smap_snd by metis lemma stream_rel_szip[iff]: "stream_all2 (rel_prod A B) (u ||| v) (r ||| s) \ stream_all2 A u r \ stream_all2 B v s" proof safe assume [transfer_rule]: "stream_all2 (rel_prod A B) (u ||| v) (r ||| s)" have "stream_all2 A (smap fst (u ||| v)) (smap fst (r ||| s))" by transfer_prover then show "stream_all2 A u r" by simp have "stream_all2 B (smap snd (u ||| v)) (smap snd (r ||| s))" by transfer_prover then show "stream_all2 B v s" by simp next assume [transfer_rule]: "stream_all2 A u r" "stream_all2 B v s" show "stream_all2 (rel_prod A B) (u ||| v) (r ||| s)" by transfer_prover qed lemma szip_shift[simp]: assumes "length u = length s" shows "u @- v ||| s @- t = (u || s) @- (v ||| t)" using assms by (simp add: eq_shift stake_shift sdrop_shift) lemma szip_sset_fst[simp]: "fst ` sset (u ||| v) = sset u" by (metis stream.set_map szip_smap_fst) lemma szip_sset_snd[simp]: "snd ` sset (u ||| v) = sset v" by (metis stream.set_map szip_smap_snd) lemma szip_sset_elim[elim]: assumes "(a, b) \ sset (u ||| v)" obtains "a \ sset u" "b \ sset v" using assms by (metis image_eqI fst_conv snd_conv szip_sset_fst szip_sset_snd) lemma szip_sset[simp]: "sset (u ||| v) \ sset u \ sset v" by auto lemma sset_szip_finite[iff]: "finite (sset (u ||| v)) \ finite (sset u) \ finite (sset v)" proof safe assume 1: "finite (sset (u ||| v))" have 2: "finite (fst ` sset (u ||| v))" using 1 by blast have 3: "finite (snd ` sset (u ||| v))" using 1 by blast show "finite (sset u)" using 2 by simp show "finite (sset v)" using 3 by simp next assume 1: "finite (sset u)" "finite (sset v)" have "sset (u ||| v) \ sset u \ sset v" by simp also have "finite \" using 1 by simp finally show "finite (sset (u ||| v))" by this qed -end + lemma infs_szip_fst[iff]: "infs (P \ fst) (u ||| v) \ infs P u" + proof - + have "infs (P \ fst) (u ||| v) \ infs P (smap fst (u ||| v))" + by (simp add: comp_def del: szip_smap_fst) + also have "\ \ infs P u" by simp + finally show ?thesis by this + qed + lemma infs_szip_snd[iff]: "infs (P \ snd) (u ||| v) \ infs P v" + proof - + have "infs (P \ snd) (u ||| v) \ infs P (smap snd (u ||| v))" + by (simp add: comp_def del: szip_smap_snd) + also have "\ \ infs P v" by simp + finally show ?thesis by this + qed + +end \ No newline at end of file