diff --git a/thys/Buchi_Complementation/Ranking.thy b/thys/Buchi_Complementation/Ranking.thy --- a/thys/Buchi_Complementation/Ranking.thy +++ b/thys/Buchi_Complementation/Ranking.thy @@ -1,479 +1,478 @@ section \Rankings\ theory Ranking imports "Alternate" "Graph" begin subsection \Rankings\ type_synonym 'state ranking = "'state node \ nat" definition ranking :: "('label, 'state) nba \ 'label stream \ 'state ranking \ bool" where "ranking A w f \ (\ v \ gunodes A w. f v \ 2 * card (nodes A)) \ (\ v \ gunodes A w. \ u \ gusuccessors A w v. f u \ f v) \ (\ v \ gunodes A w. gaccepting A v \ even (f v)) \ (\ v \ gunodes A w. \ r k. gurun A w r v \ smap f (gtrace r v) = sconst k \ odd k)" subsection \Ranking Implies Word not in Language\ lemma ranking_stuck: assumes "ranking A w f" assumes "v \ gunodes A w" "gurun A w r v" obtains n k where "smap f (gtrace (sdrop n r) (gtarget (stake n r) v)) = sconst k" proof - have 0: "f u \ f v" if "v \ gunodes A w" "u \ gusuccessors A w v" for v u using assms(1) that unfolding ranking_def by auto have 1: "shd (v ## gtrace r v) \ gunodes A w" using assms(2) by auto have 2: "sdescending (smap f (v ## gtrace r v))" using 1 assms(3) proof (coinduction arbitrary: r v rule: sdescending.coinduct) case sdescending obtain u s where 1: "r = u ## s" using stream.exhaust by blast have 2: "v \ gunodes A w" using sdescending(1) by simp have 3: "gurun A w (u ## s) v" using sdescending(2) 1 by auto have 4: "u \ gusuccessors A w v" using 3 by auto have 5: "u \ gureachable A w v" using graph.reachable_successors 4 by blast show ?case unfolding 1 proof (intro exI conjI disjI1) show "f u \ f v" using 0 2 4 by this show "shd (u ## gtrace s u) \ gunodes A w" using 2 5 by auto show "gurun A w s u" using 3 by auto qed auto qed - obtain n k where 3: "sdrop n (smap f (v ## gtrace r v)) = sconst k" + obtain s k where 3: "smap f (v ## gtrace r v) = s @- sconst k" using sdescending_stuck[OF 2] by metis - have "gtrace (sdrop (Suc n) r) (gtarget (stake (Suc n) r) v) = sdrop (Suc n) (gtrace r v)" + have "gtrace (sdrop (Suc (length s)) r) (gtarget (stake (Suc (length s)) r) v) = sdrop (Suc (length s)) (gtrace r v)" using sscan_sdrop by rule - also have "smap f \ = sdrop n (smap f (v ## gtrace r v))" - by (auto, metis 3 id_apply sdrop_smap sdrop_stl siterate.simps(2) - sscan_const stream.map stream.map_sel(2) stream.sel(2)) - also have "\ = sconst k" using 3 by this + also have "smap f \ = sdrop (length s) (smap f (v ## gtrace r v))" + by (metis "3" id_apply sdrop_simps(2) sdrop_smap sdrop_stl shift_eq siterate.simps(2) stream.sel(2)) + also have "\ = sconst k" unfolding 3 using shift_eq by metis finally show ?thesis using that by blast qed lemma ranking_stuck_odd: assumes "ranking A w f" assumes "v \ gunodes A w" "gurun A w r v" obtains n where "Ball (sset (smap f (gtrace (sdrop n r) (gtarget (stake n r) v)))) odd" proof - obtain n k where 1: "smap f (gtrace (sdrop n r) (gtarget (stake n r) v)) = sconst k" using ranking_stuck assms by this have 2: "gtarget (stake n r) v \ gunodes A w" using assms(2, 3) by (simp add: graph.nodes_target graph.run_stake) have 3: "gurun A w (sdrop n r) (gtarget (stake n r) v)" using assms(2, 3) by (simp add: graph.run_sdrop) have 4: "odd k" using 1 2 3 assms(1) unfolding ranking_def by meson have 5: "Ball (sset (smap f (gtrace (sdrop n r) (gtarget (stake n r) v)))) odd" unfolding 1 using 4 by simp show ?thesis using that 5 by this qed lemma ranking_language: assumes "ranking A w f" shows "w \ language A" proof assume 1: "w \ language A" obtain r p where 2: "run A (w ||| r) p" "p \ initial A" "infs (accepting A) (p ## r)" using 1 by rule let ?r = "fromN 1 ||| r" let ?v = "(0, p)" have 3: "?v \ gunodes A w" "gurun A w ?r ?v" using 2(1, 2) by (auto intro: run_grun) obtain n where 4: "Ball (sset (smap f (gtrace (sdrop n ?r) (gtarget (stake n ?r) ?v)))) odd" using ranking_stuck_odd assms 3 by this let ?s = "stake n ?r" let ?t = "sdrop n ?r" let ?u = "gtarget ?s ?v" have "sset (gtrace ?t ?u) \ gureachable A w ?v" proof (intro graph.reachable_trace graph.reachable_target graph.reachable.reflexive) show "gupath A w ?s ?v" using graph.run_stake 3(2) by this show "gurun A w ?t ?u" using graph.run_sdrop 3(2) by this qed also have "\ \ gunodes A w" using 3(1) by blast finally have 7: "sset (gtrace ?t ?u) \ gunodes A w" by this have 8: "\ p. p \ gunodes A w \ gaccepting A p \ even (f p)" using assms unfolding ranking_def by auto have 9: "\ p. p \ sset (gtrace ?t ?u) \ gaccepting A p \ even (f p)" using 7 8 by auto have 19: "infs (accepting A) (smap snd ?r)" using 2(3) by simp have 18: "infs (gaccepting A) ?r" using 19 by simp have 17: "infs (gaccepting A) (gtrace ?r ?v)" using 18 unfolding gtrace_alt_def by this have 16: "infs (gaccepting A) (gtrace (?s @- ?t) ?v)" using 17 unfolding stake_sdrop by this have 15: "infs (gaccepting A) (gtrace ?t ?u)" using 16 by simp have 13: "infs (even \ f) (gtrace ?t ?u)" using infs_mono[OF _ 15] 9 by simp have 12: "infs even (smap f (gtrace ?t ?u))" using 13 by (simp add: comp_def) have 11: "Bex (sset (smap f (gtrace ?t ?u))) even" using 12 infs_any by metis show False using 4 11 by auto qed subsection \Word not in Language Implies Ranking\ subsubsection \Removal of Endangered Nodes\ definition clean :: "('label, 'state) nba \ 'label stream \ 'state node set \ 'state node set" where "clean A w V \ {v \ V. infinite (greachable A w V v)}" lemma clean_decreasing: "clean A w V \ V" unfolding clean_def by auto lemma clean_successors: assumes "v \ V" "u \ gusuccessors A w v" shows "u \ clean A w V \ v \ clean A w V" proof - assume 1: "u \ clean A w V" have 2: "u \ V" "infinite (greachable A w V u)" using 1 unfolding clean_def by auto have 3: "u \ greachable A w V v" using graph.reachable.execute assms(2) 2(1) by blast have 4: "greachable A w V u \ greachable A w V v" using 3 by blast have 5: "infinite (greachable A w V v)" using 2(2) 4 by (simp add: infinite_super) show "v \ clean A w V" unfolding clean_def using assms(1) 5 by simp qed subsubsection \Removal of Safe Nodes\ definition prune :: "('label, 'state) nba \ 'label stream \ 'state node set \ 'state node set" where "prune A w V \ {v \ V. \ u \ greachable A w V v. gaccepting A u}" lemma prune_decreasing: "prune A w V \ V" unfolding prune_def by auto lemma prune_successors: assumes "v \ V" "u \ gusuccessors A w v" shows "u \ prune A w V \ v \ prune A w V" proof - assume 1: "u \ prune A w V" have 2: "u \ V" "\ x \ greachable A w V u. gaccepting A x" using 1 unfolding prune_def by auto have 3: "u \ greachable A w V v" using graph.reachable.execute assms(2) 2(1) by blast have 4: "greachable A w V u \ greachable A w V v" using 3 by blast show "v \ prune A w V" unfolding prune_def using assms(1) 2(2) 4 by auto qed subsubsection \Run Graph Interation\ definition graph :: "('label, 'state) nba \ 'label stream \ nat \ 'state node set" where "graph A w k \ alternate (clean A w) (prune A w) k (gunodes A w)" abbreviation "level A w k l \ {v \ graph A w k. fst v = l}" lemma graph_0[simp]: "graph A w 0 = gunodes A w" unfolding graph_def by simp lemma graph_Suc[simp]: "graph A w (Suc k) = (if even k then clean A w else prune A w) (graph A w k)" unfolding graph_def by simp lemma graph_antimono: "antimono (graph A w)" using alternate_antimono clean_decreasing prune_decreasing unfolding antimono_def le_fun_def graph_def by metis lemma graph_nodes: "graph A w k \ gunodes A w" using graph_0 graph_antimono le0 antimonoD by metis lemma graph_successors: assumes "v \ gunodes A w" "u \ gusuccessors A w v" shows "u \ graph A w k \ v \ graph A w k" using assms proof (induct k arbitrary: u v) case 0 show ?case using 0(2) by simp next case (Suc k) have 1: "v \ graph A w k" using Suc using antimono_iff_le_Suc graph_antimono rev_subsetD by blast show ?case using Suc(2) clean_successors[OF 1 Suc(4)] prune_successors[OF 1 Suc(4)] by auto qed lemma graph_level_finite: assumes "finite (nodes A)" shows "finite (level A w k l)" proof - have "level A w k l \ {v \ gunodes A w. fst v = l}" by (simp add: graph_nodes subset_CollectI) also have "{v \ gunodes A w. fst v = l} \ {l} \ nodes A" using gnodes_nodes by force also have "finite ({l} \ nodes A)" using assms(1) by simp finally show ?thesis by this qed lemma find_safe: assumes "w \ language A" assumes "V \ {}" "V \ gunodes A w" assumes "\ v. v \ V \ gsuccessors A w V v \ {}" obtains v where "v \ V" "\ u \ greachable A w V v. \ gaccepting A u" proof (rule ccontr) assume 1: "\ thesis" have 2: "\ v. v \ V \ \ u \ greachable A w V v. gaccepting A u" using that 1 by auto have 3: "\ r v. v \ initial A \ run A (w ||| r) v \ fins (accepting A) r" using assms(1) by auto obtain v where 4: "v \ V" using assms(2) by force obtain x where 5: "x \ greachable A w V v" "gaccepting A x" using 2 4 by blast obtain y where 50: "gpath A w V y v" "x = gtarget y v" using 5(1) by rule obtain r where 6: "grun A w V r x" "infs (\ x. x \ V \ gaccepting A x) r" proof (rule graph.recurring_condition) show "x \ V \ gaccepting A x" using greachable_subset 4 5 by blast next fix v assume 1: "v \ V \ gaccepting A v" obtain v' where 20: "v' \ gsuccessors A w V v" using assms(4) 1 by (meson IntE equals0I) have 21: "v' \ V" using 20 by auto have 22: "\ u \ greachable A w V v'. u \ V \ gaccepting A u" using greachable_subset 2 21 by blast obtain r where 30: "gpath A w V r v'" "gtarget r v' \ V \ gaccepting A (gtarget r v')" using 22 by blast show "\ r. r \ [] \ gpath A w V r v \ gtarget r v \ V \ gaccepting A (gtarget r v)" proof (intro exI conjI) show "v' # r \ []" by simp show "gpath A w V (v' # r) v" using 20 30 by auto show "gtarget (v' # r) v \ V" using 30 by simp show "gaccepting A (gtarget (v' # r) v)" using 30 by simp qed qed auto obtain u where 100: "u \ ginitial A" "v \ gureachable A w u" using 4 assms(3) by blast have 101: "gupath A w y v" using gpath_subset 50(1) subset_UNIV by this have 102: "gurun A w r x" using grun_subset 6(1) subset_UNIV by this obtain t where 103: "gupath A w t u" "v = gtarget t u" using 100(2) by rule have 104: "gurun A w (t @- y @- r) u" using 101 102 103 50(2) by auto obtain s q where 7: "t @- y @- r = fromN (Suc 0) ||| s" "u = (0, q)" using grun_elim[OF 104] 100(1) by blast have 8: "run A (w ||| s) q" using grun_run[OF 104[unfolded 7]] by simp have 9: "q \ initial A" using 100(1) 7(2) by auto have 91: "sset (trace (w ||| s) q) \ reachable A q" using nba.reachable_trace nba.reachable.reflexive 8 by this have 10: "fins (accepting A) s" using 3 9 8 by this have 12: "infs (gaccepting A) r" using infs_mono[OF _ 6(2)] by simp have "s = smap snd (t @- y @- r)" unfolding 7(1) by simp also have "infs (accepting A) \" using 12 by (simp add: comp_def) finally have 13: "infs (accepting A) s" by this show False using 10 13 by simp qed lemma remove_run: assumes "finite (nodes A)" "w \ language A" assumes "V \ gunodes A w" "clean A w V \ {}" obtains v r where "grun A w V r v" "sset (gtrace r v) \ clean A w V" "sset (gtrace r v) \ - prune A w (clean A w V)" proof - obtain u where 1: "u \ clean A w V" "\ x \ greachable A w (clean A w V) u. \ gaccepting A x" proof (rule find_safe) show "w \ language A" using assms(2) by this show "clean A w V \ {}" using assms(4) by this show "clean A w V \ gunodes A w" using assms(3) by (meson clean_decreasing subset_iff) next fix v assume 1: "v \ clean A w V" have 2: "v \ V" using 1 clean_decreasing by blast have 3: "infinite (greachable A w V v)" using 1 clean_def by auto have "gsuccessors A w V v \ gusuccessors A w v" by auto also have "finite \" using 2 assms(1, 3) finite_nodes_gsuccessors by blast finally have 4: "finite (gsuccessors A w V v)" by this have 5: "infinite (insert v (\((greachable A w V) ` (gsuccessors A w V v))))" using graph.reachable_step 3 by metis obtain u where 6: "u \ gsuccessors A w V v" "infinite (greachable A w V u)" using 4 5 by auto have 7: "u \ clean A w V" using 6 unfolding clean_def by auto show "gsuccessors A w (clean A w V) v \ {}" using 6(1) 7 by auto qed auto have 2: "u \ V" using 1(1) unfolding clean_def by auto have 3: "infinite (greachable A w V u)" using 1(1) unfolding clean_def by simp have 4: "finite (gsuccessors A w V v)" if "v \ greachable A w V u" for v proof - have 1: "v \ V" using that greachable_subset 2 by blast have "gsuccessors A w V v \ gusuccessors A w v" by auto also have "finite \" using 1 assms(1, 3) finite_nodes_gsuccessors by blast finally show ?thesis by this qed obtain r where 5: "grun A w V r u" using graph.koenig[OF 3 4] by this have 6: "greachable A w V u \ V" using 2 greachable_subset by blast have 7: "sset (gtrace r u) \ V" using graph.reachable_trace[OF graph.reachable.reflexive 5(1)] 6 by blast have 8: "sset (gtrace r u) \ clean A w V" unfolding clean_def using 7 infinite_greachable_gtrace[OF 5(1)] by auto have 9: "sset (gtrace r u) \ greachable A w (clean A w V) u" using 5 8 by (metis graph.reachable.reflexive graph.reachable_trace grun_subset) show ?thesis proof show "grun A w V r u" using 5(1) by this show "sset (gtrace r u) \ clean A w V" using 8 by this show "sset (gtrace r u) \ - prune A w (clean A w V)" proof (intro subsetI ComplI) fix p assume 10: "p \ sset (gtrace r u)" "p \ prune A w (clean A w V)" have 20: "\ x \ greachable A w (clean A w V) p. gaccepting A x" using 10(2) unfolding prune_def by auto have 30: "greachable A w (clean A w V) p \ greachable A w (clean A w V) u" using 10(1) 9 by blast show "False" using 1(2) 20 30 by force qed qed qed lemma level_bounded: assumes "finite (nodes A)" "w \ language A" obtains n where "\ l. l \ n \ card (level A w (2 * k) l) \ card (nodes A) - k" proof (induct k arbitrary: thesis) case (0) show ?case proof (rule 0) fix l :: nat have "finite ({l} \ nodes A)" using assms(1) by simp also have "level A w 0 l \ {l} \ nodes A" using gnodes_nodes by force also (card_mono) have "card \ = card (nodes A)" using assms(1) by simp finally show "card (level A w (2 * 0) l) \ card (nodes A) - 0" by simp qed next case (Suc k) show ?case proof (cases "graph A w (Suc (2 * k)) = {}") case True have 3: "graph A w (2 * Suc k) = {}" using True prune_decreasing by simp blast show ?thesis using Suc(2) 3 by simp next case False obtain v r where 1: "grun A w (graph A w (2 * k)) r v" "sset (gtrace r v) \ graph A w (Suc (2 * k))" "sset (gtrace r v) \ - graph A w (Suc (Suc (2 * k)))" proof (rule remove_run) show "finite (nodes A)" "w \ language A" using assms by this show "clean A w (graph A w (2 * k)) \ {}" using False by simp show "graph A w (2 * k) \ gunodes A w" using graph_nodes by this qed auto obtain l q where 2: "v = (l, q)" by force obtain n where 90: "\ l. n \ l \ card (level A w (2 * k) l) \ card (nodes A) - k" using Suc(1) by blast show ?thesis proof (rule Suc(2)) fix j assume 100: "n + Suc l \ j" have 6: "graph A w (Suc (Suc (2 * k))) \ graph A w (Suc (2 * k))" using graph_antimono antimono_iff_le_Suc by blast have 101: "gtrace r v !! (j - Suc l) \ graph A w (Suc (2 * k))" using 1(2) snth_sset by auto have 102: "gtrace r v !! (j - Suc l) \ graph A w (Suc (Suc (2 * k)))" using 1(3) snth_sset by blast have 103: "gtrace r v !! (j - Suc l) \ level A w (Suc (2 * k)) j" using 1(1) 100 101 2 by (auto elim: grun_elim) have 104: "gtrace r v !! (j - Suc l) \ level A w (Suc (Suc (2 * k))) j" using 100 102 by simp have "level A w (2 * Suc k) j = level A w (Suc (Suc (2 * k))) j" by simp also have "\ \ level A w (Suc (2 * k)) j" using 103 104 6 by blast also have "\ \ level A w (2 * k) j" by (simp add: Collect_mono clean_def) finally have 105: "level A w (2 * Suc k) j \ level A w (2 * k) j" by this have "card (level A w (2 * Suc k) j) < card (level A w (2 * k) j)" using assms(1) 105 by (simp add: graph_level_finite psubset_card_mono) also have "\ \ card (nodes A) - k" using 90 100 by simp finally show "card (level A w (2 * Suc k) j) \ card (nodes A) - Suc k" by simp qed qed qed lemma graph_empty: assumes "finite (nodes A)" "w \ language A" shows "graph A w (Suc (2 * card (nodes A))) = {}" proof - obtain n where 1: "\ l. l \ n \ card (level A w (2 * card (nodes A)) l) = 0" using level_bounded[OF assms(1, 2), of "card (nodes A)"] by auto have "graph A w (2 * card (nodes A)) = (\ l \ {..< n}. level A w (2 * card (nodes A)) l) \ (\ l \ {n ..}. level A w (2 * card (nodes A)) l)" by auto also have "(\ l \ {n ..}. level A w (2 * card (nodes A)) l) = {}" using graph_level_finite assms(1) 1 by fastforce also have "finite ((\ l \ {..< n}. level A w (2 * card (nodes A)) l) \ {})" using graph_level_finite assms(1) by auto finally have 100: "finite (graph A w (2 * card (nodes A)))" by this have 101: "finite (greachable A w (graph A w (2 * card (nodes A))) v)" for v using 100 greachable_subset[of A w "graph A w (2 * card (nodes A))" v] using finite_insert infinite_super by auto show ?thesis using 101 by (simp add: clean_def) qed lemma graph_le: assumes "finite (nodes A)" "w \ language A" assumes "v \ graph A w k" shows "k \ 2 * card (nodes A)" using graph_empty graph_antimono assms by (metis (no_types, lifting) Suc_leI antimono_def basic_trans_rules(30) empty_iff not_le_imp_less) subsection \Node Ranks\ definition rank :: "('label, 'state) nba \ 'label stream \ 'state node \ nat" where "rank A w v \ GREATEST k. v \ graph A w k" lemma rank_member: assumes "finite (nodes A)" "w \ language A" "v \ gunodes A w" shows "v \ graph A w (rank A w v)" unfolding rank_def proof (rule GreatestI_nat) show "v \ graph A w 0" using assms(3) by simp show "k \ 2 * card (nodes A)" if "v \ graph A w k" for k using graph_le assms(1, 2) that by blast qed lemma rank_removed: assumes "finite (nodes A)" "w \ language A" shows "v \ graph A w (Suc (rank A w v))" proof assume "v \ graph A w (Suc (rank A w v))" then have 2: "Suc (rank A w v) \ rank A w v" unfolding rank_def using Greatest_le_nat graph_le assms by metis then show "False" by auto qed lemma rank_le: assumes "finite (nodes A)" "w \ language A" assumes "v \ gunodes A w" "u \ gusuccessors A w v" shows "rank A w u \ rank A w v" unfolding rank_def proof (rule Greatest_le_nat) have 1: "u \ gureachable A w v" using graph.reachable_successors assms(4) by blast have 2: "u \ gunodes A w" using assms(3) 1 by auto show "v \ graph A w (GREATEST k. u \ graph A w k)" unfolding rank_def[symmetric] proof (rule graph_successors) show "v \ gunodes A w" using assms(3) by this show "u \ gusuccessors A w v" using assms(4) by this show "u \ graph A w (rank A w u)" using rank_member assms(1, 2) 2 by this qed show "k \ 2 * card (nodes A)" if "v \ graph A w k" for k using graph_le assms(1, 2) that by blast qed lemma language_ranking: assumes "finite (nodes A)" "w \ language A" shows "ranking A w (rank A w)" unfolding ranking_def proof (intro conjI ballI allI impI) fix v assume 1: "v \ gunodes A w" have 2: "v \ graph A w (rank A w v)" using rank_member assms 1 by this show "rank A w v \ 2 * card (nodes A)" using graph_le assms 2 by this next fix v u assume 1: "v \ gunodes A w" "u \ gusuccessors A w v" show "rank A w u \ rank A w v" using rank_le assms 1 by this next fix v assume 1: "v \ gunodes A w" "gaccepting A v" have 2: "v \ graph A w (rank A w v)" using rank_member assms 1(1) by this have 3: "v \ graph A w (Suc (rank A w v))" using rank_removed assms by this have 4: "v \ prune A w (graph A w (rank A w v))" using 2 1(2) unfolding prune_def by auto have 5: "graph A w (Suc (rank A w v)) \ prune A w (graph A w (rank A w v))" using 3 4 by blast show "even (rank A w v)" using 5 by auto next fix v r k assume 1: "v \ gunodes A w" "gurun A w r v" "smap (rank A w) (gtrace r v) = sconst k" have "sset (gtrace r v) \ gureachable A w v" using 1(2) by (metis graph.reachable.reflexive graph.reachable_trace) then have 6: "sset (gtrace r v) \ gunodes A w" using 1(1) by blast have 60: "rank A w ` sset (gtrace r v) \ {k}" using 1(3) by (metis equalityD1 sset_sconst stream.set_map) have 50: "sset (gtrace r v) \ graph A w k" using rank_member[OF assms] subsetD[OF 6] 60 unfolding image_subset_iff by auto have 70: "grun A w (graph A w k) r v" using grun_subset 1(2) 50 by this have 7: "sset (gtrace r v) \ clean A w (graph A w k)" unfolding clean_def using 50 infinite_greachable_gtrace[OF 70] by auto have 8: "sset (gtrace r v) \ graph A w (Suc k) = {}" using rank_removed[OF assms] 60 by blast have 9: "sset (gtrace r v) \ {}" using stream.set_sel(1) by auto have 10: "graph A w (Suc k) \ clean A w (graph A w k)" using 7 8 9 by blast show "odd k" using 10 unfolding graph_Suc by auto qed subsection \Correctness Theorem\ theorem language_ranking_iff: assumes "finite (nodes A)" shows "w \ language A \ (\ f. ranking A w f)" using ranking_language language_ranking assms by blast end diff --git a/thys/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,470 +1,491 @@ 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_scons: + assumes "sdescending (a ## w)" + shows "sdescending w" + using assms by (auto elim: sdescending.cases) + lemma sdescending_sappend: + assumes "sdescending (u @- v)" + obtains "sdescending v" + using assms by (induct u) (auto elim: sdescending.cases) 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)) + using assms by (metis sdescending_sappend stake_sdrop) + + lemma sdescending_sset_scons: + assumes "sdescending (a ## w)" + assumes "b \ sset w" + shows "a \ b" + proof - + have "pred_stream (\ b. a \ b) w" if "sdescending w" "a \ shd w" for w + using that by (coinduction arbitrary: w) (auto elim: sdescending.cases) + then show ?thesis using assms unfolding stream.pred_set by force + qed + lemma sdescending_sset_sappend: + assumes "sdescending (u @- v)" + assumes "a \ set u" "b \ sset v" + shows "a \ b" + using assms by (induct u) (auto elim: sdescending.cases dest: sdescending_sset_scons) 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)" + obtains u a + where "w = u @- sconst a" using assms - proof (induct "w !! 0" arbitrary: w thesis rule: less_induct) + proof (induct "shd w" arbitrary: w thesis rule: less_induct) case less show ?case - proof (cases "w = sconst (w !! 0)") + proof (cases "w = sconst (shd w)") case True - show ?thesis using less(2)[of 0] True by simp + show ?thesis using shift_replicate_sconst less(2) True by metis 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 + then obtain u v where 1: "w = u @- v" "u \ []" "shd w \ shd v" + by (metis empty_iff eqI_snth insert_iff sdrop_simps(1) shift.simps(1) snth_sset sset_sconst stake_sdrop) + have 2: "shd w \ shd v" using sdescending_sset_sappend less(3) 1 by (metis hd_in_set shd_sset shift_simps(1)) + have 3: "shd w > shd v" using 1(3) 2 by simp + obtain s a where 4: "v = s @- sconst a" using sdescending_sappend less(1, 3) 1(1) 3 by metis + have 5: "w = (u @ s) @- sconst a" unfolding 1(1) 4 by simp + show ?thesis using less(2) 5 by this qed qed end