diff --git a/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy b/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy --- a/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy +++ b/thys/Transition_Systems_and_Automata/Automata/Deterministic.thy @@ -1,564 +1,560 @@ 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 - (* TODO: maybe this should be stated in terms of the lists and streams constants (adjust language_alphabet proof) *) - lemma path_alt_def: "path A w p \ set w \ alphabet A" - unfolding lists_iff_set[symmetric] + lemma path_alt_def: "path A w p \ w \ lists (alphabet A)" 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] + lemma run_alt_def: "run A w p \ w \ streams (alphabet A)" 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_path) *) locale automaton_run = 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 \ 'label stream \ 'state stream \ 'state \ bool" begin definition language :: "'automaton \ 'label stream set" where "language A \ {w. run A w (initial A) \ test (condition A) w (trace A w (initial A)) (initial A)}" lemma language[intro]: assumes "run A w (initial A)" "test (condition A) w (trace A w (initial A)) (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) w (trace A w (initial A)) (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 + lemma language_alphabet: "language A \ streams (alphabet A)" using run_alt_def 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 \ 'item 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 \ 'item pred gen" and automaton\<^sub>2 :: "'label set \ 'state degen \ ('label, 'state degen) trans \ 'item_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 \ 'item_degen pred" + fixes item :: "'state \ 'label \ 'state \ 'item" fixes translate :: "'item_degen \ 'item degen" 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) (item (p, a, transition\<^sub>1 A a p)) k)) (degen (condition\<^sub>1 A) \ translate)" 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) (item (p, a, transition\<^sub>1 A a p)) k)" "condition\<^sub>2 (degeneralize A) = degen (condition\<^sub>1 A) \ translate" 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) \ item) (p # a.states A w p || w || 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) \ item) (p # a.states A w p || w || 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) \ item) (p ## a.trace A w p ||| w ||| a.trace A w p) k" by (coinduction arbitrary: w p k) (auto, metis sscan.code) 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 finite_subset degeneralize_nodes that by blast 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_run = 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 item translate + a: automaton_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run 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 \ 'item 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 \ 'item pred gen" and test\<^sub>1 :: "'item pred gen \ 'label stream \ 'state stream \ 'state \ bool" and automaton\<^sub>2 :: "'label set \ 'state degen \ ('label, 'state degen) trans \ 'item_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 \ 'item_degen pred" and test\<^sub>2 :: "'item_degen pred \ 'label stream \ 'state degen stream \ 'state degen \ bool" and item :: "'state \ 'label \ 'state \ 'item" and translate :: "'item_degen \ 'item degen" + assumes test[iff]: "test\<^sub>2 (degen cs \ translate) w (r ||| sscan (count cs \ item) (p ## r ||| w ||| r) k) (p, k) \ test\<^sub>1 cs w r p" begin lemma degeneralize_language[simp]: "b.language (degeneralize A) = a.language A" by force 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_run = 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_run 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 \ 'label stream \ 'state\<^sub>1 stream \ 'state\<^sub>1 \ 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 \ 'label stream \ 'state\<^sub>2 stream \ 'state\<^sub>2 \ 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 \ '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 (r ||| s) (p, q) \ test\<^sub>1 c\<^sub>1 w r p \ test\<^sub>2 c\<^sub>2 w s q" begin lemma combine_language[simp]: "c.language (combine A B) = a.language A \ b.language B" by force end locale automaton_union_run = 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_run 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 \ 'label stream \ 'state\<^sub>1 stream \ 'state\<^sub>1 \ 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 \ 'label stream \ 'state\<^sub>2 stream \ 'state\<^sub>2 \ 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 \ '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 (r ||| s) (p, q) \ test\<^sub>1 c\<^sub>1 w r p \ test\<^sub>2 c\<^sub>2 w s q" 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) + (\ a ps. map2 (\ A p. transition\<^sub>1 A a p) AA ps) (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" + "transition\<^sub>2 (combine AA) a ps = map2 (\ A p. transition\<^sub>1 A a p) AA ps" "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 and listset *) 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) + assumes "length ps = length AA" "k < length AA" + shows "smap (\ ps. ps ! k) (b.trace (combine AA) w ps) = a.trace (AA ! k) w (ps ! k)" + using assms by (coinduction arbitrary: w ps) (force) lemma combine_nodes_length: - assumes "pp \ b.nodes (combine AA)" - shows "length pp = length AA" + assumes "ps \ b.nodes (combine AA)" + shows "length ps = 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)" + assumes "ps \ b.nodes (combine AA)" "k < length ps" + shows "ps ! 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) (* TODO: this is used more than once, make top level theorem *) show "b.nodes (combine AA) \ listset (map a.nodes AA)" by (force simp: listset_member list_all2_conv_all_nth combine_nodes_length) show "finite (listset (map a.nodes AA))" using list.pred_map assms by auto 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) show "finite (listset (map a.nodes AA))" using list.pred_map assms by auto 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_run = 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run 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 \ 'label stream \ 'state stream \ 'state \ 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 \ 'label stream \ 'state list stream \ 'state list \ bool" and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + assumes test[iff]: "test\<^sub>2 (condition cs) w rs ps \ (\ k < length cs. test\<^sub>1 (cs ! k) w (smap (\ ps. ps ! k) rs) (ps ! k))" 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 + unfolding a.run_alt_def b.run_alt_def streams_iff_sset by (fastforce simp: set_conv_nth combine_trace_smap) end locale automaton_union_list_run = 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run 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 \ 'label stream \ 'state stream \ 'state \ 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 \ 'label stream \ 'state list stream \ 'state list \ bool" and condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" + assumes test[iff]: "test\<^sub>2 (condition cs) w rs ps \ (\ k < length cs. test\<^sub>1 (cs ! k) w (smap (\ ps. ps ! k) rs) (ps ! k))" 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 + unfolding a.run_alt_def b.run_alt_def streams_iff_sset by (fastforce simp: set_conv_nth combine_trace_smap) end end \ No newline at end of file diff --git a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy --- a/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy +++ b/thys/Transition_Systems_and_Automata/Automata/NBA/NBA_Combine.thy @@ -1,90 +1,90 @@ section \Nondeterministic Büchi Automata Combinations\ theory NBA_Combine imports NBA NGBA begin global_interpretation degeneralization: automaton_degeneralization_run ngba ngba.alphabet ngba.initial ngba.transition ngba.accepting "\ P w r p. gen infs P r" nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" fst id defines degeneralize = degeneralization.degeneralize - by unfold_locales auto + by (unfold_locales) (auto simp flip: sscan_smap) lemmas degeneralize_language[simp] = degeneralization.degeneralize_language[folded NBA.language_def] lemmas degeneralize_nodes_finite[iff] = degeneralization.degeneralize_nodes_finite[folded NBA.nodes_def] global_interpretation intersection: automaton_intersection_run nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" ngba ngba.alphabet ngba.initial ngba.transition ngba.accepting "\ P w r p. gen infs P r" "\ c\<^sub>1 c\<^sub>2. [c\<^sub>1 \ fst, c\<^sub>2 \ snd]" defines intersect' = intersection.intersect by unfold_locales auto lemmas intersect'_language[simp] = intersection.intersect_language[folded NGBA.language_def] lemmas intersect'_nodes_finite[intro] = intersection.intersect_nodes_finite[folded NGBA.nodes_def] + global_interpretation union: automaton_union_run + nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" + nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" + nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" + case_sum + defines union = union.union + by (unfold_locales) (auto simp: comp_def) + + lemmas union_language = union.union_language + lemmas union_nodes_finite = union.union_nodes_finite + global_interpretation intersection_list: automaton_intersection_list_run nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" ngba ngba.alphabet ngba.initial ngba.transition ngba.accepting "\ P w r p. gen infs P r" "\ cs. map (\ k ps. (cs ! k) (ps ! k)) [0 ..< length cs]" defines intersect_list' = intersection_list.intersect proof unfold_locales fix cs :: "('b \ bool) list" and rs :: "'b stream list" and w :: "'a stream" and ps :: "'b list" assume 1: "length rs = length cs" "length ps = length cs" have "gen infs (map (\ k pp. (cs ! k) (pp ! k)) [0 ..< length cs]) (stranspose rs) \ (\ k < length cs. infs (\ pp. (cs ! k) (pp ! k)) (stranspose rs))" by (auto simp: gen_def) also have "\ \ (\ k < length cs. infs (cs ! k) (smap (\ pp. pp ! k) (stranspose rs)))" by (simp add: comp_def) also have "\ \ (\ k < length cs. infs (cs ! k) (rs ! k))" using 1 by simp also have "\ \ list_all (\ (c, r, p). infs c r) (cs || rs || ps)" using 1 unfolding list_all_length by simp finally show "gen infs (map (\ k ps. (cs ! k) (ps ! k)) [0 ..< length cs]) (stranspose rs) \ list_all (\ (c, r, p). infs c r) (cs || rs || ps)" by this qed lemmas intersect_list'_language[simp] = intersection_list.intersect_language[folded NGBA.language_def] lemmas intersect_list'_nodes_finite[intro] = intersection_list.intersect_nodes_finite[folded NGBA.nodes_def] - global_interpretation union: automaton_union_run - nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" - nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" - nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" - case_sum - defines union = union.union - by (unfold_locales) (auto simp: comp_def) - - lemmas union_language = union.union_language - lemmas union_nodes_finite = union.union_nodes_finite - global_interpretation union_list: automaton_union_list_run nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" nba nba.alphabet nba.initial nba.transition nba.accepting "\ P w r p. infs P r" "\ cs (k, p). (cs ! k) p" defines union_list = union_list.union by (unfold_locales) (auto simp: szip_sconst_smap_fst comp_def) lemmas union_list_language = union_list.union_language lemmas union_list_nodes_finite = union_list.union_nodes_finite abbreviation intersect where "intersect A B \ degeneralize (intersect' A B)" lemma intersect_language[simp]: "NBA.language (intersect A B) = NBA.language A \ NBA.language B" by simp lemma intersect_nodes_finite[intro]: assumes "finite (NBA.nodes A)" "finite (NBA.nodes B)" shows "finite (NBA.nodes (intersect A B))" using intersect'_nodes_finite assms by simp abbreviation intersect_list where "intersect_list AA \ degeneralize (intersect_list' AA)" lemma intersect_list_language[simp]: "NBA.language (intersect_list AA) = \ (NBA.language ` set AA)" by simp lemma intersect_list_nodes_finite[intro]: assumes "list_all (finite \ NBA.nodes) AA" shows "finite (NBA.nodes (intersect_list AA))" using intersect_list'_nodes_finite assms by simp 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,895 +1,893 @@ theory Nondeterministic imports "../Transition_Systems/Transition_System" "../Transition_Systems/Transition_System_Extra" "../Transition_Systems/Transition_System_Construction" "../Basic/Degeneralization" begin (* TODO: might not be needed anymore *) type_synonym ('label, 'state) trans = "'label \ 'state \ 'state set" (* TODO: many of these type annotations can be left out *) 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 path_alphabet: assumes "length r = length w" "path A (w || r) p" shows "w \ lists (alphabet A)" using assms by (induct arbitrary: p rule: list_induct2) (auto) 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 locale automaton_path = 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 list \ 'state list \ 'state \ bool" begin definition language :: "'automaton \ 'label list set" where "language A \ {w |w r p. length r = length w \ p \ initial A \ path A (w || r) p \ test (condition A) w r p}" lemma language[intro]: assumes "length r = length w" "p \ initial A" "path 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 "length r = length w" "p \ initial A" "path A (w || r) p" "test (condition A) w r p" using assms unfolding language_def by auto lemma language_alphabet: "language A \ lists (alphabet A)" by (auto dest: path_alphabet) lemma restrict_language[simp]: "language (restrict A) = language A" by force end locale automaton_run = 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 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 \ 'item 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 \ 'item pred gen" and automaton\<^sub>2 :: "'label set \ 'state degen set \ ('label, 'state degen) trans \ 'item_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 \ 'item_degen pred" + fixes item :: "'state \ 'label \ 'state \ 'item" fixes translate :: "'item_degen \ 'item degen" 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). {(q, count (condition\<^sub>1 A) (item (p, a, q)) k) |q. q \ transition\<^sub>1 A a p}) (degen (condition\<^sub>1 A) \ translate)" 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) = {(q, count (condition\<^sub>1 A) (item (p, a, q)) k) |q. q \ transition\<^sub>1 A a p}" "condition\<^sub>2 (degeneralize A) = degen (condition\<^sub>1 A) \ translate" unfolding degeneralize_def by auto - abbreviation "counter cs w r p k \ sscan (count cs) (smap item (p ## r ||| w ||| r)) k" - lemma run_degeneralize: assumes "a.run A (w ||| r) p" - shows "b.run (degeneralize A) (w ||| r ||| counter (condition\<^sub>1 A) w r p k) (p, k)" + shows "b.run (degeneralize A) (w ||| r ||| sscan (count (condition\<^sub>1 A) \ item) (p ## r ||| w ||| 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 = counter (condition\<^sub>1 A) w r p k" + where "rs = r ||| s" "pk = (p, k)" "a.run A (w ||| r) p" "s = sscan (count (condition\<^sub>1 A) \ item) (p ## r ||| w ||| r) k" proof - obtain r s p k where 1: "rs = r ||| s" "pk = (p, k)" using szip_smap surjective_pairing by metis show ?thesis proof show "rs = r ||| s" "pk = (p, k)" using 1 by this show "a.run A (w ||| r) p" using assms unfolding 1 by (coinduction arbitrary: w r s p k) (force elim: b.run.cases) - show "s = counter (condition\<^sub>1 A) w r p k" + show "s = sscan (count (condition\<^sub>1 A) \ item) (p ## r ||| w ||| r) k" using assms unfolding 1 by (coinduction arbitrary: w r s p k) (erule b.run.cases, force) qed 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) (item (p, aq)) 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_run = 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 item translate + a: automaton_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run 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 \ 'item 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 \ 'item pred gen" and test\<^sub>1 :: "'item pred gen \ 'label stream \ 'state stream \ 'state \ bool" and automaton\<^sub>2 :: "'label set \ 'state degen set \ ('label, 'state degen) trans \ 'item_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 \ 'item_degen pred" and test\<^sub>2 :: "'item_degen pred \ 'label stream \ 'state degen stream \ 'state degen \ bool" and item :: "'state \ 'label \ 'state \ 'item" and translate :: "'item_degen \ 'item degen" + - assumes test[iff]: "test\<^sub>2 (degen cs \ translate) w (r ||| counter cs w r p k) (p, k) \ test\<^sub>1 cs w r p" + assumes test[iff]: "test\<^sub>2 (degen cs \ translate) w (r ||| sscan (count cs \ item) (p ## r ||| w ||| 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) 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_target[simp]: assumes "length w = length r" "length r = length s" shows "c.target (w || r || s) (p, q) = (a.target (w || r) p, b.target (w || s) q)" using assms by (induct arbitrary: p q rule: list_induct3) (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 (* TODO: format these the same as degeneralization *) locale automaton_intersection_path = 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_path automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_path automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_path 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 list \ 'state\<^sub>1 list \ '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 list \ 'state\<^sub>2 list \ '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 list \ ('state\<^sub>1 \ 'state\<^sub>2) list \ 'state\<^sub>1 \ 'state\<^sub>2 \ bool" and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + assumes test[iff]: "length u = length w \ length v = length w \ 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 (force iff: split_zip) end locale automaton_intersection_run = 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_run 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 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 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 ps = length AA" - assumes "b.run (intersect AA) (w ||| r) ps" - assumes "qs \ sset r" - shows "length qs = length AA" - proof - - 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 ps = length AA" - assumes "b.run (intersect AA) (w ||| r) ps" - obtains rs where "r = stranspose rs" "length rs = length AA" - proof - 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 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 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) 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 rs ! k) (map shd rs ! k)" - using run by (force elim: a.run.cases) - qed auto - qed - 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 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) (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 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 "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_run = - 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + - b: automaton_run 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)" - 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 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 "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 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 "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 path_union_a: assumes "length r = length w" "a.path A (w || r) p" shows "c.path (union A B) (w || map Inl r) (Inl p)" using assms by (induct arbitrary: p rule: list_induct2) (auto) lemma path_union_b: assumes "length s = length w" "b.path B (w || s) q" shows "c.path (union A B) (w || map Inr s) (Inr q)" using assms by (induct arbitrary: q rule: list_induct2) (auto) lemma union_path: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" assumes "length rs = length w" "c.path (union A B) (w || rs) pq" obtains (a) r p where "rs = map Inl r" "pq = Inl p" "a.path A (w || r) p" | (b) s q where "rs = map Inr s" "pq = Inr q" "b.path B (w || s) q" proof (cases pq) case (Inl p) have 1: "rs = map Inl (map projl rs)" using assms(2, 3) unfolding Inl by (induct arbitrary: p rule: list_induct2) (auto) have 2: "a.path A (w || map projl rs) p" using assms(2, 1, 3) unfolding Inl by (induct arbitrary: p rule: list_induct2) (auto) show ?thesis using a 1 Inl 2 by this next case (Inr q) have 1: "rs = map Inr (map projr rs)" using assms(2, 3) unfolding Inr by (induct arbitrary: q rule: list_induct2) (auto) have 2: "b.path B (w || map projr rs) q" using assms(2, 1, 3) unfolding Inr by (induct arbitrary: q rule: list_induct2) (auto) show ?thesis using b 1 Inr 2 by this qed 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 union_run: 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_path = 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_path automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_path automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_path 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 list \ 'state\<^sub>1 list \ '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 list \ 'state\<^sub>2 list \ '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 list \ ('state\<^sub>1 + 'state\<^sub>2) list \ 'state\<^sub>1 + 'state\<^sub>2 \ bool" and condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" + assumes test\<^sub>1[iff]: "length u = length w \ test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (map Inl u) (Inl p) \ test\<^sub>1 c\<^sub>1 w u p" assumes test\<^sub>2[iff]: "length v = length w \ test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (map 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 (force intro: path_union_a path_union_b elim!: union_path) end locale automaton_union_run = 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_run 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 intro: run_union_a run_union_b elim!: union_run) 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 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 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 ps = length AA" + assumes "b.run (intersect AA) (w ||| r) ps" + assumes "qs \ sset r" + shows "length qs = length AA" + proof - + 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 ps = length AA" + assumes "b.run (intersect AA) (w ||| r) ps" + obtains rs where "r = stranspose rs" "length rs = length AA" + proof + 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 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 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) 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 rs ! k) (map shd rs ! k)" + using run by (force elim: a.run.cases) + qed auto + qed + 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 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) (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 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 "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_run = + 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + + b: automaton_run 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)" + 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 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 "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 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 "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_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 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 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_run = 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_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run 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" + 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: union_run) show "\ (a.language ` set AA) \ b.language (union AA)" using assms unfolding a.language_def b.language_def by (force elim!: run_union) qed end end \ No newline at end of file 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,468 +1,470 @@ 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 list_choice: "list_all (\ x. \ y. P x y) xs \ (\ ys. list_all2 P xs ys)" by (induct xs) (auto simp: list_all2_Cons1) 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_finite'[intro]: assumes "list_all finite XS" shows "finite (listset XS)" using infinite_imp_nonempty assms by blast 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 streams_int[simp]: "streams (A \ B) = streams A \ streams B" by (auto iff: streams_iff_sset) + lemma streams_Int[simp]: "streams (\ S) = \ (streams ` S)" by (auto iff: streams_iff_sset) 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 \Transposing Streams\ primcorec (transfer) stranspose :: "'a stream list \ 'a list stream" where "stranspose ws = map shd ws ## stranspose (map stl ws)" lemma stranspose_eq_scons[iff]: "stranspose ws = a ## w \ map shd ws = a \ stranspose (map stl ws) = w" using stranspose.ctr stream.inject by metis lemma scons_eq_stranspose[iff]: "a ## w = stranspose ws \ a = map shd ws \ w = stranspose (map stl ws)" using stranspose.ctr stream.inject by metis lemma stranspose_nil[simp]: "stranspose [] = sconst []" by coinduction auto lemma stranspose_cons[simp]: "stranspose (w # ws) = smap2 Cons w (stranspose ws)" by (coinduction arbitrary: w ws) (metis list.simps(9) smap2.simps stranspose.simps stream.sel) lemma snth_stranspose[simp]: "stranspose ws !! k = map (\ w. w !! k) ws" by (induct k arbitrary: ws) (auto) lemma stranspose_nth[simp]: assumes "k < length ws" shows "smap (\ xs. xs ! k) (stranspose ws) = ws ! k" using assms by (auto intro: eqI_snth) 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