diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Graphs.thy @@ -1,147 +1,147 @@ section \Connecting Nondeterministic Büchi Automata to CAVA Automata Structures\ theory NBA_Graphs imports NBA CAVA_Automata.Automata_Impl begin no_notation build (infixr "##" 65) subsection \Regular Graphs\ definition nba_g :: "('label, 'state) nba \ 'state graph_rec" where "nba_g A \ \ g_V = UNIV, g_E = E_of_succ (successors A), g_V0 = initial A \" lemma nba_g_graph[simp]: "graph (nba_g A)" unfolding nba_g_def graph_def by simp lemma nba_g_V0: "g_V0 (nba_g A) = initial A" unfolding nba_g_def by simp lemma nba_g_E_rtrancl: "(g_E (nba_g A))\<^sup>* = {(p, q). q \ reachable A p}" unfolding nba_g_def graph_rec.simps E_of_succ_def proof safe show "(p, q) \ {(p, q). q \ successors A p}\<^sup>*" if "q \ reachable A p" for p q using that by (induct) (auto intro: rtrancl_into_rtrancl) show "q \ reachable A p" if "(p, q) \ {(p, q). q \ successors A p}\<^sup>*" for p q using that by induct auto qed lemma nba_g_rtrancl_path: "(g_E (nba_g A))\<^sup>* = {(p, target r p) |r p. NBA.path A r p}" unfolding nba_g_E_rtrancl by blast lemma nba_g_trancl_path: "(g_E (nba_g A))\<^sup>+ = {(p, target r p) |r p. NBA.path A r p \ r \ []}" unfolding nba_g_def graph_rec.simps E_of_succ_def proof safe show "\ r p. (x, y) = (p, target r p) \ NBA.path A r p \ r \ []" if "(x, y) \ {(p, q). q \ successors A p}\<^sup>+" for x y using that proof induct case (base y) obtain a where 1: "a \ alphabet A" "y \ transition A a x" using base by auto show ?case proof (intro exI conjI) show "(x, y) = (x, target [(a, y)] x)" by simp show "NBA.path A [(a, y)] x" using 1 by auto show "[(a, y)] \ []" by simp qed next case (step y z) obtain r where 1: "y = target r x" "NBA.path A r x" "r \ []" using step(3) by auto obtain a where 2: "a \ alphabet A" "z \ transition A a y" using step(2) by auto show ?case proof (intro exI conjI) show "(x, z) = (x, target (r @ [(a, z)]) x)" by simp show "NBA.path A (r @ [(a, z)]) x" using 1 2 by auto show "r @ [(a, z)] \ []" by simp qed qed show "(p, target r p) \ {(u, v). v \ successors A u}\<^sup>+" if "NBA.path A r p" "r \ []" for r p using that by (induct) (fastforce intro: trancl_into_trancl2)+ qed lemma nba_g_ipath_run: assumes "ipath (g_E (nba_g A)) r" obtains w where "run A (w ||| smap (r \ Suc) nats) (r 0)" proof - have 1: "\ a \ alphabet A. r (Suc i) \ transition A a (r i)" for i using assms unfolding ipath_def nba_g_def E_of_succ_def by auto obtain wr where 2: "run A wr (r 0)" "\ i. target (stake i wr) (r 0) = r i" proof (rule nba.invariant_run_index) show "\ aq. (fst aq \ alphabet A \ snd aq \ transition A (fst aq) p) \ snd aq = r (Suc i) \ True" if "p = r i" for i p using that 1 by auto show "r 0 = r 0" by rule qed auto have 3: "smap (r \ Suc) nats = smap snd wr" proof (rule eqI_snth) fix i have "smap (r \ Suc) nats !! i = r (Suc i)" by simp also have "\ = target (stake (Suc i) wr) (r 0)" unfolding 2(2) by rule also have "\ = (r 0 ## trace wr (r 0)) !! Suc i" by simp also have "\ = smap snd wr !! i" unfolding nba.trace_alt_def by simp finally show "smap (r \ Suc) nats !! i = smap snd wr !! i" by this qed show ?thesis proof show "run A (smap fst wr ||| smap (r \ Suc) nats) (r 0)" using 2(1) unfolding 3 by auto qed qed lemma nba_g_run_ipath: assumes "run A (w ||| r) p" shows "ipath (g_E (nba_g A)) (snth (p ## r))" proof fix i have 1: "w !! i \ alphabet A" "r !! i \ transition A (w !! i) (target (stake i (w ||| r)) p)" using assms by (auto dest: nba.run_snth) have 2: "r !! i \ successors A ((p ## r) !! i)" using 1 unfolding sscan_scons_snth[symmetric] nba.trace_alt_def by auto show "((p ## r) !! i, (p ## r) !! Suc i) \ g_E (nba_g A)" using 2 unfolding nba_g_def graph_rec.simps E_of_succ_def by simp qed subsection \Indexed Generalized Büchi Graphs\ definition nba_igbg :: "('label, 'state) nba \ 'state igb_graph_rec" where "nba_igbg A \ graph_rec.extend (nba_g A) \ igbg_num_acc = 1, igbg_acc = \ p. if accepting A p then {0} else {} \" lemma acc_run_language: assumes "igb_graph (nba_igbg A)" shows "Ex (igb_graph.is_acc_run (nba_igbg A)) \ language A \ {}" proof interpret igb_graph "nba_igbg A" using assms by this have [simp]: "V0 = g_V0 (nba_g A)" "E = g_E (nba_g A)" "num_acc = 1" "0 \ acc p \ accepting A p" for p unfolding nba_igbg_def graph_rec.defs by simp+ show "language A \ {}" if run: "Ex is_acc_run" proof - obtain r where 1: "is_acc_run r" using run by rule have 2: "r 0 \ V0" "ipath E r" "is_acc r" using 1 unfolding is_acc_run_def graph_defs.is_run_def by auto obtain w where 3: "run A (w ||| smap (r \ Suc) nats) (r 0)" using nba_g_ipath_run 2(2) by auto have 4: "r 0 ## smap (r \ Suc) nats = smap r nats" by (simp) (metis stream.map_comp smap_siterate) have 5: "infs (accepting A) (r 0 ## smap (r \ Suc) nats)" using 2(3) unfolding infs_infm is_acc_def 4 by simp have "w \ language A" proof show "r 0 \ initial A" using nba_g_V0 2(1) by force show "run A (w ||| smap (r \ Suc) nats) (r 0)" using 3 by this show "infs (accepting A) (smap (r \ Suc) nats)" using 5 by simp qed then show ?thesis by auto qed show "Ex is_acc_run" if language: "language A \ {}" proof - obtain w where 1: "w \ language A" using language by auto obtain r p where 2: "p \ initial A" "run A (w ||| r) p" "infs (accepting A) r" using 1 by rule have 3: "infs (accepting A) (p ## r)" using 2(3) by simp have "is_acc_run (snth (p ## r))" unfolding is_acc_run_def graph_defs.is_run_def proof safe show "(p ## r) !! 0 \ V0" using nba_g_V0 2(1) by force show "ipath E (snth (p ## r))" using nba_g_run_ipath 2(2) by force - show "is_acc (snth (p ## r))" using 3 unfolding infs_infm is_acc_def nba.trace_alt_def by simp + show "is_acc (snth (p ## r))" using 3 unfolding infs_infm is_acc_def by simp qed then show ?thesis by auto qed qed end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy b/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy --- a/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy +++ b/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy @@ -1,742 +1,734 @@ theory Nondeterministic 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 set" locale automaton = fixes automaton :: "'label set \ 'state set \ ('label, 'state) trans \ 'condition \ 'automaton" fixes alphabet :: "'automaton \ 'label set" fixes initial :: "'automaton \ 'state set" 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 sublocale transition_system_initial "\ a p. snd a" "\ a p. fst a \ alphabet A \ snd a \ transition A (fst a) p" "\ p. p \ initial A" for A defines path' = path and run' = run and reachable' = reachable and nodes' = nodes by this lemma states_alt_def: "states r p = map snd r" by (induct r arbitrary: p) (auto) lemma trace_alt_def: "trace r p = smap snd r" by (coinduction arbitrary: r p) (auto) lemma successors_alt_def: "successors A p = (\ a \ alphabet A. transition A a p)" by auto lemma reachable_transition[intro]: assumes "a \ alphabet A" "q \ reachable A p" "r \ transition A a q" shows "r \ reachable A p" using reachable.execute assms by force lemma nodes_transition[intro]: assumes "a \ alphabet A" "p \ nodes A" "q \ transition A a p" shows "q \ nodes A" using nodes.execute assms by force + lemma run_alphabet: + assumes "run A (w ||| r) p" + shows "w \ streams (alphabet A)" + using assms by (coinduction arbitrary: w r p) (metis run.cases stream.map szip_smap szip_smap_fst) + definition restrict :: "'automaton \ 'automaton" where "restrict A \ automaton (alphabet A) (initial A) (\ a p. if a \ alphabet A then transition A a p else {}) (condition A)" lemma restrict_simps[simp]: "alphabet (restrict A) = alphabet A" "initial (restrict A) = initial A" "transition (restrict A) a p = (if a \ alphabet A then transition A a p else {})" "condition (restrict A) = condition A" unfolding restrict_def by auto lemma restrict_path[simp]: "path (restrict A) = path A" proof (intro ext iffI) show "path A wr p" if "path (restrict A) wr p" for wr p using that by induct auto show "path (restrict A) wr p" if "path A wr p" for wr p using that by induct auto qed lemma restrict_run[simp]: "run (restrict A) = run A" proof (intro ext iffI) show "run A wr p" if "run (restrict A) wr p" for wr p using that by coinduct auto show "run (restrict A) wr p" if "run A wr p" for wr p using that by coinduct auto qed end (* TODO: create analogous thing for NFAs (automaton_target) *) - (* TODO: this should be _stream *) - (* TODO: remove all trace_alt_def *) + (* TODO: this should be _stream or _run *) locale automaton_trace = automaton automaton alphabet initial transition condition for automaton :: "'label set \ 'state set \ ('label, 'state) trans \ 'condition \ 'automaton" and alphabet :: "'automaton \ 'label set" and initial :: "'automaton \ 'state set" and transition :: "'automaton \ ('label, 'state) trans" and condition :: "'automaton \ 'condition" + fixes test :: "'condition \ 'label stream \ 'state stream \ 'state \ bool" begin definition language :: "'automaton \ 'label stream set" where "language A \ {w |w r p. p \ initial A \ run A (w ||| r) p \ test (condition A) w r p}" lemma language[intro]: assumes "p \ initial A" "run A (w ||| r) p" "test (condition A) w r p" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains r p where "p \ initial A" "run A (w ||| r) p" "test (condition A) w r p" using assms unfolding language_def by auto - (* TODO: move this up in the hierarchy *) - lemma run_alphabet: - assumes "run A (w ||| r) p" - shows "w \ streams (alphabet A)" - using assms by (coinduction arbitrary: w r p) (metis run.cases stream.map szip_smap szip_smap_fst) - lemma language_alphabet: "language A \ streams (alphabet A)" - unfolding language_def by (auto dest: run_alphabet) + lemma language_alphabet: "language A \ streams (alphabet A)" by (auto dest: run_alphabet) lemma restrict_language[simp]: "language (restrict A) = language A" by force 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 set \ ('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 set" 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 set \ ('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 set" 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 run_degeneralize: assumes "a.run A (w ||| r) p" shows "b.run (degeneralize A) (w ||| r ||| sscan (count (condition\<^sub>1 A)) (p ## r) k) (p, k)" using assms by (coinduction arbitrary: w r p k) (force elim: a.run.cases) lemma degeneralize_run: assumes "b.run (degeneralize A) (w ||| rs) pk" obtains r s p k where "rs = r ||| s" "pk = (p, k)" "a.run A (w ||| r) p" "s = sscan (count (condition\<^sub>1 A)) (p ## r) k" proof show "rs = smap fst rs ||| smap snd rs" "pk = (fst pk, snd pk)" by auto show "a.run A (w ||| smap fst rs) (fst pk)" using assms by (coinduction arbitrary: w rs pk) (force elim: b.run.cases) show "smap snd rs = sscan (count (condition\<^sub>1 A)) (fst pk ## smap fst rs) (snd pk)" using assms by (coinduction arbitrary: w rs pk) (force elim: b.run.cases) qed lemma degeneralize_nodes: "b.nodes (degeneralize A) \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" proof fix pk assume "pk \ b.nodes (degeneralize A)" then show "pk \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" by (induct) (force, cases "condition\<^sub>1 A = []", auto) qed lemma nodes_degeneralize: "a.nodes A \ fst ` b.nodes (degeneralize A)" proof fix p assume "p \ a.nodes A" then show "p \ fst ` b.nodes (degeneralize A)" proof induct case (initial p) have "(p, 0) \ b.nodes (degeneralize A)" using initial by auto then show ?case using image_iff fst_conv by force next case (execute p aq) obtain k where "(p, k) \ b.nodes (degeneralize A)" using execute(2) by auto then have "(snd aq, count (condition\<^sub>1 A) p k) \ b.nodes (degeneralize A)" using execute(3) by auto then show ?case using image_iff snd_conv by force qed qed 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 finite_subset nodes_degeneralize that by blast show "finite (b.nodes (degeneralize A))" if "finite (a.nodes A)" using finite_subset degeneralize_nodes that by blast 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 set \ ('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 set" 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 \ 'label stream \ 'state stream \ 'state \ bool" and automaton\<^sub>2 :: "'label set \ 'state degen set \ ('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 set" 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 \ 'label stream \ 'state degen stream \ 'state degen \ bool" + assumes test[iff]: "test\<^sub>2 (degen cs) w (r ||| sscan (count cs) (p ## r) k) (p, k) \ test\<^sub>1 cs w r p" begin lemma degeneralize_language[simp]: "b.language (degeneralize A) = a.language A" - unfolding a.language_def b.language_def - by (auto dest: run_degeneralize elim!: degeneralize_run) + unfolding a.language_def b.language_def by (auto dest: run_degeneralize elim!: degeneralize_run) end locale automaton_intersection = 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 set \ ('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 set" 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 set \ ('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 set" 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) set \ ('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) set" 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 intersect :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where "intersect 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 intersect_simps[simp]: "alphabet\<^sub>3 (intersect A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" "initial\<^sub>3 (intersect A B) = initial\<^sub>1 A \ initial\<^sub>2 B" "transition\<^sub>3 (intersect A B) a (p, q) = transition\<^sub>1 A a p \ transition\<^sub>2 B a q" "condition\<^sub>3 (intersect A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" unfolding intersect_def by auto lemma intersect_path[iff]: assumes "length w = length r" "length r = length s" shows "c.path (intersect A B) (w || r || s) (p, q) \ a.path A (w || r) p \ b.path B (w || s) q" using assms by (induct arbitrary: p q rule: list_induct3) (auto) lemma intersect_run[iff]: "c.run (intersect A B) (w ||| r ||| s) (p, q) \ a.run A (w ||| r) p \ b.run B (w ||| s) q" proof safe show "a.run A (w ||| r) p" if "c.run (intersect A B) (w ||| r ||| s) (p, q)" using that by (coinduction arbitrary: w r s p q) (force elim: c.run.cases) show "b.run B (w ||| s) q" if "c.run (intersect A B) (w ||| r ||| s) (p, q)" using that by (coinduction arbitrary: w r s p q) (force elim: c.run.cases) show "c.run (intersect A B) (w ||| r ||| s) (p, q)" if "a.run A (w ||| r) p" "b.run B (w ||| s) q" using that by (coinduction arbitrary: w r s p q) (auto elim: a.run.cases b.run.cases) qed lemma intersect_nodes: "c.nodes (intersect A B) \ a.nodes A \ b.nodes B" proof fix pq assume "pq \ c.nodes (intersect A B)" then show "pq \ a.nodes A \ b.nodes B" by induct auto qed lemma intersect_nodes_finite[intro]: assumes "finite (a.nodes A)" "finite (b.nodes B)" shows "finite (c.nodes (intersect A B))" using finite_subset intersect_nodes assms by blast end locale automaton_intersection_trace = automaton_intersection 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 set \ ('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 set" 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 \ 'label stream \ 'state\<^sub>1 stream \ 'state\<^sub>1 \ bool" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('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 set" 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 \ 'label stream \ 'state\<^sub>2 stream \ 'state\<^sub>2 \ bool" and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ ('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) set" 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 \ 'label stream \ ('state\<^sub>1 \ 'state\<^sub>2) stream \ 'state\<^sub>1 \ 'state\<^sub>2 \ 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) w (u ||| v) (p, q) \ test\<^sub>1 c\<^sub>1 w u p \ test\<^sub>2 c\<^sub>2 w v q" begin lemma intersect_language[simp]: "c.language (intersect A B) = a.language A \ b.language B" unfolding a.language_def b.language_def c.language_def by (fastforce iff: split_szip) end locale automaton_intersection_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 set \ ('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 set" 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 set \ ('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 set" 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 intersect :: "'automaton\<^sub>1 list \ 'automaton\<^sub>2" where "intersect AA \ automaton\<^sub>2 (\ (alphabet\<^sub>1 ` set AA)) (listset (map initial\<^sub>1 AA)) - (\ a pp. listset (map2 (\ A p. transition\<^sub>1 A a p) AA pp)) + (\ a ps. listset (map2 (\ A p. transition\<^sub>1 A a p) AA ps)) (condition (map condition\<^sub>1 AA))" lemma intersect_simps[simp]: "alphabet\<^sub>2 (intersect AA) = \ (alphabet\<^sub>1 ` set AA)" "initial\<^sub>2 (intersect AA) = listset (map initial\<^sub>1 AA)" - "transition\<^sub>2 (intersect AA) a pp = listset (map2 (\ A p. transition\<^sub>1 A a p) AA pp)" + "transition\<^sub>2 (intersect AA) a ps = listset (map2 (\ A p. transition\<^sub>1 A a p) AA ps)" "condition\<^sub>2 (intersect AA) = condition (map condition\<^sub>1 AA)" unfolding intersect_def by auto lemma intersect_run_length: - assumes "length pp = length AA" - assumes "b.run (intersect AA) (w ||| r) pp" - assumes "qq \ sset r" - shows "length qq = length AA" + assumes "length ps = length AA" + assumes "b.run (intersect AA) (w ||| r) ps" + assumes "qs \ sset r" + shows "length qs = length AA" proof - - have "pred_stream (\ qq. length qq = length AA) r" - using assms(1, 2) by (coinduction arbitrary: w r pp) + have "pred_stream (\ qs. length qs = length AA) r" + using assms(1, 2) by (coinduction arbitrary: w r ps) (force elim: b.run.cases simp: listset_member list_all2_conv_all_nth) then show ?thesis using assms(3) unfolding stream.pred_set by auto qed lemma intersect_run_stranspose: - assumes "length pp = length AA" - assumes "b.run (intersect AA) (w ||| r) pp" - obtains rr where "r = stranspose rr" "length rr = length AA" + assumes "length ps = length AA" + assumes "b.run (intersect AA) (w ||| r) ps" + obtains rs where "r = stranspose rs" "length rs = length AA" proof - define rr where "rr \ map (\ k. smap (\ pp. pp ! k) r) [0 ..< length AA]" - have "length qq = length AA" if "qq \ sset r" for qq using intersect_run_length assms that by this - then show "r = stranspose rr" - unfolding rr_def by (coinduction arbitrary: r) (auto intro: nth_equalityI simp: comp_def) - show "length rr = length AA" unfolding rr_def by auto + define rs where "rs \ map (\ k. smap (\ ps. ps ! k) r) [0 ..< length AA]" + have "length qs = length AA" if "qs \ sset r" for qs using intersect_run_length assms that by this + then show "r = stranspose rs" + unfolding rs_def by (coinduction arbitrary: r) (auto intro: nth_equalityI simp: comp_def) + show "length rs = length AA" unfolding rs_def by auto qed - lemma intersect_run: - assumes "length rr = length AA" "length pp = length AA" - assumes "\ k. k < length AA \ a.run (AA ! k) (w ||| rr ! k) (pp ! k)" - shows "b.run (intersect AA) (w ||| stranspose rr) pp" + lemma run_intersect: + assumes "length rs = length AA" "length ps = length AA" + assumes "\ k. k < length AA \ a.run (AA ! k) (w ||| rs ! k) (ps ! k)" + shows "b.run (intersect AA) (w ||| stranspose rs) ps" using assms - proof (coinduction arbitrary: w rr pp) + proof (coinduction arbitrary: w rs ps) case (run ap r) then show ?case proof (intro conjI exI) show "fst ap \ alphabet\<^sub>2 (intersect AA)" using run by (force elim: a.run.cases simp: set_conv_nth) - show "snd ap \ transition\<^sub>2 (intersect AA) (fst ap) pp" + show "snd ap \ transition\<^sub>2 (intersect AA) (fst ap) ps" using run by (force elim: a.run.cases simp: listset_member list_all2_conv_all_nth) - show "\ k < length AA. a.run' (AA ! k) (stl w ||| map stl rr ! k) (map shd rr ! k)" + show "\ k < length AA. a.run' (AA ! k) (stl w ||| map stl rs ! k) (map shd rs ! k)" using run by (force elim: a.run.cases) qed auto qed - lemma intersect_run_elim: - assumes "length rr = length AA" "length pp = length AA" - assumes "b.run (intersect AA) (w ||| stranspose rr) pp" - shows "k < length AA \ a.run (AA ! k) (w ||| rr ! k) (pp ! k)" + lemma intersect_run: + assumes "length rs = length AA" "length ps = length AA" + assumes "b.run (intersect AA) (w ||| stranspose rs) ps" + shows "k < length AA \ a.run (AA ! k) (w ||| rs ! k) (ps ! k)" using assms - proof (coinduction arbitrary: w rr pp) + proof (coinduction arbitrary: w rs ps) case (run ap wr) then show ?case proof (intro exI conjI) show "fst ap \ alphabet\<^sub>1 (AA ! k)" using run by (force elim: b.run.cases) - show "snd ap \ transition\<^sub>1 (AA ! k) (fst ap) (pp ! k)" + show "snd ap \ transition\<^sub>1 (AA ! k) (fst ap) (ps ! k)" using run by (force elim: b.run.cases simp: listset_member list_all2_conv_all_nth) - show "b.run' (intersect AA) (stl w ||| stranspose (map stl rr)) (shd (stranspose rr))" + show "b.run' (intersect AA) (stl w ||| stranspose (map stl rs)) (shd (stranspose rs))" using run by (force elim: b.run.cases) qed auto qed lemma intersect_nodes: "b.nodes (intersect AA) \ listset (map a.nodes AA)" proof - show "pp \ listset (map a.nodes AA)" if "pp \ b.nodes (intersect AA)" for pp + show "ps \ listset (map a.nodes AA)" if "ps \ b.nodes (intersect AA)" for ps using that by (induct) (auto 0 3 simp: listset_member list_all2_conv_all_nth) qed lemma intersect_nodes_finite[intro]: assumes "list_all (finite \ a.nodes) AA" shows "finite (b.nodes (intersect AA))" proof (rule finite_subset) show "b.nodes (intersect AA) \ listset (map a.nodes AA)" using intersect_nodes by this show "finite (listset (map a.nodes AA))" using list.pred_map assms by auto qed end locale automaton_intersection_list_trace = automaton_intersection_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 set \ ('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 set" 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 \ 'label stream \ 'state stream \ 'state \ bool" and automaton\<^sub>2 :: "'label set \ 'state list set \ ('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 set" 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 \ 'label stream \ 'state list stream \ 'state list \ bool" and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + assumes test[iff]: "length rs = length cs \ length ps = length cs \ - test\<^sub>2 (condition cs) w (stranspose rs) ps \ - list_all (\ (c, r, p). test\<^sub>1 c w r p) (cs || rs || ps)" + test\<^sub>2 (condition cs) w (stranspose rs) ps \ list_all (\ (c, r, p). test\<^sub>1 c w r p) (cs || rs || ps)" begin lemma intersect_language[simp]: "b.language (intersect AA) = \ (a.language ` set AA)" proof safe fix A w assume 1: "w \ b.language (intersect AA)" "A \ set AA" - obtain r pp where 2: - "pp \ initial\<^sub>2 (intersect AA)" - "b.run (intersect AA) (w ||| r) pp" - "test\<^sub>2 (condition\<^sub>2 (intersect AA)) w r pp" - using 1(1) unfolding b.language_def by auto - have 3: "length pp = length AA" using 2(1) by (simp add: listset_member list_all2_conv_all_nth) - obtain rr where 4: "r = stranspose rr" "length rr = length AA" - using intersect_run_stranspose 3 2(2) by blast + obtain r ps where 2: + "ps \ initial\<^sub>2 (intersect AA)" + "b.run (intersect AA) (w ||| r) ps" + "test\<^sub>2 (condition\<^sub>2 (intersect AA)) w r ps" + using 1(1) by auto + have 3: "length ps = length AA" using 2(1) by (simp add: listset_member list_all2_conv_all_nth) + obtain rs where 4: "r = stranspose rs" "length rs = length AA" + using intersect_run_stranspose 3 2(2) by this obtain k where 5: "k < length AA" "A = AA ! k" using 1(2) unfolding set_conv_nth by auto show "w \ a.language A" proof - show "pp ! k \ initial\<^sub>1 A" using 2(1) 5 by (auto simp: listset_member list_all2_conv_all_nth) - show "a.run A (w ||| rr ! k) (pp ! k)" using intersect_run_elim 2(2) 3 4 5 by auto - show "test\<^sub>1 (condition\<^sub>1 A) w (rr ! k) (pp ! k)" - using 2(3) 3 4 5 by (simp add: list_all_length) + show "ps ! k \ initial\<^sub>1 A" using 2(1) 5 by (auto simp: listset_member list_all2_conv_all_nth) + show "a.run A (w ||| rs ! k) (ps ! k)" using 2(2) 3 4 5 by (auto intro: intersect_run) + show "test\<^sub>1 (condition\<^sub>1 A) w (rs ! k) (ps ! k)" using 2(3) 3 4 5 by (simp add: list_all_length) qed next fix w assume 1: "w \ \ (a.language ` set AA)" have 2: "\ A \ set AA. \ r p. p \ initial\<^sub>1 A \ a.run A (w ||| r) p \ test\<^sub>1 (condition\<^sub>1 A) w r p" - using 1 unfolding a.language_def by auto - obtain rr pp where 3: - "length rr = length AA" "length pp = length AA" - "\ k. k < length AA \ pp ! k \ initial\<^sub>1 (AA ! k)" - "\ k. k < length AA \ a.run (AA ! k) (w ||| rr ! k) (pp ! k)" - "\ k. k < length AA \ test\<^sub>1 (condition\<^sub>1 (AA ! k)) w (rr ! k) (pp ! k)" + using 1 by blast + obtain rs ps where 3: + "length rs = length AA" "length ps = length AA" + "\ k. k < length AA \ ps ! k \ initial\<^sub>1 (AA ! k)" + "\ k. k < length AA \ a.run (AA ! k) (w ||| rs ! k) (ps ! k)" + "\ k. k < length AA \ test\<^sub>1 (condition\<^sub>1 (AA ! k)) w (rs ! k) (ps ! k)" using 2 unfolding Ball_set list_choice_zip list_choice_pair unfolding list.pred_set set_conv_nth by force show "w \ b.language (intersect AA)" proof - show "pp \ initial\<^sub>2 (intersect AA)" - using 3 by (force simp: listset_member list_all2_iff set_conv_nth) - show "b.run (intersect AA) (w ||| stranspose rr) pp" using intersect_run 3 by auto - show "test\<^sub>2 (condition\<^sub>2 (intersect AA)) w (stranspose rr) pp" - using 3 by (force simp: list_all_length) + show "ps \ initial\<^sub>2 (intersect AA)" using 3 by (auto simp: listset_member list_all2_conv_all_nth) + show "b.run (intersect AA) (w ||| stranspose rs) ps" using 3 by (auto intro: run_intersect) + show "test\<^sub>2 (condition\<^sub>2 (intersect AA)) w (stranspose rs) ps" using 3 by (auto simp: list_all_length) qed qed end locale automaton_union = 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 set \ ('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 set" 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 set \ ('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 set" 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) set \ ('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) set" 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 union :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where "union A B \ automaton\<^sub>3 (alphabet\<^sub>1 A \ alphabet\<^sub>2 B) (initial\<^sub>1 A <+> initial\<^sub>2 B) (\ a. \ Inl p \ Inl ` transition\<^sub>1 A a p | Inr q \ Inr ` transition\<^sub>2 B a q) (condition (condition\<^sub>1 A) (condition\<^sub>2 B))" lemma union_simps[simp]: "alphabet\<^sub>3 (union A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" "initial\<^sub>3 (union A B) = initial\<^sub>1 A <+> initial\<^sub>2 B" "transition\<^sub>3 (union A B) a (Inl p) = Inl ` transition\<^sub>1 A a p" "transition\<^sub>3 (union A B) a (Inr q) = Inr ` transition\<^sub>2 B a q" "condition\<^sub>3 (union A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" unfolding union_def by auto lemma run_union_a: assumes "a.run A (w ||| r) p" shows "c.run (union A B) (w ||| smap Inl r) (Inl p)" using assms by (coinduction arbitrary: w r p) (force elim: a.run.cases) lemma run_union_b: assumes "b.run B (w ||| s) q" shows "c.run (union A B) (w ||| smap Inr s) (Inr q)" using assms by (coinduction arbitrary: w s q) (force elim: b.run.cases) lemma run_union: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" assumes "c.run (union A B) (w ||| rs) pq" obtains (a) r p where "rs = smap Inl r" "pq = Inl p" "a.run A (w ||| r) p" | (b) s q where "rs = smap Inr s" "pq = Inr q" "b.run B (w ||| s) q" proof (cases pq) case (Inl p) have 1: "rs = smap Inl (smap projl rs)" using assms(2) unfolding Inl by (coinduction arbitrary: w rs p) (force elim: c.run.cases) have 2: "a.run A (w ||| smap projl rs) p" using assms unfolding Inl by (coinduction arbitrary: w rs p) (force elim: c.run.cases) show ?thesis using a 1 Inl 2 by this next case (Inr q) have 1: "rs = smap Inr (smap projr rs)" using assms(2) unfolding Inr by (coinduction arbitrary: w rs q) (force elim: c.run.cases) have 2: "b.run B (w ||| smap projr rs) q" using assms unfolding Inr by (coinduction arbitrary: w rs q) (force elim: c.run.cases) show ?thesis using b 1 Inr 2 by this qed lemma union_nodes: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "c.nodes (union A B) \ a.nodes A <+> b.nodes B" proof fix pq assume "pq \ c.nodes (union A B)" then show "pq \ a.nodes A <+> b.nodes B" using assms by (induct) (auto 0 3) qed lemma union_nodes_finite[intro]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" assumes "finite (a.nodes A)" "finite (b.nodes B)" shows "finite (c.nodes (union A B))" using finite_subset union_nodes assms by (auto intro: finite_Plus) end locale automaton_union_trace = automaton_union 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 set \ ('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 set" 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 \ 'label stream \ 'state\<^sub>1 stream \ 'state\<^sub>1 \ bool" and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('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 set" 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 \ 'label stream \ 'state\<^sub>2 stream \ 'state\<^sub>2 \ bool" and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 + 'state\<^sub>2) set \ ('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) set" 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 \ 'label stream \ ('state\<^sub>1 + 'state\<^sub>2) stream \ 'state\<^sub>1 + 'state\<^sub>2 \ bool" and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + assumes test\<^sub>1[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (smap Inl u) (Inl p) \ test\<^sub>1 c\<^sub>1 w u p" assumes test\<^sub>2[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (smap Inr v) (Inr q) \ test\<^sub>2 c\<^sub>2 w v q" begin lemma union_language[simp]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "c.language (union A B) = a.language A \ b.language B" using assms unfolding a.language_def b.language_def c.language_def by (auto dest: run_union_a run_union_b elim!: run_union) end locale automaton_union_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 set \ ('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 set" 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 \ (nat \ 'state) set \ ('label, nat \ 'state) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ (nat \ 'state) set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, nat \ 'state) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" + fixes condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" begin definition union :: "'automaton\<^sub>1 list \ 'automaton\<^sub>2" where "union AA \ automaton\<^sub>2 (\ (alphabet\<^sub>1 ` set AA)) (\ k < length AA. {k} \ initial\<^sub>1 (AA ! k)) (\ a (k, p). {k} \ transition\<^sub>1 (AA ! k) a p) (condition (map condition\<^sub>1 AA))" lemma union_simps[simp]: "alphabet\<^sub>2 (union AA) = \ (alphabet\<^sub>1 ` set AA)" "initial\<^sub>2 (union AA) = (\ k < length AA. {k} \ initial\<^sub>1 (AA ! k))" "transition\<^sub>2 (union AA) a (k, p) = {k} \ transition\<^sub>1 (AA ! k) a p" "condition\<^sub>2 (union AA) = condition (map condition\<^sub>1 AA)" unfolding union_def by auto - lemma union_run: + lemma run_union: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" assumes "A \ set AA" assumes "a.run A (w ||| s) p" obtains k where "k < length AA" "A = AA ! k" "b.run (union AA) (w ||| sconst k ||| s) (k, p)" proof - obtain k where 1: "k < length AA" "A = AA ! k" using assms(2) unfolding set_conv_nth by auto show ?thesis proof show "k < length AA" "A = AA ! k" using 1 by this show "b.run (union AA) (w ||| sconst k ||| s) (k, p)" using assms 1(2) by (coinduction arbitrary: w s p) (force elim: a.run.cases) qed qed - lemma run_union: + lemma union_run: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" assumes "k < length AA" assumes "b.run (union AA) (w ||| r) (k, p)" obtains s where "r = sconst k ||| s" "a.run (AA ! k) (w ||| s) p" proof show "r = sconst k ||| smap snd r" using assms by (coinduction arbitrary: w r p) (force elim: b.run.cases) show "a.run (AA ! k) (w ||| smap snd r) p" using assms by (coinduction arbitrary: w r p) (force elim: b.run.cases) qed lemma union_nodes: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" shows "b.nodes (union AA) \ (\ k < length AA. {k} \ a.nodes (AA ! k))" proof show "kp \ (\ k < length AA. {k} \ a.nodes (AA ! k))" if "kp \ b.nodes (union AA)" for kp using that assms by (induct) (auto 0 4) qed lemma union_nodes_finite[intro]: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" assumes "list_all (finite \ a.nodes) AA" shows "finite (b.nodes (union AA))" proof (rule finite_subset) show "b.nodes (union AA) \ (\ k < length AA. {k} \ a.nodes (AA ! k))" using union_nodes assms(1) by this show "finite (\ k < length AA. {k} \ a.nodes' (AA ! k))" using assms(2) unfolding list_all_length by auto qed end locale automaton_union_list_trace = automaton_union_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 set \ ('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 set" 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 \ 'label stream \ 'state stream \ 'state \ bool" and automaton\<^sub>2 :: "'label set \ (nat \ 'state) set \ ('label, nat \ 'state) trans \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 :: "'automaton\<^sub>2 \ 'label set" and initial\<^sub>2 :: "'automaton\<^sub>2 \ (nat \ 'state) set" and transition\<^sub>2 :: "'automaton\<^sub>2 \ ('label, nat \ 'state) trans" and condition\<^sub>2 :: "'automaton\<^sub>2 \ 'condition\<^sub>2" and test\<^sub>2 :: "'condition\<^sub>2 \ 'label stream \ (nat \ 'state) stream \ nat \ 'state \ bool" and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + - (* TODO: why do we not need wellformedness here but we do in the intersection_list case? *) - assumes test[iff]: "test\<^sub>2 (condition cs) w (sconst k ||| r) (k, p) \ test\<^sub>1 (cs ! k) w r p" + assumes test[iff]: "k < length cs \ test\<^sub>2 (condition cs) w (sconst k ||| r) (k, p) \ test\<^sub>1 (cs ! k) w r p" begin lemma union_language[simp]: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" shows "b.language (union AA) = \ (a.language ` set AA)" proof show "b.language (union AA) \ \ (a.language ` set AA)" - using assms unfolding a.language_def b.language_def by (force elim: run_union) + using assms unfolding a.language_def b.language_def by (force elim: union_run) show "\ (a.language ` set AA) \ b.language (union AA)" - using assms unfolding a.language_def b.language_def by (force elim!: union_run) + using assms unfolding a.language_def b.language_def by (force elim!: run_union) qed end end \ No newline at end of file