diff --git a/thys/Auto2_HOL/HOL/Lists_Thms.thy b/thys/Auto2_HOL/HOL/Lists_Thms.thy --- a/thys/Auto2_HOL/HOL/Lists_Thms.thy +++ b/thys/Auto2_HOL/HOL/Lists_Thms.thy @@ -1,300 +1,302 @@ (* File: Lists_Thms.thy Author: Bohua Zhan Setup for proof steps related to lists. *) section \Setup for lists\ theory Lists_Thms imports Set_Thms begin subsection \Definition of lists\ setup \add_resolve_prfstep @{thm list.distinct(2)}\ setup \add_forward_prfstep (equiv_forward_th @{thm list.simps(1)})\ setup \fold add_rewrite_rule @{thms List.list.sel}\ setup \add_rewrite_rule @{thm list.collapse}\ setup \add_var_induct_rule @{thm list.induct}\ subsection \Length\ setup \add_rewrite_rule @{thm List.list.size(3)}\ lemma length_one [rewrite]: "length [x] = 1" by simp lemma length_Cons [rewrite]: "length (a # b) = length b + 1" by simp lemma length_snoc [rewrite]: "length (xs @ [x]) = length xs + 1" by auto lemma length_zero_is_nil [forward]: "length xs = 0 \ xs = []" by simp lemma length_gt_zero [forward]: "length xs > 0 \ xs \ []" by simp subsection \Append\ setup \add_rewrite_rule @{thm List.length_append}\ setup \add_rewrite_rule_cond @{thm List.append.simps(2)} [with_cond "?xs \ []"]\ setup \add_rewrite_rule @{thm List.hd_append2}\ lemma append_is_empty [forward]: "xs @ ys = [] \ xs = [] \ ys = []" by simp lemma cons_to_append [rewrite_back]: "a # b = [a] @ b" by simp ML_file \list_ac.ML\ ML_file \list_ac_test.ML\ subsection \Showing two lists are equal\ setup \add_backward2_prfstep_cond @{thm nth_equalityI} [with_filt (order_filter "xs" "ys")]\ subsection \Set of elements of a list\ setup \add_rewrite_rule @{thm List.set_simps(1)}\ lemma set_one [rewrite]: "set [u] = {u}" by simp lemma set_two [rewrite]: "set [u, v] = {u, v}" by simp lemma set_simps2: "set (x # xs) = {x} \ set xs" by simp setup \add_rewrite_rule_cond @{thm set_simps2} [with_cond "?xs \ []", with_cond "?xs \ [?y]"]\ setup \add_rewrite_rule @{thm List.set_append}\ setup \add_rewrite_rule @{thm List.set_rev}\ setup \add_resolve_prfstep @{thm List.finite_set}\ setup \add_backward_prfstep (equiv_forward_th @{thm in_set_conv_nth})\ subsection \hd\ setup \register_wellform_data ("hd xs", ["xs \ []"])\ setup \add_forward_prfstep_cond @{thm List.hd_in_set} [with_term "hd ?xs"]\ subsection \tl\ setup \add_rewrite_rule @{thm length_tl}\ lemma nth_tl' [rewrite]: "i < length (tl xs) \ tl xs ! i = xs ! (i + 1)" by (simp add: nth_tl) lemma set_tl_subset [forward_arg1]: "set (tl xs) \ set xs" by (metis list.set_sel(2) subsetI tl_Nil) subsection \nth\ setup \register_wellform_data ("xs ! i", ["i < length xs"])\ setup \add_prfstep_check_req ("xs ! i", "i < length xs")\ setup \add_rewrite_rule_back @{thm hd_conv_nth}\ setup \add_rewrite_rule @{thm List.nth_Cons'}\ setup \add_rewrite_rule @{thm List.nth_append}\ setup \add_forward_prfstep_cond @{thm nth_mem} [with_term "?xs ! ?n"]\ subsection \sorted\ lemma sorted_Nil [resolve]: "sorted []" by simp lemma sorted_single [resolve]: "sorted [x]" by simp -setup \add_backward_prfstep (equiv_backward_th @{thm sorted.simps(2)})\ +setup \add_backward_prfstep (equiv_backward_th @{thm sorted_simps(2)})\ lemma sorted_ConsD1 [forward]: "sorted (x # xs) \ sorted xs" - using sorted.simps(2) by blast + by simp lemma sorted_ConsD2 [forward, backward2]: "sorted (x # xs) \ y \ set xs \ x \ y" - using sorted.simps(2) by blast + by simp lemma sorted_appendI [backward]: "sorted xs \ sorted ys \ (\x\set xs. \y\set ys. x \ y) \ sorted (xs @ ys)" by (simp add: sorted_append) lemma sorted_appendE [forward]: "sorted (xs @ ys) \ sorted xs \ sorted ys" by (simp add: sorted_append) lemma sorted_appendE2 [forward]: "sorted (xs @ ys) \ x \ set xs \ \y\set ys. x \ y" using sorted_append by blast lemma sorted_nth_mono' [backward]: "sorted xs \ j < length xs \ i \ j \ xs ! i \ xs ! j" using sorted_nth_mono by auto +(*NOT CLEAR WHY THIS THEOREM IS CURRENTLY REJECTED WHEN IT WAS ALLOWED BEFORE: + "gen_prfstep: new schematic variable in pats_r / cases" lemma sorted_nth_mono_less [forward]: "sorted xs \ i < length xs \ xs ! i < xs ! j \ i < j" by (meson leD not_le_imp_less sorted_nth_mono) - +*) subsection \sort\ setup \add_forward_prfstep_cond @{thm sorted_sort} [with_term "sort ?xs"]\ setup \add_rewrite_rule @{thm length_sort}\ setup \add_rewrite_rule @{thm mset_sort}\ setup \add_rewrite_rule @{thm set_sort}\ setup \add_backward_prfstep @{thm properties_for_sort}\ lemma sort_Nil [rewrite]: "sort [] = []" by auto lemma sort_singleton [rewrite]: "sort [a] = [a]" by auto subsection \distinct\ lemma distinct_Nil [resolve]: "distinct []" by simp setup \add_resolve_prfstep @{thm List.distinct_singleton}\ setup \add_rewrite_rule_cond @{thm distinct.simps(2)} [with_cond "?xs \ []"]\ setup \add_rewrite_rule @{thm distinct_append}\ setup \add_rewrite_rule @{thm distinct_rev}\ setup \add_rewrite_rule @{thm distinct_sort}\ setup \add_resolve_prfstep (equiv_backward_th @{thm distinct_conv_nth})\ lemma distinct_nthE [forward]: "distinct xs \ i < length xs \ j < length xs \ xs ! i = xs ! j \ i = j" using nth_eq_iff_index_eq by blast subsection \map function\ setup \fold add_rewrite_rule @{thms List.list.map}\ setup \add_rewrite_rule @{thm List.length_map}\ setup \add_rewrite_rule @{thm List.nth_map}\ setup \add_rewrite_rule @{thm List.map_append}\ subsection \Replicate\ setup \add_rewrite_arg_rule @{thm length_replicate}\ setup \add_rewrite_rule @{thm List.nth_replicate}\ subsection \last\ setup \register_wellform_data ("last xs", ["xs \ []"])\ lemma last_eval1 [rewrite]: "last [x] = x" by simp lemma last_eval2 [rewrite]: "last [u, v] = v" by simp setup \add_rewrite_rule @{thm List.last_ConsR}\ setup \add_rewrite_rule @{thm List.last_appendR}\ setup \add_rewrite_rule @{thm List.last_snoc}\ setup \add_rewrite_rule_back @{thm last_conv_nth}\ setup \add_forward_prfstep_cond @{thm List.last_in_set} [with_term "last ?as"]\ subsection \butlast\ setup \add_rewrite_arg_rule @{thm List.length_butlast}\ setup \add_rewrite_rule @{thm List.nth_butlast}\ setup \add_rewrite_rule_back @{thm List.butlast_conv_take}\ setup \add_rewrite_rule @{thm List.butlast_snoc}\ lemma butlast_eval1 [rewrite]: "butlast [x] = []" by simp lemma butlast_eval2 [rewrite]: "butlast [x, y] = [x]" by simp lemma butlast_cons [rewrite]: "as \ [] \ butlast (a # as) = a # butlast as" by simp lemma butlast_append' [rewrite]: "bs \ [] \ butlast (as @ bs) = as @ butlast bs" by (simp add: butlast_append) setup \add_rewrite_rule @{thm List.append_butlast_last_id}\ lemma set_butlast_is_subset: "set (butlast xs) \ set xs" by (simp add: in_set_butlastD subsetI) setup \add_forward_arg1_prfstep @{thm set_butlast_is_subset}\ subsection \List update\ setup \register_wellform_data ("xs[i := x]", ["i < length xs"])\ setup \add_rewrite_arg_rule @{thm List.length_list_update}\ setup \add_rewrite_rule @{thm List.nth_list_update_eq}\ setup \add_rewrite_rule @{thm List.nth_list_update_neq}\ setup \add_rewrite_rule @{thm List.nth_list_update}\ subsection \take\ setup \register_wellform_data ("take n xs", ["n \ length xs"])\ setup \add_prfstep_check_req ("take n xs", "n \ length xs")\ lemma length_take [rewrite_arg]: "n \ length xs \ length (take n xs) = n" by simp lemma nth_take [rewrite]: "i < length (take n xs) \ take n xs ! i = xs ! i" by simp setup \add_rewrite_rule @{thm List.take_0}\ setup \add_rewrite_rule @{thm List.take_Suc_conv_app_nth}\ lemma take_length [rewrite]: "take (length xs) xs = xs" by simp setup \add_forward_arg1_prfstep @{thm List.set_take_subset}\ lemma take_Suc [rewrite]: "Suc n \ length xs \ take (Suc n) xs = take n xs @ [nth xs n]" using Suc_le_lessD take_Suc_conv_app_nth by blast setup \add_rewrite_rule @{thm List.take_update_cancel}\ setup \add_rewrite_rule @{thm List.append_take_drop_id}\ setup \add_rewrite_rule @{thm List.take_all}\ subsection \drop\ setup \add_rewrite_arg_rule @{thm List.length_drop}\ lemma nth_drop [rewrite]: "i < length (drop n xs) \ drop n xs ! i = xs ! (n + i)" by simp setup \add_rewrite_rule @{thm List.drop_0}\ setup \add_rewrite_rule @{thm List.drop_all}\ setup \add_rewrite_rule_back @{thm List.take_drop}\ setup \add_rewrite_rule @{thm List.drop_drop}\ subsection \rev\ setup \add_rewrite_arg_rule @{thm List.length_rev}\ setup \fold add_rewrite_rule @{thms List.rev.simps}\ setup \add_rewrite_rule @{thm List.rev_append}\ setup \add_rewrite_rule @{thm List.rev_rev_ident}\ subsection \filter\ setup \fold add_rewrite_rule @{thms filter.simps}\ setup \add_rewrite_rule @{thm filter_append}\ setup \add_rewrite_rule_bidir @{thm rev_filter}\ subsection \concat\ setup \fold add_rewrite_rule @{thms concat.simps}\ subsection \mset\ setup \add_rewrite_rule @{thm mset.simps(1)}\ lemma mset_simps_2 [rewrite]: "mset (a # x) = mset x + {#a#}" by simp setup \add_rewrite_rule @{thm mset_append}\ setup \add_rewrite_rule @{thm mset_eq_setD}\ setup \add_rewrite_rule_cond @{thm in_multiset_in_set} [with_term "set ?xs"]\ setup \add_rewrite_rule_back_cond @{thm in_multiset_in_set} [with_term "mset ?xs"]\ setup \add_backward_prfstep @{thm Multiset.nth_mem_mset}\ lemma in_mset_conv_nth [resolve]: "x \# mset xs \ \i [] \ hd xs \# mset xs" by simp setup \add_forward_prfstep_cond @{thm hd_in_mset} [with_term "hd ?xs", with_term "mset ?xs"]\ lemma last_in_mset: "xs \ [] \ last xs \# mset xs" by simp setup \add_forward_prfstep_cond @{thm last_in_mset} [with_term "last ?xs", with_term "mset ?xs"]\ subsection \Relationship between mset and set of lists\ lemma mset_butlast [rewrite]: "xs \ [] \ mset (butlast xs) = mset xs - {# last xs #}" by (metis add_diff_cancel_right' append_butlast_last_id mset.simps(1) mset.simps(2) union_code) lemma insert_mset_to_set [rewrite]: "mset xs' = mset xs + {# x #} \ set xs' = set xs \ {x}" by (metis set_mset_mset set_mset_single set_mset_union) lemma delete_mset_to_set [rewrite]: "distinct xs \ mset xs' = mset xs - {# x #} \ set xs' = set xs - {x}" by (metis mset_eq_setD mset_remove1 set_remove1_eq) lemma update_mset_to_set [rewrite]: "distinct xs \ mset xs' = {# y #} + (mset xs - {# x #}) \ set xs' = (set xs - {x}) \ {y}" by (metis insert_mset_to_set mset_remove1 set_remove1_eq union_commute) lemma mset_update' [rewrite]: "i < length ls \ mset (ls[i := v]) = {#v#} + (mset ls - {# ls ! i #})" using mset_update by fastforce subsection \swap\ setup \add_rewrite_rule @{thm mset_swap}\ setup \add_rewrite_rule @{thm set_swap}\ subsection \upto lists\ lemma upt_zero_length [rewrite]: "length [0.. [0..Lambda lists\ definition list :: "(nat \ 'a) \ nat \ 'a list" where "list s n = map s [0 ..< n]" lemma list_length [rewrite_arg]: "length (list s n) = n" by (simp add: list_def) lemma list_nth [rewrite]: "i < length (list s n) \ (list s n) ! i = s i" by (simp add: list_def) subsection \Splitting of lists\ setup \add_resolve_prfstep @{thm split_list}\ setup \add_backward_prfstep @{thm not_distinct_decomp}\ subsection \Finiteness\ setup \add_resolve_prfstep @{thm finite_lists_length_le}\ subsection \Cardinality\ setup \add_rewrite_rule @{thm distinct_card}\ end diff --git a/thys/Automatic_Refinement/Lib/Misc.thy b/thys/Automatic_Refinement/Lib/Misc.thy --- a/thys/Automatic_Refinement/Lib/Misc.thy +++ b/thys/Automatic_Refinement/Lib/Misc.thy @@ -1,4525 +1,4523 @@ (* Title: Miscellaneous Definitions and Lemmas Author: Peter Lammich Maintainer: Peter Lammich Thomas Tuerk *) (* CHANGELOG: 2010-05-09: Removed AC, AI locales, they are superseeded by concepts from OrderedGroups 2010-09-22: Merges with ext/Aux 2017-06-12: Added a bunch of lemmas from various other projects *) section \Miscellaneous Definitions and Lemmas\ theory Misc imports Main "HOL-Library.Multiset" "HOL-ex.Quicksort" "HOL-Library.Option_ord" "HOL-Library.Infinite_Set" "HOL-Eisbach.Eisbach" begin text_raw \\label{thy:Misc}\ subsection \Isabelle Distribution Move Proposals\ subsubsection \Pure\ lemma TERMI: "TERM x" unfolding Pure.term_def . subsubsection \HOL\ (* Stronger disjunction elimination rules. *) lemma disjE1: "\ P \ Q; P \ R; \\P;Q\ \ R \ \ R" by metis lemma disjE2: "\ P \ Q; \P; \Q\ \ R; Q \ R \ \ R" by blast lemma imp_mp_iff[simp]: "a \ (a \ b) \ a \ b" "(a \ b) \ a \ a \ b" (* is Inductive.imp_conj_iff, but not in simpset by default *) by blast+ lemma atomize_Trueprop_eq[atomize]: "(Trueprop x \ Trueprop y) \ Trueprop (x=y)" apply rule apply (rule) apply (erule equal_elim_rule1) apply assumption apply (erule equal_elim_rule2) apply assumption apply simp done subsubsection \Set\ lemma inter_compl_diff_conv[simp]: "A \ -B = A - B" by auto lemma pair_set_inverse[simp]: "{(a,b). P a b}\ = {(b,a). P a b}" by auto lemma card_doubleton_eq_2_iff[simp]: "card {a,b} = 2 \ a\b" by auto subsubsection \List\ (* TODO: Move, analogous to List.length_greater_0_conv *) thm List.length_greater_0_conv lemma length_ge_1_conv[iff]: "Suc 0 \ length l \ l\[]" by (cases l) auto \ \Obtains a list from the pointwise characterization of its elements\ lemma obtain_list_from_elements: assumes A: "\ili. P li i)" obtains l where "length l = n" "\il. length l=n \ (\iii l!j" by (simp add: sorted_iff_nth_mono) also from nth_eq_iff_index_eq[OF D] B have "l!i \ l!j" by auto finally show ?thesis . qed lemma distinct_sorted_strict_mono_iff: assumes "distinct l" "sorted l" assumes "i il!j \ i\j" by (metis assms distinct_sorted_strict_mono_iff leD le_less_linear) (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] We could, analogously, declare rules for "map _ _ = _@_" as dest!, or use elim!, or declare the _conv-rule as simp *) lemma map_eq_appendE: assumes "map f ls = fl@fl'" obtains l l' where "ls=l@l'" and "map f l=fl" and "map f l' = fl'" using assms proof (induction fl arbitrary: ls thesis) case (Cons x xs) then obtain l ls' where [simp]: "ls = l#ls'" "f l = x" by force with Cons.prems(2) have "map f ls' = xs @ fl'" by simp from Cons.IH[OF _ this] guess ll ll' . with Cons.prems(1)[of "l#ll" ll'] show thesis by simp qed simp lemma map_eq_append_conv: "map f ls = fl@fl' \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: map_eq_appendE) lemmas append_eq_mapE = map_eq_appendE[OF sym] lemma append_eq_map_conv: "fl@fl' = map f ls \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: append_eq_mapE) lemma distinct_mapI: "distinct (map f l) \ distinct l" by (induct l) auto lemma map_distinct_upd_conv: "\i \ (map f l)[i := x] = map (f(l!i := x)) l" \ \Updating a mapped distinct list is equal to updating the mapping function\ by (auto simp: nth_eq_iff_index_eq intro: nth_equalityI) lemma zip_inj: "\length a = length b; length a' = length b'; zip a b = zip a' b'\ \ a=a' \ b=b'" proof (induct a b arbitrary: a' b' rule: list_induct2) case Nil then show ?case by (cases a'; cases b'; auto) next case (Cons x xs y ys) then show ?case by (cases a'; cases b'; auto) qed lemma zip_eq_zip_same_len[simp]: "\ length a = length b; length a' = length b' \ \ zip a b = zip a' b' \ a=a' \ b=b'" by (auto dest: zip_inj) lemma upt_merge[simp]: "i\j \ j\k \ [i.. (ys \ [] \ butlast ys = xs \ last ys = x)" by auto (* Case distinction how two elements of a list can be related to each other *) lemma list_match_lel_lel: assumes "c1 @ qs # c2 = c1' @ qs' # c2'" obtains (left) c21' where "c1 = c1' @ qs' # c21'" "c2' = c21' @ qs # c2" | (center) "c1' = c1" "qs' = qs" "c2' = c2" | (right) c21 where "c1' = c1 @ qs # c21" "c2 = c21 @ qs' # c2'" using assms apply (clarsimp simp: append_eq_append_conv2) subgoal for us by (cases us) auto done lemma xy_in_set_cases[consumes 2, case_names EQ XY YX]: assumes A: "x\set l" "y\set l" and C: "!!l1 l2. \ x=y; l=l1@y#l2 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@x#l2@y#l3 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@y#l2@x#l3 \ \ P" shows P proof (cases "x=y") case True with A(1) obtain l1 l2 where "l=l1@y#l2" by (blast dest: split_list) with C(1) True show ?thesis by blast next case False from A(1) obtain l1 l2 where S1: "l=l1@x#l2" by (blast dest: split_list) from A(2) obtain l1' l2' where S2: "l=l1'@y#l2'" by (blast dest: split_list) from S1 S2 have M: "l1@x#l2 = l1'@y#l2'" by simp thus P proof (cases rule: list_match_lel_lel[consumes 1, case_names 1 2 3]) case (1 c) with S1 have "l=l1'@y#c@x#l2" by simp with C(3) False show ?thesis by blast next case 2 with False have False by blast thus ?thesis .. next case (3 c) with S1 have "l=l1@x#c@y#l2'" by simp with C(2) False show ?thesis by blast qed qed lemma list_e_eq_lel[simp]: "[e] = l1@e'#l2 \ l1=[] \ e'=e \ l2=[]" "l1@e'#l2 = [e] \ l1=[] \ e'=e \ l2=[]" apply (cases l1, auto) [] apply (cases l1, auto) [] done lemma list_ee_eq_leel[simp]: "([e1,e2] = l1@e1'#e2'#l2) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" "(l1@e1'#e2'#l2 = [e1,e2]) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" apply (cases l1, auto) [] apply (cases l1, auto) [] done subsubsection \Transitive Closure\ text \A point-free induction rule for elements reachable from an initial set\ lemma rtrancl_reachable_induct[consumes 0, case_names base step]: assumes I0: "I \ INV" assumes IS: "E``INV \ INV" shows "E\<^sup>*``I \ INV" by (metis I0 IS Image_closed_trancl Image_mono subset_refl) lemma acyclic_empty[simp, intro!]: "acyclic {}" by (unfold acyclic_def) auto lemma acyclic_union: "acyclic (A\B) \ acyclic A" "acyclic (A\B) \ acyclic B" by (metis Un_upper1 Un_upper2 acyclic_subset)+ subsubsection \Lattice Syntax\ (* Providing the syntax in a locale makes it more usable, without polluting the global namespace*) locale Lattice_Syntax begin notation bot ("\") and top ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\_" [900] 900) and Sup ("\_" [900] 900) end text \Here we provide a collection of miscellaneous definitions and helper lemmas\ subsection "Miscellaneous (1)" text \This stuff is used in this theory itself, and thus occurs in first place or is simply not sorted into any other section of this theory.\ lemma IdD: "(a,b)\Id \ a=b" by simp text \Conversion Tag\ definition [simp]: "CNV x y \ x=y" lemma CNV_I: "CNV x x" by simp lemma CNV_eqD: "CNV x y \ x=y" by simp lemma CNV_meqD: "CNV x y \ x\y" by simp (* TODO: Move. Shouldn't this be detected by simproc? *) lemma ex_b_b_and_simp[simp]: "(\b. b \ Q b) \ Q True" by auto lemma ex_b_not_b_and_simp[simp]: "(\b. \b \ Q b) \ Q False" by auto method repeat_all_new methods m = m;(repeat_all_new \m\)? subsubsection "AC-operators" text \Locale to declare AC-laws as simplification rules\ locale Assoc = fixes f assumes assoc[simp]: "f (f x y) z = f x (f y z)" locale AC = Assoc + assumes commute[simp]: "f x y = f y x" lemma (in AC) left_commute[simp]: "f x (f y z) = f y (f x z)" by (simp only: assoc[symmetric]) simp lemmas (in AC) AC_simps = commute assoc left_commute text \Locale to define functions from surjective, unique relations\ locale su_rel_fun = fixes F and f assumes unique: "\(A,B)\F; (A,B')\F\ \ B=B'" assumes surjective: "\!!B. (A,B)\F \ P\ \ P" assumes f_def: "f A == THE B. (A,B)\F" lemma (in su_rel_fun) repr1: "(A,f A)\F" proof (unfold f_def) obtain B where "(A,B)\F" by (rule surjective) with theI[where P="\B. (A,B)\F", OF this] show "(A, THE x. (A, x) \ F) \ F" by (blast intro: unique) qed lemma (in su_rel_fun) repr2: "(A,B)\F \ B=f A" using repr1 by (blast intro: unique) lemma (in su_rel_fun) repr: "(f A = B) = ((A,B)\F)" using repr1 repr2 by (blast) \ \Contract quantification over two variables to pair\ lemma Ex_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma All_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma nat_geq_1_eq_neqz: "x\1 \ x\(0::nat)" by auto lemma nat_in_between_eq: "(a b\Suc a) \ b = Suc a" "(a\b \ b b = a" by auto lemma Suc_n_minus_m_eq: "\ n\m; m>1 \ \ Suc (n - m) = n - (m - 1)" by simp lemma Suc_to_right: "Suc n = m \ n = m - Suc 0" by simp lemma Suc_diff[simp]: "\n m. n\m \ m\1 \ Suc (n - m) = n - (m - 1)" by simp lemma if_not_swap[simp]: "(if \c then a else b) = (if c then b else a)" by auto lemma all_to_meta: "Trueprop (\a. P a) \ (\a. P a)" apply rule by auto lemma imp_to_meta: "Trueprop (P\Q) \ (P\Q)" apply rule by auto (* for some reason, there is no such rule in HOL *) lemma iffI2: "\P \ Q; \ P \ \ Q\ \ P \ Q" by metis lemma iffExI: "\ \x. P x \ Q x; \x. Q x \ P x \ \ (\x. P x) \ (\x. Q x)" by metis lemma bex2I[intro?]: "\ (a,b)\S; (a,b)\S \ P a b \ \ \a b. (a,b)\S \ P a b" by blast (* TODO: Move lemma to HOL! *) lemma cnv_conj_to_meta: "(P \ Q \ PROP X) \ (\P;Q\ \ PROP X)" by (rule BNF_Fixpoint_Base.conj_imp_eq_imp_imp) subsection \Sets\ lemma remove_subset: "x\S \ S-{x} \ S" by auto lemma subset_minus_empty: "A\B \ A-B = {}" by auto lemma insert_minus_eq: "x\y \ insert x A - {y} = insert x (A - {y})" by auto lemma set_notEmptyE: "\S\{}; !!x. x\S \ P\ \ P" by (metis equals0I) lemma subset_Collect_conv: "S \ Collect P \ (\x\S. P x)" by auto lemma memb_imp_not_empty: "x\S \ S\{}" by auto lemma disjoint_mono: "\ a\a'; b\b'; a'\b'={} \ \ a\b={}" by auto lemma disjoint_alt_simp1: "A-B = A \ A\B = {}" by auto lemma disjoint_alt_simp2: "A-B \ A \ A\B \ {}" by auto lemma disjoint_alt_simp3: "A-B \ A \ A\B \ {}" by auto lemma disjointI[intro?]: "\ \x. \x\a; x\b\ \ False \ \ a\b={}" by auto lemmas set_simps = subset_minus_empty disjoint_alt_simp1 disjoint_alt_simp2 disjoint_alt_simp3 Un_absorb1 Un_absorb2 lemma set_minus_singleton_eq: "x\X \ X-{x} = X" by auto lemma set_diff_diff_left: "A-B-C = A-(B\C)" by auto lemma image_update[simp]: "x\A \ f(x:=n)`A = f`A" by auto lemma eq_or_mem_image_simp[simp]: "{f l |l. l = a \ l\B} = insert (f a) {f l|l. l\B}" by blast lemma set_union_code [code_unfold]: "set xs \ set ys = set (xs @ ys)" by auto lemma in_fst_imageE: assumes "x \ fst`S" obtains y where "(x,y)\S" using assms by auto lemma in_snd_imageE: assumes "y \ snd`S" obtains x where "(x,y)\S" using assms by auto lemma fst_image_mp: "\fst`A \ B; (x,y)\A \ \ x\B" by (metis Domain.DomainI fst_eq_Domain in_mono) lemma snd_image_mp: "\snd`A \ B; (x,y)\A \ \ y\B" by (metis Range.intros rev_subsetD snd_eq_Range) lemma inter_eq_subsetI: "\ S\S'; A\S' = B\S' \ \ A\S = B\S" by auto text \ Decompose general union over sum types. \ lemma Union_plus: "(\ x \ A <+> B. f x) = (\ a \ A. f (Inl a)) \ (\b \ B. f (Inr b))" by auto lemma Union_sum: "(\x. f (x::'a+'b)) = (\l. f (Inl l)) \ (\r. f (Inr r))" (is "?lhs = ?rhs") proof - have "?lhs = (\x \ UNIV <+> UNIV. f x)" by simp thus ?thesis by (simp only: Union_plus) qed subsubsection \Finite Sets\ lemma card_1_singletonI: "\finite S; card S = 1; x\S\ \ S={x}" proof (safe, rule ccontr, goal_cases) case prems: (1 x') hence "finite (S-{x})" "S-{x} \ {}" by auto hence "card (S-{x}) \ 0" by auto moreover from prems(1-3) have "card (S-{x}) = 0" by auto ultimately have False by simp thus ?case .. qed lemma card_insert_disjoint': "\finite A; x \ A\ \ card (insert x A) - Suc 0 = card A" by (drule (1) card_insert_disjoint) auto lemma card_eq_UNIV[simp]: "card (S::'a::finite set) = card (UNIV::'a set) \ S=UNIV" proof (auto) fix x assume A: "card S = card (UNIV::'a set)" show "x\S" proof (rule ccontr) assume "x\S" hence "S\UNIV" by auto with psubset_card_mono[of UNIV S] have "card S < card (UNIV::'a set)" by auto with A show False by simp qed qed lemma card_eq_UNIV2[simp]: "card (UNIV::'a set) = card (S::'a::finite set) \ S=UNIV" using card_eq_UNIV[of S] by metis lemma card_ge_UNIV[simp]: "card (UNIV::'a::finite set) \ card (S::'a set) \ S=UNIV" using card_mono[of "UNIV::'a::finite set" S, simplified] by auto lemmas length_remdups_card = length_remdups_concat[of "[l]", simplified] for l lemma fs_contract: "fst ` { p | p. f (fst p) (snd p) \ S } = { a . \b. f a b \ S }" by (simp add: image_Collect) lemma finite_Collect: "finite S \ inj f \ finite {a. f a : S}" by(simp add: finite_vimageI vimage_def[symmetric]) \ \Finite sets have an injective mapping to an initial segments of the natural numbers\ (* This lemma is also in the standard library (from Isabelle2009-1 on) as @{thm [source] Finite_Set.finite_imp_inj_to_nat_seg}. However, it is formulated with HOL's \ there rather then with the meta-logic obtain *) lemma finite_imp_inj_to_nat_seg': fixes A :: "'a set" assumes A: "finite A" obtains f::"'a \ nat" and n::"nat" where "f`A = {i. i finite (lists P \ { l. length l = n })" proof (induct n) case 0 thus ?case by auto next case (Suc n) have "lists P \ { l. length l = Suc n } = (\(a,l). a#l) ` (P \ (lists P \ {l. length l = n}))" apply auto apply (case_tac x) apply auto done moreover from Suc have "finite \" by auto ultimately show ?case by simp qed lemma lists_of_len_fin2: "finite P \ finite (lists P \ { l. n = length l })" proof - assume A: "finite P" have S: "{ l. n = length l } = { l. length l = n }" by auto have "finite (lists P \ { l. n = length l }) \ finite (lists P \ { l. length l = n })" by (subst S) simp thus ?thesis using lists_of_len_fin1[OF A] by auto qed lemmas lists_of_len_fin = lists_of_len_fin1 lists_of_len_fin2 (* Try (simp only: cset_fin_simps, fastforce intro: cset_fin_intros) when reasoning about finiteness of collected sets *) lemmas cset_fin_simps = Ex_prod_contract fs_contract[symmetric] image_Collect[symmetric] lemmas cset_fin_intros = finite_imageI finite_Collect inj_onI lemma Un_interval: fixes b1 :: "'a::linorder" assumes "b1\b2" and "b2\b3" shows "{ f i | i. b1\i \ i { f i | i. b2\i \ ii \ i The standard library proves that a generalized union is finite if the index set is finite and if for every index the component set is itself finite. Conversely, we show that every component set must be finite when the union is finite. \ lemma finite_UNION_then_finite: "finite (\(B ` A)) \ a \ A \ finite (B a)" by (metis Set.set_insert UN_insert Un_infinite) lemma finite_if_eq_beyond_finite: "finite S \ finite {s. s - S = s' - S}" proof (rule finite_subset[where B="(\s. s \ (s' - S)) ` Pow S"], clarsimp) fix s have "s = (s \ S) \ (s - S)" by auto also assume "s - S = s' - S" finally show "s \ (\s. s \ (s' - S)) ` Pow S" by blast qed blast lemma distinct_finite_subset: assumes "finite x" shows "finite {ys. set ys \ x \ distinct ys}" (is "finite ?S") proof (rule finite_subset) from assms show "?S \ {ys. set ys \ x \ length ys \ card x}" by clarsimp (metis distinct_card card_mono) from assms show "finite ..." by (rule finite_lists_length_le) qed lemma distinct_finite_set: shows "finite {ys. set ys = x \ distinct ys}" (is "finite ?S") proof (cases "finite x") case False hence "{ys. set ys = x} = {}" by auto thus ?thesis by simp next case True show ?thesis proof (rule finite_subset) show "?S \ {ys. set ys \ x \ length ys \ card x}" using distinct_card by force from True show "finite ..." by (rule finite_lists_length_le) qed qed lemma finite_set_image: assumes f: "finite (set ` A)" and dist: "\xs. xs \ A \ distinct xs" shows "finite A" proof (rule finite_subset) from f show "finite (set -` (set ` A) \ {xs. distinct xs})" proof (induct rule: finite_induct) case (insert x F) have "finite (set -` {x} \ {xs. distinct xs})" apply (simp add: vimage_def) by (metis Collect_conj_eq distinct_finite_set) with insert show ?case apply (subst vimage_insert) apply (subst Int_Un_distrib2) apply (rule finite_UnI) apply simp_all done qed simp moreover from dist show "A \ ..." by (auto simp add: vimage_image_eq) qed subsubsection \Infinite Set\ lemma INFM_nat_inductI: assumes P0: "P (0::nat)" assumes PS: "\i. P i \ \j>i. P j \ Q j" shows "\\<^sub>\i. Q i" proof - have "\i. \j>i. P j \ Q j" proof fix i show "\j>i. P j \ Q j" apply (induction i) using PS[OF P0] apply auto [] by (metis PS Suc_lessI) qed thus ?thesis unfolding INFM_nat by blast qed subsection \Functions\ lemma fun_neq_ext_iff: "m\m' \ (\x. m x \ m' x)" by auto definition "inv_on f A x == SOME y. y\A \ f y = x" lemma inv_on_f_f[simp]: "\inj_on f A; x\A\ \ inv_on f A (f x) = x" by (auto simp add: inv_on_def inj_on_def) lemma f_inv_on_f: "\ y\f`A \ \ f (inv_on f A y) = y" by (auto simp add: inv_on_def intro: someI2) lemma inv_on_f_range: "\ y \ f`A \ \ inv_on f A y \ A" by (auto simp add: inv_on_def intro: someI2) lemma inj_on_map_inv_f [simp]: "\set l \ A; inj_on f A\ \ map (inv_on f A) (map f l) = l" apply (simp) apply (induct l) apply auto done lemma comp_cong_right: "x = y \ f o x = f o y" by (simp) lemma comp_cong_left: "x = y \ x o f = y o f" by (simp) lemma fun_comp_eq_conv: "f o g = fg \ (\x. f (g x) = fg x)" by auto abbreviation comp2 (infixl "oo" 55) where "f oo g \ \x. f o (g x)" abbreviation comp3 (infixl "ooo" 55) where "f ooo g \ \x. f oo (g x)" notation comp2 (infixl "\\" 55) and comp3 (infixl "\\\" 55) definition [code_unfold, simp]: "swap_args2 f x y \ f y x" subsection \Multisets\ (* The following is a syntax extension for multisets. Unfortunately, it depends on a change in the Library/Multiset.thy, so it is commented out here, until it will be incorporated into Library/Multiset.thy by its maintainers. The required change in Library/Multiset.thy is removing the syntax for single: - single :: "'a => 'a multiset" ("{#_#}") + single :: "'a => 'a multiset" And adding the following translations instead: + syntax + "_multiset" :: "args \ 'a multiset" ("{#(_)#}") + translations + "{#x, xs#}" == "{#x#} + {#xs#}" + "{# x #}" == "single x" This translates "{# \ #}" into a sum of singletons, that is parenthesized to the right. ?? Can we also achieve left-parenthesizing ?? *) (* Let's try what happens if declaring AC-rules for multiset union as simp-rules *) (*declare union_ac[simp] -- don't do it !*) lemma count_mset_set_finite_iff: "finite S \ count (mset_set S) a = (if a \ S then 1 else 0)" by simp lemma ex_Melem_conv: "(\x. x \# A) = (A \ {#})" by (simp add: ex_in_conv) subsubsection \Count\ lemma count_ne_remove: "\ x ~= t\ \ count S x = count (S-{#t#}) x" by (auto) lemma mset_empty_count[simp]: "(\p. count M p = 0) = (M={#})" by (auto simp add: multiset_eq_iff) subsubsection \Union, difference and intersection\ lemma size_diff_se: "t \# S \ size S = size (S - {#t#}) + 1" proof (unfold size_multiset_overloaded_eq) let ?SIZE = "sum (count S) (set_mset S)" assume A: "t \# S" from A have SPLITPRE: "finite (set_mset S) & {t}\(set_mset S)" by auto hence "?SIZE = sum (count S) (set_mset S - {t}) + sum (count S) {t}" by (blast dest: sum.subset_diff) hence "?SIZE = sum (count S) (set_mset S - {t}) + count (S) t" by auto moreover with A have "count S t = count (S-{#t#}) t + 1" by auto ultimately have D: "?SIZE = sum (count S) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (arith) moreover have "sum (count S) (set_mset S - {t}) = sum (count (S-{#t#})) (set_mset S - {t})" proof - have "\x\(set_mset S - {t}). count S x = count (S-{#t#}) x" by (auto iff add: count_ne_remove) thus ?thesis by simp qed ultimately have D: "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (simp) moreover { assume CASE: "t \# S - {#t#}" from CASE have "set_mset S - {t} = set_mset (S - {#t#})" by (auto simp add: in_diff_count split: if_splits) with CASE D have "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by (simp add: not_in_iff) } moreover { assume CASE: "t \# S - {#t#}" from CASE have "t \# S" by (rule in_diffD) with CASE have 1: "set_mset S = set_mset (S-{#t#})" by (auto simp add: in_diff_count split: if_splits) moreover from D have "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t} + 1" by simp moreover from SPLITPRE sum.subset_diff have "sum (count (S-{#t#})) (set_mset S) = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t}" by (blast) ultimately have "?SIZE = sum (count (S-{#t#})) (set_mset (S-{#t#})) + 1" by simp } ultimately show "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by blast qed (* TODO: Check whether this proof can be done simpler *) lemma mset_union_diff_comm: "t \# S \ T + (S - {#t#}) = (T + S) - {#t#}" proof - assume "t \# S" then obtain S' where S: "S = add_mset t S'" by (metis insert_DiffM) then show ?thesis by auto qed (* lemma mset_diff_diff_left: "A-B-C = A-((B::'a multiset)+C)" proof - have "\e . count (A-B-C) e = count (A-(B+C)) e" by auto thus ?thesis by (simp add: multiset_eq_conv_count_eq) qed lemma mset_diff_commute: "A-B-C = A-C-(B::'a multiset)" proof - have "A-B-C = A-(B+C)" by (simp add: mset_diff_diff_left) also have "\ = A-(C+B)" by (simp add: union_commute) thus ?thesis by (simp add: mset_diff_diff_left) qed lemma mset_diff_same_empty[simp]: "(S::'a multiset) - S = {#}" proof - have "\e . count (S-S) e = 0" by auto hence "\e . ~ (e : set_mset (S-S))" by auto hence "set_mset (S-S) = {}" by blast thus ?thesis by (auto) qed *) lemma mset_right_cancel_union: "\a \# A+B; ~(a \# B)\ \ a\#A" by (simp) lemma mset_left_cancel_union: "\a \# A+B; ~(a \# A)\ \ a\#B" by (simp) lemmas mset_cancel_union = mset_right_cancel_union mset_left_cancel_union lemma mset_right_cancel_elem: "\a \# A+{#b#}; a~=b\ \ a\#A" by simp lemma mset_left_cancel_elem: "\a \# {#b#}+A; a~=b\ \ a\#A" by simp lemmas mset_cancel_elem = mset_right_cancel_elem mset_left_cancel_elem lemma mset_diff_cancel1elem[simp]: "~(a \# B) \ {#a#}-B = {#a#}" by (auto simp add: not_in_iff intro!: multiset_eqI) (* lemma diff_union_inverse[simp]: "A + B - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) lemma diff_union_inverse2[simp]: "B + A - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) *) (*lemma union_diff_assoc_se2: "t \# A \ (A+B)-{#t#} = (A-{#t#}) + B" by (auto iff add: multiset_eq_conv_count_eq) lemmas union_diff_assoc_se = union_diff_assoc_se1 union_diff_assoc_se2*) lemma union_diff_assoc: "C-B={#} \ (A+B)-C = A + (B-C)" by (simp add: multiset_eq_iff) lemmas mset_neutral_cancel1 = union_left_cancel[where N="{#}", simplified] union_right_cancel[where N="{#}", simplified] declare mset_neutral_cancel1[simp] lemma mset_union_2_elem: "{#a, b#} = add_mset c M \ {#a#}=M & b=c | a=c & {#b#}=M" by (auto simp: add_eq_conv_diff) lemma mset_un_cases[cases set, case_names left right]: "\a \# A + B; a \# A \ P; a \# B \ P\ \ P" by (auto) lemma mset_unplusm_dist_cases[cases set, case_names left right]: assumes A: "{#s#}+A = B+C" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P proof - from A[symmetric] have "s \# B+C" by simp thus ?thesis proof (cases rule: mset_un_cases) case left hence 1: "B={#s#}+(B-{#s#})" by simp with A have "{#s#}+A = {#s#}+((B-{#s#})+C)" by (simp add: union_ac) hence 2: "A = (B-{#s#})+C" by (simp) from L[OF 1 2] show ?thesis . next case right hence 1: "C={#s#}+(C-{#s#})" by simp with A have "{#s#}+A = {#s#}+(B+(C-{#s#}))" by (simp add: union_ac) hence 2: "A = B+(C-{#s#})" by (simp) from R[OF 1 2] show ?thesis . qed qed lemma mset_unplusm_dist_cases2[cases set, case_names left right]: assumes A: "B+C = {#s#}+A" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P using mset_unplusm_dist_cases[OF A[symmetric]] L R by blast lemma mset_single_cases[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'={#s#}+(c'-{#s#}); c={#r'#}+(c-{#r'#}); c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - { assume CASE: "s=r'" with A have "c=c'" by simp with CASE CASES have ?thesis by auto } moreover { assume CASE: "s\r'" have "s \# {#s#}+c" by simp with A have "s \# {#r'#}+c'" by simp with CASE have "s \# c'" by simp hence 1: "c' = {#s#} + (c' - {#s#})" by simp with A have "{#s#}+c = {#s#}+({#r'#}+(c' - {#s#}))" by (auto simp add: union_ac) hence 2: "c={#r'#}+(c' - {#s#})" by (auto) hence 3: "c-{#r'#} = (c' - {#s#})" by auto from 1 2 3 CASES have ?thesis by auto } ultimately show ?thesis by blast qed lemma mset_single_cases'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'={#s#}+cc; c={#r'#}+cc; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases) lemma mset_single_cases2[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'=(c'-{#s#})+{#s#}; c=(c-{#r'#})+{#r'#}; c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - from A have "add_mset s c = add_mset r' c'" by (simp add: union_ac) thus ?thesis using CASES by (cases rule: mset_single_cases) simp_all qed lemma mset_single_cases2'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'=cc+{#s#}; c=cc+{#r'#}; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases2) lemma mset_un_single_un_cases [consumes 1, case_names left right]: assumes A: "add_mset a A = B + C" and CASES: "a \# B \ A = (B - {#a#}) + C \ P" "a \# C \ A = B + (C - {#a#}) \ P" shows "P" proof - have "a \# A+{#a#}" by simp with A have "a \# B+C" by auto thus ?thesis proof (cases rule: mset_un_cases) case left hence "B=B-{#a#}+{#a#}" by auto with A have "A+{#a#} = (B-{#a#})+C+{#a#}" by (auto simp add: union_ac) hence "A=(B-{#a#})+C" by simp with CASES(1)[OF left] show ?thesis by blast next case right hence "C=C-{#a#}+{#a#}" by auto with A have "A+{#a#} = B+(C-{#a#})+{#a#}" by (auto simp add: union_ac) hence "A=B+(C-{#a#})" by simp with CASES(2)[OF right] show ?thesis by blast qed qed (* TODO: Can this proof be done more automatically ? *) lemma mset_distrib[consumes 1, case_names dist]: assumes A: "(A::'a multiset)+B = M+N" "!!Am An Bm Bn. \A=Am+An; B=Bm+Bn; M=Am+Bm; N=An+Bn\ \ P" shows "P" proof - have BN_MA: "B - N = M - A" by (metis (no_types) add_diff_cancel_right assms(1) union_commute) have H: "A = A\# C + (A - C) \# D" if "A + B = C + D" for A B C D :: "'a multiset" by (metis add.commute diff_intersect_left_idem mset_subset_eq_add_left subset_eq_diff_conv subset_mset.add_diff_inverse subset_mset.inf_absorb1 subset_mset.inf_le1 that) have A': "A = A\# M + (A - M) \# N" using A(1) H by blast moreover have B': "B = (B - N) \# M + B\# N" using A(1) H[of B A N M] by (auto simp: ac_simps) moreover have "M = A \# M + (B - N) \# M" using H[of M N A B] BN_MA[symmetric] A(1) by (metis (no_types) diff_intersect_left_idem diff_union_cancelR multiset_inter_commute subset_mset.diff_add subset_mset.inf.cobounded1 union_commute) moreover have "N = (A - M) \# N + B \# N" by (metis A' assms(1) diff_union_cancelL inter_union_distrib_left inter_union_distrib_right mset_subset_eq_multiset_union_diff_commute subset_mset.inf.cobounded1 subset_mset.inf.commute) ultimately show P using A(2) by blast qed subsubsection \Singleton multisets\ lemma mset_size_le1_cases[case_names empty singleton,consumes 1]: "\ size M \ Suc 0; M={#} \ P; !!m. M={#m#} \ P \ \ P" by (cases M) auto lemma diff_union_single_conv2: "a \# J \ J + I - {#a#} = (J - {#a#}) + I" by simp lemmas diff_union_single_convs = diff_union_single_conv diff_union_single_conv2 lemma mset_contains_eq: "(m \# M) = ({#m#}+(M-{#m#})=M)" using diff_single_trivial by fastforce subsubsection \Pointwise ordering\ (*declare mset_le_trans[trans] Seems to be in there now. Why is this not done in Multiset.thy or order-class ? *) lemma mset_le_incr_right1: "a\#(b::'a multiset) \ a\#b+c" using mset_subset_eq_mono_add[of a b "{#}" c, simplified] . lemma mset_le_incr_right2: "a\#(b::'a multiset) \ a\#c+b" using mset_le_incr_right1 by (auto simp add: union_commute) lemmas mset_le_incr_right = mset_le_incr_right1 mset_le_incr_right2 lemma mset_le_decr_left1: "a+c\#(b::'a multiset) \ a\#b" using mset_le_incr_right1 mset_subset_eq_mono_add_right_cancel by blast lemma mset_le_decr_left2: "c+a\#(b::'a multiset) \ a\#b" using mset_le_decr_left1 by (auto simp add: union_ac) lemma mset_le_add_mset_decr_left1: "add_mset c a\#(b::'a multiset) \ a\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemma mset_le_add_mset_decr_left2: "add_mset c a\#(b::'a multiset) \ {#c#}\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemmas mset_le_decr_left = mset_le_decr_left1 mset_le_decr_left2 mset_le_add_mset_decr_left1 mset_le_add_mset_decr_left2 lemma mset_le_subtract: "A\#B \ A-C \# B-(C::'a multiset)" apply (unfold subseteq_mset_def count_diff) apply clarify apply (subgoal_tac "count A a \ count B a") apply arith apply simp done lemma mset_union_subset: "A+B \# C \ A\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_add_mset: "add_mset x B \# C \ {#x#}\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_subtract_left: "A+B \# (X::'a multiset) \ B \# X-A \ A\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "A"] mset_union_subset) lemma mset_le_subtract_right: "A+B \# (X::'a multiset) \ A \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "B"] mset_union_subset) lemma mset_le_subtract_add_mset_left: "add_mset x B \# (X::'a multiset) \ B \# X-{#x#} \ {#x#}\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "{#x#}"] mset_le_add_mset) lemma mset_le_subtract_add_mset_right: "add_mset x B \# (X::'a multiset) \ {#x#} \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "B"] mset_le_add_mset) lemma mset_le_addE: "\ xs \# (ys::'a multiset); !!zs. ys=xs+zs \ P \ \ P" using mset_subset_eq_exists_conv by blast declare subset_mset.diff_add[simp, intro] lemma mset_2dist2_cases: assumes A: "{#a#}+{#b#} \# A+B" assumes CASES: "{#a#}+{#b#} \# A \ P" "{#a#}+{#b#} \# B \ P" "\a \# A; b \# B\ \ P" "\a \# B; b \# A\ \ P" shows "P" proof - { assume C: "a \# A" "b \# A-{#a#}" with mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "A-{#a#}"] have "{#a#}+{#b#} \# A" by auto } moreover { assume C: "a \# A" "\ (b \# A-{#a#})" with A have "b \# B" by (metis diff_union_single_conv2 mset_le_subtract_left mset_subset_eq_insertD mset_un_cases) } moreover { assume C: "\ (a \# A)" "b \# B-{#a#}" with A have "a \# B" by (auto dest: mset_subset_eqD) with C mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "B-{#a#}"] have "{#a#}+{#b#} \# B" by auto } moreover { assume C: "\ (a \# A)" "\ (b \# B-{#a#})" with A have "a \# B \ b \# A" apply (intro conjI) apply (auto dest!: mset_subset_eq_insertD simp: insert_union_subset_iff; fail)[] by (metis mset_diff_cancel1elem mset_le_subtract_left multiset_diff_union_assoc single_subset_iff subset_eq_diff_conv) } ultimately show P using CASES by blast qed lemma mset_union_subset_s: "{#a#}+B \# C \ a \# C \ B \# C" by (drule mset_union_subset) simp (* TODO: Check which of these lemmas are already introduced by order-classes ! *) lemma mset_le_single_cases[consumes 1, case_names empty singleton]: "\M\#{#a#}; M={#} \ P; M={#a#} \ P\ \ P" by (induct M) auto lemma mset_le_distrib[consumes 1, case_names dist]: "\(X::'a multiset)\#A+B; !!Xa Xb. \X=Xa+Xb; Xa\#A; Xb\#B\ \ P \ \ P" by (auto elim!: mset_le_addE mset_distrib) lemma mset_le_mono_add_single: "\a \# ys; b \# ws\ \ {#a#} + {#b#} \# ys + ws" by (meson mset_subset_eq_mono_add single_subset_iff) lemma mset_size1elem: "\size P \ 1; q \# P\ \ P={#q#}" by (auto elim: mset_size_le1_cases) lemma mset_size2elem: "\size P \ 2; {#q#}+{#q'#} \# P\ \ P={#q#}+{#q'#}" by (auto elim: mset_le_addE) subsubsection \Image under function\ notation image_mset (infixr "`#" 90) lemma mset_map_split_orig: "!!M1 M2. \f `# P = M1+M2; !!P1 P2. \P=P1+P2; f `# P1 = M1; f `# P2 = M2\ \ Q \ \ Q" by (induct P) (force elim!: mset_un_single_un_cases)+ lemma mset_map_id: "\!!x. f (g x) = x\ \ f `# g `# X = X" by (induct X) auto text \The following is a very specialized lemma. Intuitively, it splits the original multiset by a splitting of some pointwise supermultiset of its image. Application: This lemma came in handy when proving the correctness of a constraint system that collects at most k sized submultisets of the sets of spawned threads. \ lemma mset_map_split_orig_le: assumes A: "f `# P \# M1+M2" and EX: "!!P1 P2. \P=P1+P2; f `# P1 \# M1; f `# P2 \# M2\ \ Q" shows "Q" using A EX by (auto elim: mset_le_distrib mset_map_split_orig) subsection \Lists\ lemma len_greater_imp_nonempty[simp]: "length l > x \ l\[]" by auto lemma list_take_induct_tl2: "\length xs = length ys; \n \ \n < length (tl xs). P ((tl ys) ! n) ((tl xs) ! n)" by (induct xs ys rule: list_induct2) auto lemma not_distinct_split_distinct: assumes "\ distinct xs" obtains y ys zs where "distinct ys" "y \ set ys" "xs = ys@[y]@zs" using assms proof (induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) thus ?case by (cases "distinct xs") auto qed lemma distinct_length_le: assumes d: "distinct ys" and eq: "set ys = set xs" shows "length ys \ length xs" proof - from d have "length ys = card (set ys)" by (simp add: distinct_card) also from eq List.card_set have "card (set ys) = length (remdups xs)" by simp also have "... \ length xs" by simp finally show ?thesis . qed lemma find_SomeD: "List.find P xs = Some x \ P x" "List.find P xs = Some x \ x\set xs" by (auto simp add: find_Some_iff) lemma in_hd_or_tl_conv[simp]: "l\[] \ x=hd l \ x\set (tl l) \ x\set l" by (cases l) auto lemma length_dropWhile_takeWhile: assumes "x < length (dropWhile P xs)" shows "x + length (takeWhile P xs) < length xs" using assms by (induct xs) auto text \Elim-version of @{thm neq_Nil_conv}.\ lemma neq_NilE: assumes "l\[]" obtains x xs where "l=x#xs" using assms by (metis list.exhaust) lemma length_Suc_rev_conv: "length xs = Suc n \ (\ys y. xs=ys@[y] \ length ys = n)" by (cases xs rule: rev_cases) auto subsubsection \List Destructors\ lemma not_hd_in_tl: "x \ hd xs \ x \ set xs \ x \ set (tl xs)" by (induct xs) simp_all lemma distinct_hd_tl: "distinct xs \ x = hd xs \ x \ set (tl (xs))" by (induct xs) simp_all lemma in_set_tlD: "x \ set (tl xs) \ x \ set xs" by (induct xs) simp_all lemma nth_tl: "xs \ [] \ tl xs ! n = xs ! Suc n" by (induct xs) simp_all lemma tl_subset: "xs \ [] \ set xs \ A \ set (tl xs) \ A" by (metis in_set_tlD rev_subsetD subsetI) lemma tl_last: "tl xs \ [] \ last xs = last (tl xs)" by (induct xs) simp_all lemma tl_obtain_elem: assumes "xs \ []" "tl xs = []" obtains e where "xs = [e]" using assms by (induct xs rule: list_nonempty_induct) simp_all lemma butlast_subset: "xs \ [] \ set xs \ A \ set (butlast xs) \ A" by (metis in_set_butlastD rev_subsetD subsetI) lemma butlast_rev_tl: "xs \ [] \ butlast (rev xs) = rev (tl xs)" by (induct xs rule: rev_induct) simp_all lemma hd_butlast: "length xs > 1 \ hd (butlast xs) = hd xs" by (induct xs) simp_all lemma butlast_upd_last_eq[simp]: "length l \ 2 \ (butlast l) [ length l - 2 := x ] = take (length l - 2) l @ [x]" apply (case_tac l rule: rev_cases) apply simp apply simp apply (case_tac ys rule: rev_cases) apply simp apply simp done lemma distinct_butlast_swap[simp]: "distinct pq \ distinct (butlast (pq[i := last pq]))" apply (cases pq rule: rev_cases) apply (auto simp: list_update_append distinct_list_update split: nat.split) done subsubsection \Splitting list according to structure of other list\ context begin private definition "SPLIT_ACCORDING m l \ length l = length m" private lemma SPLIT_ACCORDINGE: assumes "length m = length l" obtains "SPLIT_ACCORDING m l" unfolding SPLIT_ACCORDING_def using assms by auto private lemma SPLIT_ACCORDING_simp: "SPLIT_ACCORDING m (l1@l2) \ (\m1 m2. m=m1@m2 \ SPLIT_ACCORDING m1 l1 \ SPLIT_ACCORDING m2 l2)" "SPLIT_ACCORDING m (x#l') \ (\y m'. m=y#m' \ SPLIT_ACCORDING m' l')" apply (fastforce simp: SPLIT_ACCORDING_def intro: exI[where x = "take (length l1) m"] exI[where x = "drop (length l1) m"]) apply (cases m;auto simp: SPLIT_ACCORDING_def) done text \Split structure of list @{term m} according to structure of list @{term l}.\ method split_list_according for m :: "'a list" and l :: "'b list" = (rule SPLIT_ACCORDINGE[of m l], (simp; fail), ( simp only: SPLIT_ACCORDING_simp, elim exE conjE, simp only: SPLIT_ACCORDING_def ) ) end subsubsection \\list_all2\\ lemma list_all2_induct[consumes 1, case_names Nil Cons]: assumes "list_all2 P l l'" assumes "Q [] []" assumes "\x x' ls ls'. \ P x x'; list_all2 P ls ls'; Q ls ls' \ \ Q (x#ls) (x'#ls')" shows "Q l l'" using list_all2_lengthD[OF assms(1)] assms apply (induct rule: list_induct2) apply auto done subsubsection \Indexing\ lemma ran_nth_set_encoding_conv[simp]: "ran (\i. if ii l[j:=x]!i = l!i" apply (induction l arbitrary: i j) apply (auto split: nat.splits) done lemma nth_list_update': "l[i:=x]!j = (if i=j \ i length l \ n\0 \ last (take n l) = l!(n - 1)" apply (induction l arbitrary: n) apply (auto simp: take_Cons split: nat.split) done lemma nth_append_first[simp]: "i (l@l')!i = l!i" by (simp add: nth_append) lemma in_set_image_conv_nth: "f x \ f`set l \ (\ii. i f (l!i) = f (l'!i)" shows "f`set l = f`set l'" using assms by (fastforce simp: in_set_conv_nth in_set_image_conv_nth) lemma insert_swap_set_eq: "i insert (l!i) (set (l[i:=x])) = insert x (set l)" by (auto simp: in_set_conv_nth nth_list_update split: if_split_asm) subsubsection \Reverse lists\ lemma neq_Nil_revE: assumes "l\[]" obtains ll e where "l = ll@[e]" using assms by (cases l rule: rev_cases) auto lemma neq_Nil_rev_conv: "l\[] \ (\xs x. l = xs@[x])" by (cases l rule: rev_cases) auto text \Caution: Same order of case variables in snoc-case as @{thm [source] rev_exhaust}, the other way round than @{thm [source] rev_induct} !\ lemma length_compl_rev_induct[case_names Nil snoc]: "\P []; !! l e . \!! ll . length ll <= length l \ P ll\ \ P (l@[e])\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs" rule: rev_cases) apply(auto) done lemma list_append_eq_Cons_cases[consumes 1]: "\ys@zs = x#xs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: append_eq_Cons_conv) lemma list_Cons_eq_append_cases[consumes 1]: "\x#xs = ys@zs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: Cons_eq_append_conv) lemma map_of_rev_distinct[simp]: "distinct (map fst m) \ map_of (rev m) = map_of m" apply (induct m) apply simp apply simp apply (subst map_add_comm) apply force apply simp done \ \Tail-recursive, generalized @{const rev}. May also be used for tail-recursively getting a list with all elements of the two operands, if the order does not matter, e.g. when implementing sets by lists.\ fun revg where "revg [] b = b" | "revg (a#as) b = revg as (a#b)" lemma revg_fun[simp]: "revg a b = rev a @ b" by (induct a arbitrary: b) auto lemma rev_split_conv[simp]: "l \ [] \ rev (tl l) @ [hd l] = rev l" by (induct l) simp_all lemma rev_butlast_is_tl_rev: "rev (butlast l) = tl (rev l)" by (induct l) auto lemma hd_last_singletonI: "\xs \ []; hd xs = last xs; distinct xs\ \ xs = [hd xs]" by (induct xs rule: list_nonempty_induct) auto lemma last_filter: "\xs \ []; P (last xs)\ \ last (filter P xs) = last xs" by (induct xs rule: rev_nonempty_induct) simp_all (* As the following two rules are similar in nature to list_induct2', they are named accordingly. *) lemma rev_induct2' [case_names empty snocl snocr snoclr]: assumes empty: "P [] []" and snocl: "\x xs. P (xs@[x]) []" and snocr: "\y ys. P [] (ys@[y])" and snoclr: "\x xs y ys. P xs ys \ P (xs@[x]) (ys@[y])" shows "P xs ys" proof (induct xs arbitrary: ys rule: rev_induct) case Nil thus ?case using empty snocr by (cases ys rule: rev_exhaust) simp_all next case snoc thus ?case using snocl snoclr by (cases ys rule: rev_exhaust) simp_all qed lemma rev_nonempty_induct2' [case_names single snocl snocr snoclr, consumes 2]: assumes "xs \ []" "ys \ []" assumes single': "\x y. P [x] [y]" and snocl: "\x xs y. xs \ [] \ P (xs@[x]) [y]" and snocr: "\x y ys. ys \ [] \ P [x] (ys@[y])" and snoclr: "\x xs y ys. \P xs ys; xs \ []; ys\[]\ \ P (xs@[x]) (ys@[y])" shows "P xs ys" using assms(1,2) proof (induct xs arbitrary: ys rule: rev_nonempty_induct) case single then obtain z zs where "ys = zs@[z]" by (metis rev_exhaust) thus ?case using single' snocr by (cases "zs = []") simp_all next case (snoc x xs) then obtain z zs where zs: "ys = zs@[z]" by (metis rev_exhaust) thus ?case using snocl snoclr snoc by (cases "zs = []") simp_all qed subsubsection "Folding" text "Ugly lemma about foldl over associative operator with left and right neutral element" lemma foldl_A1_eq: "!!i. \ !! e. f n e = e; !! e. f e n = e; !!a b c. f a (f b c) = f (f a b) c \ \ foldl f i ww = f i (foldl f n ww)" proof (induct ww) case Nil thus ?case by simp next case (Cons a ww i) note IHP[simplified]=this have "foldl f i (a # ww) = foldl f (f i a) ww" by simp also from IHP have "\ = f (f i a) (foldl f n ww)" by blast also from IHP(4) have "\ = f i (f a (foldl f n ww))" by simp also from IHP(1)[OF IHP(2,3,4), where i=a] have "\ = f i (foldl f a ww)" by simp also from IHP(2)[of a] have "\ = f i (foldl f (f n a) ww)" by simp also have "\ = f i (foldl f n (a#ww))" by simp finally show ?case . qed lemmas foldl_conc_empty_eq = foldl_A1_eq[of "(@)" "[]", simplified] lemmas foldl_un_empty_eq = foldl_A1_eq[of "(\)" "{}", simplified, OF Un_assoc[symmetric]] lemma foldl_set: "foldl (\) {} l = \{x. x\set l}" apply (induct l) apply simp_all apply (subst foldl_un_empty_eq) apply auto done lemma (in monoid_mult) foldl_absorb1: "x*foldl (*) 1 zs = foldl (*) x zs" apply (rule sym) apply (rule foldl_A1_eq) apply (auto simp add: mult.assoc) done text \Towards an invariant rule for foldl\ lemma foldl_rule_aux: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" shows "I (foldl f \0 l0) []" using initial step apply (induct l0 arbitrary: \0) apply auto done lemma foldl_rule_aux_P: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" assumes final: "!!\. I \ [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule_aux[of I \0 l0, OF initial, OF step] final by simp lemma foldl_rule: fixes I :: "'\ \ 'a list \ 'a list \ bool" assumes initial: "I \0 [] l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" shows "I (foldl f \0 l0) l0 []" using initial step apply (rule_tac I="\\ lr. \ll. l0=ll@lr \ I \ ll lr" in foldl_rule_aux_P) apply auto done text \ Invariant rule for foldl. The invariant is parameterized with the state, the list of items that have already been processed and the list of items that still have to be processed. \ lemma foldl_rule_P: fixes I :: "'\ \ 'a list \ 'a list \ bool" \ \The invariant holds for the initial state, no items processed yet and all items to be processed:\ assumes initial: "I \0 [] l0" \ \The invariant remains valid if one item from the list is processed\ assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" \ \The proposition follows from the invariant in the final state, i.e. all items processed and nothing to be processed\ assumes final: "!!\. I \ l0 [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule[of I, OF initial step] by (simp add: final) text \Invariant reasoning over @{const foldl} for distinct lists. Invariant rule makes no assumptions about ordering.\ lemma distinct_foldl_invar: "\ distinct S; I (set S) \0; \x it \. \x \ it; it \ set S; I it \\ \ I (it - {x}) (f \ x) \ \ I {} (foldl f \0 S)" proof (induct S arbitrary: \0) case Nil thus ?case by auto next case (Cons x S) note [simp] = Cons.prems(1)[simplified] show ?case apply simp apply (rule Cons.hyps) proof - from Cons.prems(1) show "distinct S" by simp from Cons.prems(3)[of x "set (x#S)", simplified, OF Cons.prems(2)[simplified]] show "I (set S) (f \0 x)" . fix xx it \ assume A: "xx\it" "it \ set S" "I it \" show "I (it - {xx}) (f \ xx)" using A(2) apply (rule_tac Cons.prems(3)) apply (simp_all add: A(1,3)) apply blast done qed qed lemma foldl_length_aux: "foldl (\i x. Suc i) a l = a + length l" by (induct l arbitrary: a) auto lemmas foldl_length[simp] = foldl_length_aux[where a=0, simplified] lemma foldr_length_aux: "foldr (\x i. Suc i) l a = a + length l" by (induct l arbitrary: a rule: rev_induct) auto lemmas foldr_length[simp] = foldr_length_aux[where a=0, simplified] context comp_fun_commute begin lemma foldl_f_commute: "f a (foldl (\a b. f b a) b xs) = foldl (\a b. f b a) (f a b) xs" by(induct xs arbitrary: b)(simp_all add: fun_left_comm) lemma foldr_conv_foldl: "foldr f xs a = foldl (\a b. f b a) a xs" by(induct xs arbitrary: a)(simp_all add: foldl_f_commute) end lemma filter_conv_foldr: "filter P xs = foldr (\x xs. if P x then x # xs else xs) xs []" by(induct xs) simp_all lemma foldr_Cons: "foldr Cons xs [] = xs" by(induct xs) simp_all lemma foldr_snd_zip: "length xs \ length ys \ foldr (\(x, y). f y) (zip xs ys) b = foldr f ys b" proof(induct ys arbitrary: xs) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma foldl_snd_zip: "length xs \ length ys \ foldl (\b (x, y). f b y) b (zip xs ys) = foldl f b ys" proof(induct ys arbitrary: xs b) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma fst_foldl: "fst (foldl (\(a, b) x. (f a x, g a b x)) (a, b) xs) = foldl f a xs" by(induct xs arbitrary: a b) simp_all lemma foldl_foldl_conv_concat: "foldl (foldl f) a xs = foldl f a (concat xs)" by(induct xs arbitrary: a) simp_all lemma foldl_list_update: "n < length xs \ foldl f a (xs[n := x]) = foldl f (f (foldl f a (take n xs)) x) (drop (Suc n) xs)" by(simp add: upd_conv_take_nth_drop) lemma map_by_foldl: fixes l :: "'a list" and f :: "'a \ 'b" shows "foldl (\l x. l@[f x]) [] l = map f l" proof - { fix l' have "foldl (\l x. l@[f x]) l' l = l'@map f l" by (induct l arbitrary: l') auto } thus ?thesis by simp qed subsubsection \Sorting\ lemma sorted_in_between: assumes A: "0\i" "i x" "xk" and "kx" and "xk. i\k \ k l!k\x \ x x") case True from True Suc.hyps have "d = j - (i + 1)" by simp moreover from True have "i+1 < j" by (metis Suc.prems Suc_eq_plus1 Suc_lessI not_less) moreover from True have "0\i+1" by simp ultimately obtain k where "i+1\k" "k x" "xk" "k x" "xsorted l; l\[]\ \ hd l \ last l" -by (metis eq_iff hd_Cons_tl last_in_set not_hd_in_tl sorted.simps(2)) +by (metis eq_iff hd_Cons_tl last_in_set not_hd_in_tl sorted_wrt.simps(2)) lemma (in linorder) sorted_hd_min: "\xs \ []; sorted xs\ \ \x \ set xs. hd xs \ x" by (induct xs, auto) lemma sorted_append_bigger: "\sorted xs; \x \ set xs. x \ y\ \ sorted (xs @ [y])" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) then have s: "sorted xs" by (cases xs) simp_all from Cons have a: "\x\set xs. x \ y" by simp from Cons(1)[OF s a] Cons(2-) show ?case by (cases xs) simp_all qed lemma sorted_filter': "sorted l \ sorted (filter P l)" using sorted_filter[where f=id, simplified] . subsubsection \Map\ (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] *) lemma map_eq_consE: "\map f ls = fa#fl; !!a l. \ ls=a#l; f a=fa; map f l = fl \ \ P\ \ P" by auto lemma map_fst_mk_snd[simp]: "map fst (map (\x. (x,k)) l) = l" by (induct l) auto lemma map_snd_mk_fst[simp]: "map snd (map (\x. (k,x)) l) = l" by (induct l) auto lemma map_fst_mk_fst[simp]: "map fst (map (\x. (k,x)) l) = replicate (length l) k" by (induct l) auto lemma map_snd_mk_snd[simp]: "map snd (map (\x. (x,k)) l) = replicate (length l) k" by (induct l) auto lemma map_zip1: "map (\x. (x,k)) l = zip l (replicate (length l) k)" by (induct l) auto lemma map_zip2: "map (\x. (k,x)) l = zip (replicate (length l) k) l" by (induct l) auto lemmas map_zip=map_zip1 map_zip2 (* TODO/FIXME: hope nobody changes nth to be underdefined! *) lemma map_eq_nth_eq: assumes A: "map f l = map f l'" shows "f (l!i) = f (l'!i)" proof - from A have "length l = length l'" by (metis length_map) thus ?thesis using A apply (induct arbitrary: i rule: list_induct2) apply simp apply (simp add: nth_def split: nat.split) done qed lemma map_upd_eq: "\i f (l!i) = f x\ \ map f (l[i:=x]) = map f l" by (metis list_update_beyond list_update_id map_update not_le_imp_less) lemma inj_map_inv_f [simp]: "inj f \ map (inv f) (map f l) = l" by (simp) lemma inj_on_map_the: "\D \ dom m; inj_on m D\ \ inj_on (the\m) D" apply (rule inj_onI) apply simp apply (case_tac "m x") apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply (auto intro: inj_onD) [1] apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply simp apply (rule inj_onD) apply assumption apply auto done lemma map_consI: "w=map f ww \ f a#w = map f (a#ww)" "w@l=map f ww@l \ f a#w@l = map f (a#ww)@l" by auto lemma restrict_map_subset_eq: fixes R shows "\m |` R = m'; R'\R\ \ m|` R' = m' |` R'" by (auto simp add: Int_absorb1) lemma restrict_map_self[simp]: "m |` dom m = m" apply (rule ext) apply (case_tac "m x") apply (auto simp add: restrict_map_def) done lemma restrict_map_UNIV[simp]: "f |` UNIV = f" by (auto simp add: restrict_map_def) lemma restrict_map_inv[simp]: "f |` (- dom f) = Map.empty" by (auto simp add: restrict_map_def intro: ext) lemma restrict_map_upd: "(f |` S)(k \ v) = f(k\v) |` (insert k S)" by (auto simp add: restrict_map_def intro: ext) lemma map_upd_eq_restrict: "m (x:=None) = m |` (-{x})" by (auto intro: ext) declare Map.finite_dom_map_of [simp, intro!] lemma dom_const'[simp]: "dom (\x. Some (f x)) = UNIV" by auto lemma restrict_map_eq : "((m |` A) k = None) \ (k \ dom m \ A)" "((m |` A) k = Some v) \ (m k = Some v \ k \ A)" unfolding restrict_map_def by (simp_all add: dom_def) definition "rel_of m P == {(k,v). m k = Some v \ P (k, v)}" lemma rel_of_empty[simp]: "rel_of Map.empty P = {}" by (auto simp add: rel_of_def) lemma remove1_tl: "xs \ [] \ remove1 (hd xs) xs = tl xs" by (cases xs) auto lemma set_oo_map_alt: "(set \\ map) f = (\l. f ` set l)" by auto subsubsection "Filter and Revert" primrec filter_rev_aux where "filter_rev_aux a P [] = a" | "filter_rev_aux a P (x#xs) = ( if P x then filter_rev_aux (x#a) P xs else filter_rev_aux a P xs)" lemma filter_rev_aux_alt: "filter_rev_aux a P l = filter P (rev l) @ a" by (induct l arbitrary: a) auto definition "filter_rev == filter_rev_aux []" lemma filter_rev_alt: "filter_rev P l = filter P (rev l)" unfolding filter_rev_def by (simp add: filter_rev_aux_alt) definition "remove_rev x == filter_rev (Not o (=) x)" lemma remove_rev_alt_def : "remove_rev x xs = (filter (\y. y \ x) (rev xs))" unfolding remove_rev_def apply (simp add: filter_rev_alt comp_def) by metis subsubsection "zip" declare zip_map_fst_snd[simp] lemma pair_list_split: "\ !!l1 l2. \ l = zip l1 l2; length l1=length l2; length l=length l2 \ \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) from Cons.hyps obtain l1 l2 where IHAPP: "l=zip l1 l2" "length l1 = length l2" "length l=length l2" . obtain a1 a2 where [simp]: "a=(a1,a2)" by (cases a) auto from IHAPP have "a#l = zip (a1#l1) (a2#l2)" "length (a1#l1) = length (a2#l2)" "length (a#l) = length (a2#l2)" by (simp_all only:) (simp_all (no_asm_use)) with Cons.prems show ?case by blast qed lemma set_zip_cart: "x\set (zip l l') \ x\set l \ set l'" by (auto simp add: set_zip) lemma map_prod_fun_zip: "map (\(x, y). (f x, g y)) (zip xs ys) = zip (map f xs) (map g ys)" proof(induct xs arbitrary: ys) case Nil thus ?case by simp next case (Cons x xs) thus ?case by(cases ys) simp_all qed subsubsection \Generalized Zip\ text \Zip two lists element-wise, where the combination of two elements is specified by a function. Note that this function is underdefined for lists of different length.\ fun zipf :: "('a\'b\'c) \ 'a list \ 'b list \ 'c list" where "zipf f [] [] = []" | "zipf f (a#as) (b#bs) = f a b # zipf f as bs" lemma zipf_zip: "\length l1 = length l2\ \ zipf Pair l1 l2 = zip l1 l2" apply (induct l1 arbitrary: l2) apply auto apply (case_tac l2) apply auto done \ \All quantification over zipped lists\ fun list_all_zip where "list_all_zip P [] [] \ True" | "list_all_zip P (a#as) (b#bs) \ P a b \ list_all_zip P as bs" | "list_all_zip P _ _ \ False" lemma list_all_zip_alt: "list_all_zip P as bs \ length as = length bs \ (\iP as bs rule: list_all_zip.induct) apply auto apply (case_tac i) apply auto done lemma list_all_zip_map1: "list_all_zip P (List.map f as) bs \ list_all_zip (\a b. P (f a) b) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done lemma list_all_zip_map2: "list_all_zip P as (List.map f bs) \ list_all_zip (\a b. P a (f b)) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done declare list_all_zip_alt[mono] lemma lazI[intro?]: "\ length a = length b; !!i. i P (a!i) (b!i) \ \ list_all_zip P a b" by (auto simp add: list_all_zip_alt) lemma laz_conj[simp]: "list_all_zip (\x y. P x y \ Q x y) a b \ list_all_zip P a b \ list_all_zip Q a b" by (auto simp add: list_all_zip_alt) lemma laz_len: "list_all_zip P a b \ length a = length b" by (simp add: list_all_zip_alt) lemma laz_eq: "list_all_zip (=) a b \ a=b" apply (induct a arbitrary: b) apply (case_tac b) apply simp apply simp apply (case_tac b) apply simp apply simp done lemma laz_swap_ex: assumes A: "list_all_zip (\a b. \c. P a b c) A B" obtains C where "list_all_zip (\a c. \b. P a b c) A C" "list_all_zip (\b c. \a. P a b c) B C" proof - from A have [simp]: "length A = length B" and IC: "\ici. P (A!i) (B!i) ci" by (auto simp add: list_all_zip_alt) from obtain_list_from_elements[OF IC] obtain C where "length C = length B" "\ia b. P a) A B \ (length A = length B) \ (\a\set A. P a)" by (auto simp add: list_all_zip_alt set_conv_nth) lemma laz_weak_Pb[simp]: "list_all_zip (\a b. P b) A B \ (length A = length B) \ (\b\set B. P b)" by (force simp add: list_all_zip_alt set_conv_nth) subsubsection "Collecting Sets over Lists" definition "list_collect_set f l == \{ f a | a. a\set l }" lemma list_collect_set_simps[simp]: "list_collect_set f [] = {}" "list_collect_set f [a] = f a" "list_collect_set f (a#l) = f a \ list_collect_set f l" "list_collect_set f (l@l') = list_collect_set f l \ list_collect_set f l'" by (unfold list_collect_set_def) auto lemma list_collect_set_map_simps[simp]: "list_collect_set f (map x []) = {}" "list_collect_set f (map x [a]) = f (x a)" "list_collect_set f (map x (a#l)) = f (x a) \ list_collect_set f (map x l)" "list_collect_set f (map x (l@l')) = list_collect_set f (map x l) \ list_collect_set f (map x l')" by simp_all lemma list_collect_set_alt: "list_collect_set f l = \{ f (l!i) | i. i(set (map f l))" by (unfold list_collect_set_def) auto subsubsection \Sorted List with arbitrary Relations\ lemma (in linorder) sorted_wrt_rev_linord [simp] : "sorted_wrt (\) l \ sorted (rev l)" -by (simp add: sorted_sorted_wrt sorted_wrt_rev) +by (simp add: sorted_wrt_rev) lemma (in linorder) sorted_wrt_map_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (map fst l)" -by (simp add: sorted_sorted_wrt sorted_wrt_map) +by (simp add: sorted_wrt_map) lemma (in linorder) sorted_wrt_map_rev_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (rev (map fst l))" by (induct l) (auto simp add: sorted_append) subsubsection \Take and Drop\ lemma take_update[simp]: "take n (l[i:=x]) = (take n l)[i:=x]" apply (induct l arbitrary: n i) apply (auto split: nat.split) apply (case_tac n) apply simp_all apply (case_tac n) apply simp_all done lemma take_update_last: "length list>n \ (take (Suc n) list) [n:=x] = take n list @ [x]" by (induct list arbitrary: n) (auto split: nat.split) lemma drop_upd_irrelevant: "m < n \ drop n (l[m:=x]) = drop n l" apply (induct n arbitrary: l m) apply simp apply (case_tac l) apply (auto split: nat.split) done lemma set_drop_conv: "set (drop n l) = { l!i | i. n\i \ i < length l }" (is "?L=?R") proof (intro equalityI subsetI) fix x assume "x\?L" then obtain i where L: "i = l!(n+i)" using L by simp finally show "x\?R" using L by auto next fix x assume "x\?R" then obtain i where L: "n\i" "i?L" by (auto simp add: in_set_conv_nth) qed lemma filter_upt_take_conv: "[i\[n..[n..set (drop n l) \ (\i. n\i \ i x = l!i)" apply (clarsimp simp: in_set_conv_nth) apply safe apply simp apply (metis le_add2 less_diff_conv add.commute) apply (rule_tac x="i-n" in exI) apply auto [] done lemma Union_take_drop_id: "\(set (drop n l)) \ \(set (take n l)) = \(set l)" by (metis Union_Un_distrib append_take_drop_id set_union_code sup_commute) lemma Un_set_drop_extend: "\j\Suc 0; j < length l\ \ l ! (j - Suc 0) \ \(set (drop j l)) = \(set (drop (j - Suc 0) l))" apply safe apply simp_all apply (metis diff_Suc_Suc diff_zero gr0_implies_Suc in_set_drop_conv_nth le_refl less_eq_Suc_le order.strict_iff_order) apply (metis Nat.diff_le_self set_drop_subset_set_drop subset_code(1)) by (metis diff_Suc_Suc gr0_implies_Suc in_set_drop_conv_nth less_eq_Suc_le order.strict_iff_order minus_nat.diff_0) lemma drop_take_drop_unsplit: "i\j \ drop i (take j l) @ drop j l = drop i l" proof - assume "i \ j" then obtain skf where "i + skf = j" by (metis le_iff_add) thus "drop i (take j l) @ drop j l = drop i l" by (metis append_take_drop_id diff_add_inverse drop_drop drop_take add.commute) qed lemma drop_last_conv[simp]: "l\[] \ drop (length l - Suc 0) l = [last l]" by (cases l rule: rev_cases) auto lemma take_butlast_conv[simp]: "take (length l - Suc 0) l = butlast l" by (cases l rule: rev_cases) auto lemma drop_takeWhile: assumes "i\length (takeWhile P l)" shows "drop i (takeWhile P l) = takeWhile P (drop i l)" using assms proof (induction l arbitrary: i) case Nil thus ?case by auto next case (Cons x l) thus ?case by (cases i) auto qed lemma less_length_takeWhile_conv: "i < length (takeWhile P l) \ (i (\j\i. P (l!j)))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (metis dual_order.strict_trans2 nth_mem set_takeWhileD takeWhile_nth) subgoal by (meson less_le_trans not_le_imp_less nth_length_takeWhile) done lemma eq_len_takeWhile_conv: "i=length (takeWhile P l) \ i\length l \ (\j (i \P (l!i))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (auto simp: less_length_takeWhile_conv) subgoal using nth_length_takeWhile by blast subgoal by (metis length_takeWhile_le nth_length_takeWhile order.order_iff_strict) subgoal by (metis dual_order.strict_trans2 leI less_length_takeWhile_conv linorder_neqE_nat nth_length_takeWhile) done subsubsection \Up-to\ lemma upt_eq_append_conv: "i\j \ [i.. (\k. i\k \ k\j \ [i.. [k..j" thus "\k\i. k \ j \ [i.. [k.. length l \ map (nth l) [M.. is1 = [l.. is2 = [Suc i.. l\i \ ii. i + ofs) [a..Last and butlast\ lemma butlast_upt: "butlast [m.. butlast [n..length l \ take (n - Suc 0) l = butlast (take n l)" by (simp add: butlast_take) lemma butlast_eq_cons_conv: "butlast l = x#xs \ (\xl. l=x#xs@[xl])" by (metis Cons_eq_appendI append_butlast_last_id butlast.simps butlast_snoc eq_Nil_appendI) lemma butlast_eq_consE: assumes "butlast l = x#xs" obtains xl where "l=x#xs@[xl]" using assms by (auto simp: butlast_eq_cons_conv) lemma drop_eq_ConsD: "drop n xs = x # xs' \ drop (Suc n) xs = xs'" by(induct xs arbitrary: n)(simp_all add: drop_Cons split: nat.split_asm) subsubsection \List Slices\ text \Based on Lars Hupel's code.\ definition slice :: "nat \ nat \ 'a list \ 'a list" where "slice from to list = take (to - from) (drop from list)" lemma slice_len[simp]: "\ from \ to; to \ length xs \ \ length (slice from to xs) = to - from" unfolding slice_def by simp lemma slice_head: "\ from < to; to \ length xs \ \ hd (slice from to xs) = xs ! from" unfolding slice_def proof - assume a1: "from < to" assume "to \ length xs" then have "\n. to - (to - n) \ length (take to xs)" by (metis (no_types) slice_def diff_le_self drop_take length_drop slice_len) then show "hd (take (to - from) (drop from xs)) = xs ! from" using a1 by (metis diff_diff_cancel drop_take hd_drop_conv_nth leI le_antisym less_or_eq_imp_le nth_take) qed lemma slice_eq_bounds_empty[simp]: "slice i i xs = []" unfolding slice_def by auto lemma slice_nth: "\ from < to; to \ length xs; i < to - from \ \ slice from to xs ! i = xs ! (from + i)" unfolding slice_def by (induction "to - from" arbitrary: "from" to i) simp+ lemma slice_prepend: "\ i \ k; k \ length xs \ \ let p = length ys in slice i k xs = slice (i + p) (k + p) (ys @ xs)" unfolding slice_def Let_def by force lemma slice_Nil[simp]: "slice begin end [] = []" unfolding slice_def by auto lemma slice_Cons: "slice begin end (x#xs) = (if begin=0 \ end>0 then x#slice begin (end-1) xs else slice (begin - 1) (end - 1) xs)" unfolding slice_def by (auto simp: take_Cons' drop_Cons') lemma slice_complete[simp]: "slice 0 (length xs) xs = xs" unfolding slice_def by simp subsubsection \Miscellaneous\ lemma length_compl_induct[case_names Nil Cons]: "\P []; !! e l . \!! ll . length ll <= length l \ P ll\ \ P (e#l)\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs") apply(auto) done lemma in_set_list_format: "\ e\set l; !!l1 l2. l=l1@e#l2 \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) show ?case proof (cases "a=e") case True with Cons show ?thesis by force next case False with Cons.prems(1) have "e\set l" by auto with Cons.hyps obtain l1 l2 where "l=l1@e#l2" by blast hence "a#l = (a#l1)@e#l2" by simp with Cons.prems(2) show P by blast qed qed lemma in_set_upd_cases: assumes "x\set (l[i:=y])" obtains "iset l" by (metis assms in_set_conv_nth length_list_update nth_list_update_eq nth_list_update_neq) lemma in_set_upd_eq_aux: assumes "iset (l[i:=y]) \ x=y \ (\y. x\set (l[i:=y]))" by (metis in_set_upd_cases assms list_update_overwrite set_update_memI) lemma in_set_upd_eq: assumes "iset (l[i:=y]) \ x=y \ (x\set l \ (\y. x\set (l[i:=y])))" by (metis in_set_upd_cases in_set_upd_eq_aux assms) text \Simultaneous induction over two lists, prepending an element to one of the lists in each step\ lemma list_2pre_induct[case_names base left right]: assumes BASE: "P [] []" and LEFT: "!!e w1' w2. P w1' w2 \ P (e#w1') w2" and RIGHT: "!!e w1 w2'. P w1 w2' \ P w1 (e#w2')" shows "P w1 w2" proof - { \ \The proof is done by induction over the sum of the lengths of the lists\ fix n have "!!w1 w2. \length w1 + length w2 = n; P [] []; !!e w1' w2. P w1' w2 \ P (e#w1') w2; !!e w1 w2'. P w1 w2' \ P w1 (e#w2') \ \ P w1 w2 " apply (induct n) apply simp apply (case_tac w1) apply auto apply (case_tac w2) apply auto done } from this[OF _ BASE LEFT RIGHT] show ?thesis by blast qed lemma list_decomp_1: "length l=1 \ \a. l=[a]" by (case_tac l, auto) lemma list_decomp_2: "length l=2 \ \a b. l=[a,b]" by (case_tac l, auto simp add: list_decomp_1) lemma list_rest_coinc: "\length s2 \ length s1; s1@r1 = s2@r2\ \ \r1p. r2=r1p@r1" by (metis append_eq_append_conv_if) lemma list_tail_coinc: "n1#r1 = n2#r2 \ n1=n2 & r1=r2" by (auto) lemma last_in_set[intro]: "\l\[]\ \ last l \ set l" by (induct l) auto lemma empty_append_eq_id[simp]: "(@) [] = (\x. x)" by auto lemma op_conc_empty_img_id[simp]: "((@) [] ` L) = L" by auto lemma distinct_match: "\ distinct (al@e#bl) \ \ (al@e#bl = al'@e#bl') \ (al=al' \ bl=bl')" proof (rule iffI, induct al arbitrary: al') case Nil thus ?case by (cases al') auto next case (Cons a al) note Cprems=Cons.prems note Chyps=Cons.hyps show ?case proof (cases al') case Nil with Cprems have False by auto thus ?thesis .. next case [simp]: (Cons a' all') with Cprems have [simp]: "a=a'" and P: "al@e#bl = all'@e#bl'" by auto from Cprems(1) have D: "distinct (al@e#bl)" by auto from Chyps[OF D P] have [simp]: "al=all'" "bl=bl'" by auto show ?thesis by simp qed qed simp lemma prop_match: "\ list_all P al; \P e; \P e'; list_all P bl \ \ (al@e#bl = al'@e'#bl') \ (al=al' \ e=e' \ bl=bl')" apply (rule iffI, induct al arbitrary: al') apply (case_tac al', fastforce, fastforce)+ done lemmas prop_matchD = rev_iffD1[OF _ prop_match[where P=P]] for P declare distinct_tl[simp] lemma list_se_match[simp]: "l1 \ [] \ l1@l2 = [a] \ l1 = [a] \ l2 = []" "l2 \ [] \ l1@l2 = [a] \ l1 = [] \ l2 = [a]" "l1 \ [] \ [a] = l1@l2 \ l1 = [a] \ l2 = []" "l2 \ [] \ [a] = l1@l2 \ l1 = [] \ l2 = [a]" apply (cases l1, simp_all) apply (cases l1, simp_all) apply (cases l1, auto) [] apply (cases l1, auto) [] done (* Placed here because it depends on xy_in_set_cases *) lemma distinct_map_eq: "\ distinct (List.map f l); f x = f y; x\set l; y\set l \ \ x=y" by (erule (2) xy_in_set_cases) auto lemma upt_append: assumes "iu'" assumes NP: "\i. u\i \ i \P i" shows "[i\[0..[0..[0..[0..[u..[u..[0.. P (l!i)" proof assume A: "P (l!i)" have "[0..i by (simp add: upt_append) also have "[i..i by (auto simp: upt_conv_Cons) finally have "[k\[0..[Suc i..P (l!i)\ by simp hence "j = last (i#[k\[Suc i.. \ i" proof - have "sorted (i#[k\[Suc i.. last ?l" by (rule sorted_hd_last) simp thus ?thesis by simp qed finally have "i\j" . thus False using \j by simp qed lemma all_set_conv_nth: "(\x\set l. P x) \ (\i *) lemma upt_0_eq_Nil_conv[simp]: "[0.. j=0" by auto lemma filter_eq_snocD: "filter P l = l'@[x] \ x\set l \ P x" proof - assume A: "filter P l = l'@[x]" hence "x\set (filter P l)" by simp thus ?thesis by simp qed lemma lists_image_witness: assumes A: "x\lists (f`Q)" obtains xo where "xo\lists Q" "x=map f xo" proof - have "\ x\lists (f`Q) \ \ \xo\lists Q. x=map f xo" proof (induct x) case Nil thus ?case by auto next case (Cons x xs) then obtain xos where "xos\lists Q" "xs=map f xos" by force moreover from Cons.prems have "x\f`Q" by auto then obtain xo where "xo\Q" "x=f xo" by auto ultimately show ?case by (rule_tac x="xo#xos" in bexI) auto qed thus ?thesis apply (simp_all add: A) apply (erule_tac bexE) apply (rule_tac that) apply assumption+ done qed lemma map_of_None_filterD: "map_of xs x = None \ map_of (filter P xs) x = None" by(induct xs) auto lemma map_of_concat: "map_of (concat xss) = foldr (\xs f. f ++ map_of xs) xss Map.empty" by(induct xss) simp_all lemma map_of_Some_split: "map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None" proof(induct xs) case (Cons x xs) obtain k' v' where x: "x = (k', v')" by(cases x) show ?case proof(cases "k' = k") case True with \map_of (x # xs) k = Some v\ x have "x # xs = [] @ (k, v) # xs" "map_of [] k = None" by simp_all thus ?thesis by blast next case False with \map_of (x # xs) k = Some v\ x have "map_of xs k = Some v" by simp from \map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None\[OF this] obtain ys zs where "xs = ys @ (k, v) # zs" "map_of ys k = None" by blast with False x have "x # xs = (x # ys) @ (k, v) # zs" "map_of (x # ys) k = None" by simp_all thus ?thesis by blast qed qed simp lemma map_add_find_left: "g k = None \ (f ++ g) k = f k" by(simp add: map_add_def) lemma map_add_left_None: "f k = None \ (f ++ g) k = g k" by(simp add: map_add_def split: option.split) lemma map_of_Some_filter_not_in: "\ map_of xs k = Some v; \ P (k, v); distinct (map fst xs) \ \ map_of (filter P xs) k = None" apply(induct xs) apply(auto) apply(auto simp add: map_of_eq_None_iff) done lemma distinct_map_fst_filterI: "distinct (map fst xs) \ distinct (map fst (filter P xs))" by(induct xs) auto lemma distinct_map_fstD: "\ distinct (map fst xs); (x, y) \ set xs; (x, z) \ set xs \ \ y = z" by(induct xs)(fastforce elim: notE rev_image_eqI)+ lemma concat_filter_neq_Nil: "concat [ys\xs. ys \ Nil] = concat xs" by(induct xs) simp_all lemma distinct_concat': "\distinct [ys\xs. ys \ Nil]; \ys. ys \ set xs \ distinct ys; \ys zs. \ys \ set xs; zs \ set xs; ys \ zs\ \ set ys \ set zs = {}\ \ distinct (concat xs)" by(erule distinct_concat[of "[ys\xs. ys \ Nil]", unfolded concat_filter_neq_Nil]) auto lemma distinct_idx: assumes "distinct (map f l)" assumes "im. n \ m \ m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" using assms proof(induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) show ?case proof(cases "P x") case [simp]: False from \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) thus ?thesis by(auto simp add: nth_append) next case [simp]: True show ?thesis proof(cases "n = length (filter P xs)") case False with \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp moreover hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) ultimately show ?thesis by(auto simp add: nth_append) next case [simp]: True hence "filter P (xs @ [x]) ! n = (xs @ [x]) ! length xs" by simp moreover have "length xs < length (xs @ [x])" by simp moreover have "length xs \ n" by simp moreover have "filter P (take (length xs) (xs @ [x])) = take n (filter P (xs @ [x]))" by simp ultimately show ?thesis by blast qed qed qed lemma set_map_filter: "set (List.map_filter g xs) = {y. \x. x \ set xs \ g x = Some y}" by (induct xs) (auto simp add: List.map_filter_def set_eq_iff) subsection \Quicksort by Relation\ text \A functional implementation of quicksort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun partition_rev :: "('a \ bool) \ ('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "partition_rev P (yes, no) [] = (yes, no)" | "partition_rev P (yes, no) (x # xs) = partition_rev P (if P x then (x # yes, no) else (yes, x # no)) xs" lemma partition_rev_filter_conv : "partition_rev P (yes, no) xs = (rev (filter P xs) @ yes, rev (filter (Not \ P) xs) @ no)" by (induct xs arbitrary: yes no) (simp_all) function quicksort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "quicksort_by_rel R sl [] = sl" | "quicksort_by_rel R sl (x#xs) = (let (xs_s, xs_b) = partition_rev (\y. R y x) ([],[]) xs in quicksort_by_rel R (x # (quicksort_by_rel R sl xs_b)) xs_s)" by pat_completeness simp_all termination by (relation "measure (\(_, _, xs). length xs)") (simp_all add: partition_rev_filter_conv less_Suc_eq_le) lemma quicksort_by_rel_remove_acc : "quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp1a = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2"] note ind_hyp1b = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2 @ sl"] note ind_hyp2 = ind_hyp[OF length_le(2), of sl] show ?thesis by (simp add: ind_hyp1a ind_hyp1b ind_hyp2) qed qed lemma quicksort_by_rel_remove_acc_guared : "sl \ [] \ quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" by (metis quicksort_by_rel_remove_acc) lemma quicksort_by_rel_permutes [simp]: "mset (quicksort_by_rel R sl xs) = mset (xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs'_multi_eq : "mset xs' = mset xs1 + mset xs2" unfolding partition_rev_filter_conv by (simp add: mset_filter) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: xs'_multi_eq union_assoc) qed qed lemma set_quicksort_by_rel [simp]: "set (quicksort_by_rel R sl xs) = set (xs @ sl)" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_quicksort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (quicksort_by_rel R [] xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs1_props: "\y. y \ set xs1 \ (R y x)" and xs2_props: "\y. y \ set xs2 \ \(R y x)" unfolding partition_rev_filter_conv by simp_all from xs2_props lin have xs2_props': "\y. y \ set xs2 \ (R x y)" by blast from xs2_props' xs1_props trans_R have xs1_props': "\y1 y2. y1 \ set xs1 \ y2 \ set xs2 \ (R y1 y2)" by metis from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyps = ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: quicksort_by_rel_remove_acc_guared sorted_wrt_append Ball_def xs1_props xs2_props' xs1_props') qed qed lemma sorted_quicksort_by_rel: "sorted (quicksort_by_rel (\) [] xs)" -unfolding sorted_sorted_wrt -by (rule sorted_wrt_quicksort_by_rel) auto + by (rule sorted_wrt_quicksort_by_rel) auto lemma sort_quicksort_by_rel: "sort = quicksort_by_rel (\) []" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_quicksort_by_rel) done lemma [code]: "quicksort = quicksort_by_rel (\) []" apply (subst sort_quicksort[symmetric]) by (rule sort_quicksort_by_rel) subsection \Mergesort by Relation\ text \A functional implementation of mergesort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun mergesort_by_rel_split :: "('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "mergesort_by_rel_split (xs1, xs2) [] = (xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) [x] = (x # xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) (x1 # x2 # xs) = mergesort_by_rel_split (x1 # xs1, x2 # xs2) xs" lemma list_induct_first2 [consumes 0, case_names Nil Sing Cons2]: assumes "P []" "\x. P [x]" "\x1 x2 xs. P xs \ P (x1 # x2 #xs)" shows "P xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis using assms(1) by simp next case (Cons x1 xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis using assms(2) by simp next case (Cons x2 xs'') note xs'_eq[simp] = this show ?thesis by (simp add: ind_hyp assms(3)) qed qed qed lemma mergesort_by_rel_split_length : "length (fst (mergesort_by_rel_split (xs1, xs2) xs)) = length xs1 + (length xs div 2) + (length xs mod 2) \ length (snd (mergesort_by_rel_split (xs1, xs2) xs)) = length xs2 + (length xs div 2)" by (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) (simp_all) lemma mset_mergesort_by_rel_split [simp]: "mset (fst (mergesort_by_rel_split (xs1, xs2) xs)) + mset (snd (mergesort_by_rel_split (xs1, xs2) xs)) = mset xs + mset xs1 + mset xs2" apply (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) apply (simp_all add: ac_simps) done fun mergesort_by_rel_merge :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" | "mergesort_by_rel_merge R xs [] = xs" | "mergesort_by_rel_merge R [] ys = ys" declare mergesort_by_rel_merge.simps [simp del] lemma mergesort_by_rel_merge_simps[simp] : "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" "mergesort_by_rel_merge R xs [] = xs" "mergesort_by_rel_merge R [] ys = ys" apply (simp_all add: mergesort_by_rel_merge.simps) apply (cases ys) apply (simp_all add: mergesort_by_rel_merge.simps) done lemma mergesort_by_rel_merge_induct [consumes 0, case_names Nil1 Nil2 Cons1 Cons2]: assumes "\xs::'a list. P xs []" "\ys::'b list. P [] ys" "\x xs y ys. R x y \ P xs (y # ys) \ P (x # xs) (y # ys)" "\x xs y ys. \(R x y) \ P (x # xs) ys \ P (x # xs) (y # ys)" shows "P xs ys" proof (induct xs arbitrary: ys) case Nil thus ?case using assms(2) by simp next case (Cons x xs) note P_xs = this show ?case proof (induct ys) case Nil thus ?case using assms(1) by simp next case (Cons y ys) note P_x_xs_ys = this show ?case using assms(3,4)[of x y xs ys] P_x_xs_ys P_xs by metis qed qed lemma mset_mergesort_by_rel_merge [simp]: "mset (mergesort_by_rel_merge R xs ys) = mset xs + mset ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) (simp_all add: ac_simps) lemma set_mergesort_by_rel_merge [simp]: "set (mergesort_by_rel_merge R xs ys) = set xs \ set ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) auto lemma sorted_wrt_mergesort_by_rel_merge [simp]: assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel_merge R xs ys) \ sorted_wrt R xs \ sorted_wrt R ys" proof (induct xs ys rule: mergesort_by_rel_merge_induct[where R = R]) case Nil1 thus ?case by simp next case Nil2 thus ?case by simp next case (Cons1 x xs y ys) thus ?case by (simp add: Ball_def) (metis trans_R) next case (Cons2 x xs y ys) thus ?case apply (auto simp add: Ball_def) apply (metis lin) apply (metis lin trans_R) done qed function mergesort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list" where "mergesort_by_rel R xs = (if length xs < 2 then xs else (mergesort_by_rel_merge R (mergesort_by_rel R (fst (mergesort_by_rel_split ([], []) xs))) (mergesort_by_rel R (snd (mergesort_by_rel_split ([], []) xs)))))" by auto termination by (relation "measure (\(_, xs). length xs)") (simp_all add: mergesort_by_rel_split_length not_less minus_div_mult_eq_mod [symmetric]) declare mergesort_by_rel.simps [simp del] lemma mergesort_by_rel_simps [simp, code] : "mergesort_by_rel R [] = []" "mergesort_by_rel R [x] = [x]" "mergesort_by_rel R (x1 # x2 # xs) = (let (xs1, xs2) = (mergesort_by_rel_split ([x1], [x2]) xs) in mergesort_by_rel_merge R (mergesort_by_rel R xs1) (mergesort_by_rel R xs2))" apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps[of _ "x1 # x2 # xs"] split: prod.splits) done lemma mergesort_by_rel_permutes [simp]: "mset (mergesort_by_rel R xs) = mset xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x1 xs') note xs_eq[simp] = this show ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: ac_simps) qed qed qed lemma set_mergesort_by_rel [simp]: "set (mergesort_by_rel R xs) = set xs" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_mergesort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel R xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: sorted_wrt_mergesort_by_rel_merge[OF lin trans_R]) qed qed qed lemma sorted_mergesort_by_rel: "sorted (mergesort_by_rel (\) xs)" -unfolding sorted_sorted_wrt by (rule sorted_wrt_mergesort_by_rel) auto lemma sort_mergesort_by_rel: "sort = mergesort_by_rel (\)" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_mergesort_by_rel) done definition "mergesort = mergesort_by_rel (\)" lemma sort_mergesort: "sort = mergesort" unfolding mergesort_def by (rule sort_mergesort_by_rel) subsubsection \Mergesort with Remdup\ term merge fun merge :: "'a::{linorder} list \ 'a list \ 'a list" where "merge [] l2 = l2" | "merge l1 [] = l1" | "merge (x1 # l1) (x2 # l2) = (if (x1 < x2) then x1 # (merge l1 (x2 # l2)) else (if (x1 = x2) then x1 # (merge l1 l2) else x2 # (merge (x1 # l1) l2)))" lemma merge_correct : assumes l1_OK: "distinct l1 \ sorted l1" assumes l2_OK: "distinct l2 \ sorted l2" shows "distinct (merge l1 l2) \ sorted (merge l1 l2) \ set (merge l1 l2) = set l1 \ set l2" using assms proof (induct l1 arbitrary: l2) case Nil thus ?case by simp next case (Cons x1 l1 l2) note x1_l1_props = Cons(2) note l2_props = Cons(3) from x1_l1_props have l1_props: "distinct l1 \ sorted l1" and x1_nin_l1: "x1 \ set l1" and x1_le: "\x. x \ set l1 \ x1 \ x" by (simp_all add: Ball_def) note ind_hyp_l1 = Cons(1)[OF l1_props] show ?case using l2_props proof (induct l2) case Nil with x1_l1_props show ?case by simp next case (Cons x2 l2) note x2_l2_props = Cons(2) from x2_l2_props have l2_props: "distinct l2 \ sorted l2" and x2_nin_l2: "x2 \ set l2" and x2_le: "\x. x \ set l2 \ x2 \ x" by (simp_all add: Ball_def) note ind_hyp_l2 = Cons(1)[OF l2_props] show ?case proof (cases "x1 < x2") case True note x1_less_x2 = this from ind_hyp_l1[OF x2_l2_props] x1_less_x2 x1_nin_l1 x1_le x2_le show ?thesis apply (auto simp add: Ball_def) apply (metis linorder_not_le) apply (metis linorder_not_less xt1(6) xt1(9)) done next case False note x2_le_x1 = this show ?thesis proof (cases "x1 = x2") case True note x1_eq_x2 = this from ind_hyp_l1[OF l2_props] x1_le x2_le x2_nin_l2 x1_eq_x2 x1_nin_l1 show ?thesis by (simp add: x1_eq_x2 Ball_def) next case False note x1_neq_x2 = this with x2_le_x1 have x2_less_x1 : "x2 < x1" by auto from ind_hyp_l2 x2_le_x1 x1_neq_x2 x2_le x2_nin_l2 x1_le show ?thesis apply (simp add: x2_less_x1 Ball_def) apply (metis linorder_not_le x2_less_x1 xt1(7)) done qed qed qed qed function merge_list :: "'a::{linorder} list list \ 'a list list \ 'a list" where "merge_list [] [] = []" | "merge_list [] [l] = l" | "merge_list (la # acc2) [] = merge_list [] (la # acc2)" | "merge_list (la # acc2) [l] = merge_list [] (l # la # acc2)" | "merge_list acc2 (l1 # l2 # ls) = merge_list ((merge l1 l2) # acc2) ls" by pat_completeness simp_all termination by (relation "measure (\(acc, ls). 3 * length acc + 2 * length ls)") (simp_all) lemma merge_list_correct : assumes ls_OK: "\l. l \ set ls \ distinct l \ sorted l" assumes as_OK: "\l. l \ set as \ distinct l \ sorted l" shows "distinct (merge_list as ls) \ sorted (merge_list as ls) \ set (merge_list as ls) = set (concat (as @ ls))" using assms proof (induct as ls rule: merge_list.induct) case 1 thus ?case by simp next case 2 thus ?case by simp next case 3 thus ?case by simp next case 4 thus ?case by auto next case (5 acc l1 l2 ls) note ind_hyp = 5(1) note l12_l_OK = 5(2) note acc_OK = 5(3) from l12_l_OK acc_OK merge_correct[of l1 l2] have set_merge_eq: "set (merge l1 l2) = set l1 \ set l2" by auto from l12_l_OK acc_OK merge_correct[of l1 l2] have "distinct (merge_list (merge l1 l2 # acc) ls) \ sorted (merge_list (merge l1 l2 # acc) ls) \ set (merge_list (merge l1 l2 # acc) ls) = set (concat ((merge l1 l2 # acc) @ ls))" by (rule_tac ind_hyp) auto with set_merge_eq show ?case by auto qed definition mergesort_remdups where "mergesort_remdups xs = merge_list [] (map (\x. [x]) xs)" lemma mergesort_remdups_correct : "distinct (mergesort_remdups l) \ sorted (mergesort_remdups l) \ (set (mergesort_remdups l) = set l)" proof - let ?l' = "map (\x. [x]) l" { fix xs assume "xs \ set ?l'" then obtain x where xs_eq: "xs = [x]" by auto hence "distinct xs \ sorted xs" by simp } note l'_OK = this from merge_list_correct[of "?l'" "[]", OF l'_OK] show ?thesis unfolding mergesort_remdups_def by simp qed (* TODO: Move *) lemma ex1_eqI: "\\!x. P x; P a; P b\ \ a=b" by blast lemma remdup_sort_mergesort_remdups: "remdups o sort = mergesort_remdups" (is "?lhs=?rhs") proof fix l have "set (?lhs l) = set l" and "sorted (?lhs l)" and "distinct (?lhs l)" by simp_all moreover note mergesort_remdups_correct ultimately show "?lhs l = ?rhs l" by (auto intro!: ex1_eqI[OF finite_sorted_distinct_unique[OF finite_set]]) qed subsection \Native Integers\ lemma int_of_integer_less_iff: "int_of_integer x < int_of_integer y \ x0 \ y\0 \ nat_of_integer x < nat_of_integer y \ xn' < n. \ P n') \ P (n::nat)" shows "\n' \ n. P n'" proof (rule classical) assume contra: "\ (\n'\n. P n')" hence "\n' < n. \ P n'" by auto hence "P n" by (rule hyp) thus "\n'\n. P n'" by auto qed subsubsection \Induction on nat\ lemma nat_compl_induct[case_names 0 Suc]: "\P 0; \n . \nn. nn \ n \ P nn \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nat_compl_induct'[case_names 0 Suc]: "\P 0; !! n . \!! nn . nn \ n \ P nn\ \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nz_le_conv_less: "0 k \ m \ k - Suc 0 < m" by auto lemma min_Suc_gt[simp]: "a min (Suc a) b = Suc a" "a min b (Suc a) = Suc a" by auto subsection \Integer\ text \Some setup from \int\ transferred to \integer\\ lemma atLeastLessThanPlusOne_atLeastAtMost_integer: "{l.. u \ {(0::integer).. u") apply (subst image_atLeastZeroLessThan_integer, assumption) apply (rule finite_imageI) apply auto done lemma finite_atLeastLessThan_integer [iff]: "finite {l..Functions of type @{typ "bool\bool"}\ lemma boolfun_cases_helper: "g=(\x. False) | g=(\x. x) | g=(\x. True) | g= (\x. \x)" proof - { assume "g False" "g True" hence "g = (\x. True)" by (rule_tac ext, case_tac x, auto) } moreover { assume "g False" "\g True" hence "g = (\x. \x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "g True" hence "g = (\x. x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "\g True" hence "g = (\x. False)" by (rule_tac ext, case_tac x, auto) } ultimately show ?thesis by fast qed lemma boolfun_cases[case_names False Id True Neg]: "\g=(\x. False) \ P; g=(\x. x) \ P; g=(\x. True) \ P; g=(\x. \x) \ P\ \ P" proof - note boolfun_cases_helper[of g] moreover assume "g=(\x. False) \ P" "g=(\x. x) \ P" "g=(\x. True) \ P" "g=(\x. \x) \ P" ultimately show ?thesis by fast qed subsection \Definite and indefinite description\ text "Combined definite and indefinite description for binary predicate" lemma some_theI: assumes EX: "\a b . P a b" and BUN: "!! b1 b2 . \\a . P a b1; \a . P a b2\ \ b1=b2" shows "P (SOME a . \b . P a b) (THE b . \a . P a b)" proof - from EX have "\b. P (SOME a. \b. P a b) b" by (rule someI_ex) moreover from EX have "\b. \a. P a b" by blast with BUN theI'[of "\b. \a. P a b"] have "\a. P a (THE b. \a. P a b)" by (unfold Ex1_def, blast) moreover note BUN ultimately show ?thesis by (fast) qed lemma some_insert_self[simp]: "S\{} \ insert (SOME x. x\S) S = S" by (auto intro: someI) lemma some_elem[simp]: "S\{} \ (SOME x. x\S) \ S" by (auto intro: someI) subsubsection\Hilbert Choice with option\ definition Eps_Opt where "Eps_Opt P = (if (\x. P x) then Some (SOME x. P x) else None)" lemma some_opt_eq_trivial[simp] : "Eps_Opt (\y. y = x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_sym_eq_trivial[simp] : "Eps_Opt ((=) x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_false_trivial[simp] : "Eps_Opt (\_. False) = None" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_None[simp] : "Eps_Opt P = None \ \(Ex P)" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_Some_implies : "Eps_Opt P = Some x \ P x" unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) lemma Eps_Opt_eq_Some : assumes P_prop: "\x'. P x \ P x' \ x' = x" shows "Eps_Opt P = Some x \ P x" using P_prop unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) subsection \Product Type\ lemma nested_case_prod_simp: "(\(a,b) c. f a b c) x y = (case_prod (\a b. f a b y) x)" by (auto split: prod.split) lemma fn_fst_conv: "(\x. (f (fst x))) = (\(a,_). f a)" by auto lemma fn_snd_conv: "(\x. (f (snd x))) = (\(_,b). f b)" by auto fun pairself where "pairself f (a,b) = (f a, f b)" lemma pairself_image_eq[simp]: "pairself f ` {(a,b). P a b} = {(f a, f b)| a b. P a b}" by force lemma pairself_image_cart[simp]: "pairself f ` (A\B) = f`A \ f`B" by (auto simp: image_def) lemma in_prod_fst_sndI: "fst x \ A \ snd x \ B \ x\A\B" by (cases x) auto lemma inj_Pair[simp]: "inj_on (\x. (x,c x)) S" "inj_on (\x. (c x,x)) S" by (auto intro!: inj_onI) declare Product_Type.swap_inj_on[simp] lemma img_fst [intro]: assumes "(a,b) \ S" shows "a \ fst ` S" by (rule image_eqI[OF _ assms]) simp lemma img_snd [intro]: assumes "(a,b) \ S" shows "b \ snd ` S" by (rule image_eqI[OF _ assms]) simp lemma range_prod: "range f \ (range (fst \ f)) \ (range (snd \ f))" proof fix y assume "y \ range f" then obtain x where y: "y = f x" by auto hence "y = (fst(f x), snd(f x))" by simp thus "y \ (range (fst \ f)) \ (range (snd \ f))" by (fastforce simp add: image_def) qed lemma finite_range_prod: assumes fst: "finite (range (fst \ f))" and snd: "finite (range (snd \ f))" shows "finite (range f)" proof - from fst snd have "finite (range (fst \ f) \ range (snd \ f))" by (rule finite_SigmaI) thus ?thesis by (rule finite_subset[OF range_prod]) qed lemma fstE: "x = (a,b) \ P (fst x) \ P a" by (metis fst_conv) lemma sndE: "x = (a,b) \ P (snd x) \ P b" by (metis snd_conv) subsubsection \Uncurrying\ (* TODO: Move to HOL/Product_Type? Lars H: "It's equal to case_prod, should use an abbreviation"*) definition uncurry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "uncurry f \ \(a,b). f a b" lemma uncurry_apply[simp]: "uncurry f (a,b) = f a b" unfolding uncurry_def by simp lemma curry_uncurry_id[simp]: "curry (uncurry f) = f" unfolding uncurry_def by simp lemma uncurry_curry_id[simp]: "uncurry (curry f) = f" unfolding uncurry_def by simp lemma do_curry: "f (a,b) = curry f a b" by simp lemma do_uncurry: "f a b = uncurry f (a,b)" by simp subsection \Sum Type\ lemma map_sum_Inr_conv: "map_sum fl fr s = Inr y \ (\x. s=Inr x \ y = fr x)" by (cases s) auto lemma map_sum_Inl_conv: "map_sum fl fr s = Inl y \ (\x. s=Inl x \ y = fl x)" by (cases s) auto subsection \Directed Graphs and Relations\ subsubsection "Reflexive-Transitive Closure" lemma r_le_rtrancl[simp]: "S\S\<^sup>*" by auto lemma rtrancl_mono_rightI: "S\S' \ S\S'\<^sup>*" by auto lemma trancl_sub: "R \ R\<^sup>+" by auto lemma trancl_single[simp]: "{(a,b)}\<^sup>+ = {(a,b)}" by (auto simp: trancl_insert) text \Pick first non-reflexive step\ lemma converse_rtranclE'[consumes 1, case_names base step]: assumes "(u,v)\R\<^sup>*" obtains "u=v" | vh where "u\vh" and "(u,vh)\R" and "(vh,v)\R\<^sup>*" using assms apply (induct rule: converse_rtrancl_induct) apply auto [] apply (case_tac "y=z") apply auto done lemma in_rtrancl_insert: "x\R\<^sup>* \ x\(insert r R)\<^sup>*" by (metis in_mono rtrancl_mono subset_insertI) lemma rtrancl_apply_insert: "R\<^sup>*``(insert x S) = insert x (R\<^sup>*``(S\R``{x}))" apply (auto) apply (erule converse_rtranclE) apply auto [2] apply (erule converse_rtranclE) apply (auto intro: converse_rtrancl_into_rtrancl) [2] done text \A path in a graph either does not use nodes from S at all, or it has a prefix leading to a node in S and a suffix that does not use nodes in S\ lemma rtrancl_last_visit[cases set, case_names no_visit last_visit_point]: shows "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>+; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" proof (induct rule: converse_rtrancl_induct[case_names refl step]) case refl thus ?case by auto next case (step q qh) show P proof (rule step.hyps(3)) assume A: "(qh,q')\(R-UNIV\S)\<^sup>*" show P proof (cases "qh\S") case False with step.hyps(1) A have "(q,q')\(R-UNIV\S)\<^sup>*" by (auto intro: converse_rtrancl_into_rtrancl) with step.prems(1) show P . next case True from step.hyps(1) have "(q,qh)\R\<^sup>+" by auto with step.prems(2) True A show P by blast qed next fix qt assume A: "qt\S" "(qh,qt)\R\<^sup>+" "(qt,q')\(R-UNIV\S)\<^sup>*" with step.hyps(1) have "(q,qt)\R\<^sup>+" by auto with step.prems(2) A(1,3) show P by blast qed qed text \Less general version of \rtrancl_last_visit\, but there's a short automatic proof\ lemma rtrancl_last_visit': "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (induct rule: converse_rtrancl_induct) (auto intro: converse_rtrancl_into_rtrancl) lemma rtrancl_last_visit_node: assumes "(s,s')\R\<^sup>*" shows "s\sh \ (s,s')\(R \ (UNIV \ (-{sh})))\<^sup>* \ (s,sh)\R\<^sup>* \ (sh,s')\(R \ (UNIV \ (-{sh})))\<^sup>*" using assms proof (induct rule: converse_rtrancl_induct) case base thus ?case by auto next case (step s st) moreover { assume P: "(st,s')\ (R \ UNIV \ - {sh})\<^sup>*" { assume "st=sh" with step have ?case by auto } moreover { assume "st\sh" with \(s,st)\R\ have "(s,st)\(R \ UNIV \ - {sh})\<^sup>*" by auto also note P finally have ?case by blast } ultimately have ?case by blast } moreover { assume P: "(st, sh) \ R\<^sup>* \ (sh, s') \ (R \ UNIV \ - {sh})\<^sup>*" with step(1) have ?case by (auto dest: converse_rtrancl_into_rtrancl) } ultimately show ?case by blast qed text \Find last point where a path touches a set\ lemma rtrancl_last_touch: "\ (q,q')\R\<^sup>*; q\S; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (erule rtrancl_last_visit') auto text \A path either goes over edge once, or not at all\ lemma trancl_over_edgeE: assumes "(u,w)\(insert (v1,v2) E)\<^sup>+" obtains "(u,w)\E\<^sup>+" | "(u,v1)\E\<^sup>*" and "(v2,w)\E\<^sup>*" using assms proof induct case (base z) thus ?thesis by (metis insertE prod.inject r_into_trancl' rtrancl_eq_or_trancl) next case (step y z) thus ?thesis by (metis (hide_lams, no_types) Pair_inject insertE rtrancl.simps trancl.simps trancl_into_rtrancl) qed lemma rtrancl_image_advance: "\q\R\<^sup>* `` Q0; (q,x)\R\ \ x\R\<^sup>* `` Q0" by (auto intro: rtrancl_into_rtrancl) lemma trancl_image_by_rtrancl: "(E\<^sup>+)``Vi \ Vi = (E\<^sup>*)``Vi" by (metis Image_Id Un_Image rtrancl_trancl_reflcl) lemma reachable_mono: "\R\R'; X\X'\ \ R\<^sup>*``X \ R'\<^sup>*``X'" by (metis Image_mono rtrancl_mono) lemma finite_reachable_advance: "\ finite (E\<^sup>*``{v0}); (v0,v)\E\<^sup>* \ \ finite (E\<^sup>*``{v})" by (erule finite_subset[rotated]) auto lemma rtrancl_mono_mp: "U\V \ x\U\<^sup>* \ x\V\<^sup>*" by (metis in_mono rtrancl_mono) lemma trancl_mono_mp: "U\V \ x\U\<^sup>+ \ x\V\<^sup>+" by (metis trancl_mono) lemma rtrancl_mapI: "(a,b)\E\<^sup>* \ (f a, f b)\(pairself f `E)\<^sup>*" apply (induction rule: rtrancl_induct) apply (force intro: rtrancl.intros)+ done lemma rtrancl_image_advance_rtrancl: assumes "q \ R\<^sup>*``Q0" assumes "(q,x) \ R\<^sup>*" shows "x \ R\<^sup>*``Q0" using assms by (metis rtrancl_idemp rtrancl_image_advance) lemma nth_step_trancl: "\n m. \ \ n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R \ \ n < length xs \ m < n \ (xs ! n, xs ! m) \ R\<^sup>+" proof (induction xs) case (Cons x xs) hence "\n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R" apply clarsimp by (metis One_nat_def diff_Suc_eq_diff_pred nth_Cons_Suc zero_less_diff) note IH = this[THEN Cons.IH] from Cons obtain n' where n': "Suc n' = n" by (cases n) blast+ show ?case proof (cases m) case "0" with Cons have "xs \ []" by auto with "0" Cons.prems(1)[of m] have "(xs ! 0, x) \ R" by simp moreover from IH[where m = 0] have "\n. n < length xs \ n > 0 \ (xs ! n, xs ! 0) \ R\<^sup>+" by simp ultimately have "\n. n < length xs \ (xs ! n, x) \ R\<^sup>+" by (metis trancl_into_trancl gr0I r_into_trancl') with Cons "0" show ?thesis by auto next case (Suc m') with Cons.prems n' have "n' < length xs" "m' < n'" by auto with IH have "(xs ! n', xs ! m') \ R\<^sup>+" by simp with Suc n' show ?thesis by auto qed qed simp lemma Image_empty_trancl_Image_empty: "R `` {v} = {} \ R\<^sup>+ `` {v} = {}" unfolding Image_def by (auto elim: converse_tranclE) lemma Image_empty_rtrancl_Image_id: "R `` {v} = {} \ R\<^sup>* `` {v} = {v}" unfolding Image_def by (auto elim: converse_rtranclE) lemma trans_rtrancl_eq_reflcl: "trans A \ A^* = A^=" by (simp add: rtrancl_trancl_reflcl) lemma refl_on_reflcl_Image: "refl_on B A \ C \ B \ A^= `` C = A `` C" by (auto simp add: Image_def dest: refl_onD) lemma Image_absorb_rtrancl: "\ trans A; refl_on B A; C \ B \ \ A^* `` C = A `` C" by (simp add: trans_rtrancl_eq_reflcl refl_on_reflcl_Image) lemma trancl_Image_unfold_left: "E\<^sup>+``S = E\<^sup>*``E``S" by (auto simp: trancl_unfold_left) lemma trancl_Image_unfold_right: "E\<^sup>+``S = E``E\<^sup>*``S" by (auto simp: trancl_unfold_right) lemma trancl_Image_advance_ss: "(u,v)\E \ E\<^sup>+``{v} \ E\<^sup>+``{u}" by auto lemma rtrancl_Image_advance_ss: "(u,v)\E \ E\<^sup>*``{v} \ E\<^sup>*``{u}" by auto (* FIXME: nicer name *) lemma trancl_union_outside: assumes "(v,w) \ (E\U)\<^sup>+" and "(v,w) \ E\<^sup>+" shows "\x y. (v,x) \ (E\U)\<^sup>* \ (x,y) \ U \ (y,w) \ (E\U)\<^sup>*" using assms proof (induction) case base thus ?case by auto next case (step w x) show ?case proof (cases "(v,w)\E\<^sup>+") case True from step have "(v,w)\(E\U)\<^sup>*" by simp moreover from True step have "(w,x) \ U" by (metis Un_iff trancl.simps) moreover have "(x,x) \ (E\U)\<^sup>*" by simp ultimately show ?thesis by blast next case False with step.IH obtain a b where "(v,a) \ (E\U)\<^sup>*" "(a,b) \ U" "(b,w) \ (E\U)\<^sup>*" by blast moreover with step have "(b,x) \ (E\U)\<^sup>*" by (metis rtrancl_into_rtrancl) ultimately show ?thesis by blast qed qed lemma trancl_restrict_reachable: assumes "(u,v) \ E\<^sup>+" assumes "E``S \ S" assumes "u\S" shows "(u,v) \ (E\S\S)\<^sup>+" using assms by (induction rule: converse_trancl_induct) (auto intro: trancl_into_trancl2) lemma rtrancl_image_unfold_right: "E``E\<^sup>*``V \ E\<^sup>*``V" by (auto intro: rtrancl_into_rtrancl) lemma trancl_Image_in_Range: "R\<^sup>+ `` V \ Range R" by (auto elim: trancl.induct) lemma rtrancl_Image_in_Field: "R\<^sup>* `` V \ Field R \ V" proof - from trancl_Image_in_Range have "R\<^sup>+ `` V \ Field R" unfolding Field_def by fast hence "R\<^sup>+ `` V \ V \ Field R \ V" by blast with trancl_image_by_rtrancl show ?thesis by metis qed lemma rtrancl_sub_insert_rtrancl: "R\<^sup>* \ (insert x R)\<^sup>*" by (auto elim: rtrancl.induct rtrancl_into_rtrancl) lemma trancl_sub_insert_trancl: "R\<^sup>+ \ (insert x R)\<^sup>+" by (auto elim: trancl.induct trancl_into_trancl) lemma Restr_rtrancl_mono: "(v,w) \ (Restr E U)\<^sup>* \ (v,w) \ E\<^sup>*" by (metis inf_le1 rtrancl_mono subsetCE) lemma Restr_trancl_mono: "(v,w) \ (Restr E U)\<^sup>+ \ (v,w) \ E\<^sup>+" by (metis inf_le1 trancl_mono) subsubsection "Converse Relation" lemmas converse_add_simps = converse_Times trancl_converse[symmetric] converse_Un converse_Int lemma dom_ran_disj_comp[simp]: "Domain R \ Range R = {} \ R O R = {}" by auto lemma below_Id_inv[simp]: "R\\Id \ R\Id" by (auto) subsubsection "Cyclicity" lemma cyclicE: "\\acyclic g; !!x. (x,x)\g\<^sup>+ \ P\ \ P" by (unfold acyclic_def) blast lemma acyclic_insert_cyclic: "\acyclic g; \acyclic (insert (x,y) g)\ \ (y,x)\g\<^sup>*" by (unfold acyclic_def) (auto simp add: trancl_insert) text \ This lemma makes a case distinction about a path in a graph where a couple of edges with the same endpoint have been inserted: If there is a path from a to b, then there's such a path in the original graph, or there's a path that uses an inserted edge only once. Originally, this lemma was used to reason about the graph of an updated acquisition history. Any path in this graph is either already contained in the original graph, or passes via an inserted edge. Because all the inserted edges point to the same target node, in the second case, the path can be short-circuited to use exactly one inserted edge. \ lemma trancl_multi_insert[cases set, case_names orig via]: "\ (a,b)\(r\X\{m})\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,x)\r\<^sup>*; (m,b)\r\<^sup>* \ \ P \ \ P" proof (induct arbitrary: P rule: trancl_induct) case (base b) thus ?case by auto next case (step b c) show ?case proof (rule step.hyps(3)) assume A: "(a,b)\r\<^sup>+" note step.hyps(2) moreover { assume "(b,c)\r" with A have "(a,c)\r\<^sup>+" by auto with step.prems have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto next fix x assume A: "x \ X" "(a, x) \ r\<^sup>*" "(m, b) \ r\<^sup>*" note step.hyps(2) moreover { assume "(b,c)\r" with A(3) have "(m,c)\r\<^sup>*" by auto with step.prems(2)[OF A(1,2)] have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto qed qed text \ Version of @{thm [source] trancl_multi_insert} for inserted edges with the same startpoint. \ lemma trancl_multi_insert2[cases set, case_names orig via]: "\(a,b)\(r\{m}\X)\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,m)\r\<^sup>*; (x,b)\r\<^sup>* \ \ P \ \ P" proof goal_cases case prems: 1 from prems(1) have "(b,a)\((r\{m}\X)\<^sup>+)\" by simp also have "((r\{m}\X)\<^sup>+)\ = (r\\X\{m})\<^sup>+" by (simp add: converse_add_simps) finally have "(b, a) \ (r\ \ X \ {m})\<^sup>+" . thus ?case by (auto elim!: trancl_multi_insert intro: prems(2,3) simp add: trancl_converse rtrancl_converse ) qed lemma cyclic_subset: "\ \ acyclic R; R \ S \ \ \ acyclic S" unfolding acyclic_def by (blast intro: trancl_mono) subsubsection \Wellfoundedness\ lemma wf_min: assumes A: "wf R" "R\{}" "!!m. m\Domain R - Range R \ P" shows P proof - have H: "!!x. wf R \ \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)" by (erule_tac wf_induct_rule[where P="\x. \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)"]) auto from A(2) obtain x y where "(x,y)\R" by auto with A(1,3) H show ?thesis by blast qed lemma finite_wf_eq_wf_converse: "finite R \ wf (R\) \ wf R" by (metis acyclic_converse finite_acyclic_wf finite_acyclic_wf_converse wf_acyclic) lemma wf_max: assumes A: "wf (R\)" "R\{}" and C: "!!m. m\Range R - Domain R \ P" shows "P" proof - from A(2) have NE: "R\\{}" by auto from wf_min[OF A(1) NE] obtain m where "m\Range R - Domain R" by auto thus P by (blast intro: C) qed \ \Useful lemma to show well-foundedness of some process approaching a finite upper bound\ lemma wf_bounded_supset: "finite S \ wf {(Q',Q). Q'\Q \ Q'\ S}" proof - assume [simp]: "finite S" hence [simp]: "!!x. finite (S-x)" by auto have "{(Q',Q). Q\Q' \ Q'\ S} \ inv_image ({(s'::nat,s). s'Q. card (S-Q))" proof (intro subsetI, case_tac x, simp) fix a b assume A: "b\a \ a\S" hence "S-a \ S-b" by blast thus "card (S-a) < card (S-b)" by (auto simp add: psubset_card_mono) qed moreover have "wf ({(s'::nat,s). s' Range R = {} \ wf R" apply (rule wf_no_loop) by simp text \Extend a wf-relation by a break-condition\ definition "brk_rel R \ {((False,x),(False,y)) | x y. (x,y)\R} \ {((True,x),(False,y)) | x y. True}" lemma brk_rel_wf[simp,intro!]: assumes WF[simp]: "wf R" shows "wf (brk_rel R)" proof - have "wf {((False,x),(False,y)) | x y. (x,y)\R}" proof - have "{((False,x),(False,y)) | x y. (x,y)\R} \ inv_image R snd" by auto from wf_subset[OF wf_inv_image[OF WF] this] show ?thesis . qed moreover have "wf {((True,x),(False,y)) | x y. True}" by (rule wf_no_path) auto ultimately show ?thesis unfolding brk_rel_def apply (subst Un_commute) by (blast intro: wf_Un) qed subsubsection \Restrict Relation\ definition rel_restrict :: "('a \ 'a) set \ 'a set \ ('a \ 'a) set" where "rel_restrict R A \ {(v,w). (v,w) \ R \ v \ A \ w \ A}" lemma rel_restrict_alt_def: "rel_restrict R A = R \ (-A) \ (-A)" unfolding rel_restrict_def by auto lemma rel_restrict_empty[simp]: "rel_restrict R {} = R" by (simp add: rel_restrict_def) lemma rel_restrict_notR: assumes "(x,y) \ rel_restrict A R" shows "x \ R" and "y \ R" using assms unfolding rel_restrict_def by auto lemma rel_restrict_sub: "rel_restrict R A \ R" unfolding rel_restrict_def by auto lemma rel_restrict_Int_empty: "A \ Field R = {} \ rel_restrict R A = R" unfolding rel_restrict_def Field_def by auto lemma Domain_rel_restrict: "Domain (rel_restrict R A) \ Domain R - A" unfolding rel_restrict_def by auto lemma Range_rel_restrict: "Range (rel_restrict R A) \ Range R - A" unfolding rel_restrict_def by auto lemma Field_rel_restrict: "Field (rel_restrict R A) \ Field R - A" unfolding rel_restrict_def Field_def by auto lemma rel_restrict_compl: "rel_restrict R A \ rel_restrict R (-A) = {}" unfolding rel_restrict_def by auto lemma finite_rel_restrict: "finite R \ finite (rel_restrict R A)" by (metis finite_subset rel_restrict_sub) lemma R_subset_Field: "R \ Field R \ Field R" unfolding Field_def by auto lemma homo_rel_restrict_mono: "R \ B \ B \ rel_restrict R A \ (B - A) \ (B - A)" proof - assume A: "R \ B \ B" hence "Field R \ B" unfolding Field_def by auto with Field_rel_restrict have "Field (rel_restrict R A) \ B - A" by (metis Diff_mono order_refl order_trans) with R_subset_Field show ?thesis by blast qed lemma rel_restrict_union: "rel_restrict R (A \ B) = rel_restrict (rel_restrict R A) B" unfolding rel_restrict_def by auto lemma rel_restrictI: "x \ R \ y \ R \ (x,y) \ E \ (x,y) \ rel_restrict E R" unfolding rel_restrict_def by auto lemma rel_restrict_lift: "(x,y) \ rel_restrict E R \ (x,y) \ E" unfolding rel_restrict_def by simp lemma rel_restrict_trancl_mem: "(a,b) \ (rel_restrict A R)\<^sup>+ \ (a,b) \ rel_restrict (A\<^sup>+) R" by (induction rule: trancl_induct) (auto simp add: rel_restrict_def) lemma rel_restrict_trancl_sub: "(rel_restrict A R)\<^sup>+ \ rel_restrict (A\<^sup>+) R" by (metis subrelI rel_restrict_trancl_mem) lemma rel_restrict_mono: "A \ B \ rel_restrict A R \ rel_restrict B R" unfolding rel_restrict_def by auto lemma rel_restrict_mono2: "R \ S \ rel_restrict A S \ rel_restrict A R" unfolding rel_restrict_def by auto lemma rel_restrict_Sigma_sub: "rel_restrict ((A\A)\<^sup>+) R \ ((A - R) \ (A - R))\<^sup>+" unfolding rel_restrict_def by auto (metis DiffI converse_tranclE mem_Sigma_iff r_into_trancl tranclE) lemma finite_reachable_restrictedI: assumes F: "finite Q" assumes I: "I\Q" assumes R: "Range E \ Q" shows "finite (E\<^sup>*``I)" proof - from I R have "E\<^sup>*``I \ Q" by (force elim: rtrancl.cases) also note F finally (finite_subset) show ?thesis . qed context begin private lemma rtrancl_restrictI_aux: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>* \ v\R" using assms by (induction) (auto simp: rel_restrict_def intro: rtrancl.intros) corollary rtrancl_restrictI: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>*" using rtrancl_restrictI_aux[OF assms] .. end lemma E_closed_restr_reach_cases: assumes P: "(u,v)\E\<^sup>*" assumes CL: "E``R \ R" obtains "v\R" | "u\R" "(u,v)\(rel_restrict E R)\<^sup>*" using P proof (cases rule: rtrancl_last_visit[where S=R]) case no_visit show ?thesis proof (cases "u\R") case True with P have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{u}"] by auto thus ?thesis .. next case False with no_visit have "(u,v)\(rel_restrict E R)\<^sup>*" by (rule rtrancl_restrictI) with False show ?thesis .. qed next case (last_visit_point x) from \(x, v) \ (E - UNIV \ R)\<^sup>*\ have "(x,v)\E\<^sup>*" by (rule rtrancl_mono_mp[rotated]) auto with \x\R\ have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{x}"] by auto thus ?thesis .. qed lemma rel_restrict_trancl_notR: assumes "(v,w) \ (rel_restrict E R)\<^sup>+" shows "v \ R" and "w \ R" using assms by (metis rel_restrict_trancl_mem rel_restrict_notR)+ lemma rel_restrict_tranclI: assumes "(x,y) \ E\<^sup>+" and "x \ R" "y \ R" and "E `` R \ R" shows "(x,y) \ (rel_restrict E R)\<^sup>+" using assms proof (induct) case base thus ?case by (metis r_into_trancl rel_restrictI) next case (step y z) hence "y \ R" by auto with step show ?case by (metis trancl_into_trancl rel_restrictI) qed subsubsection \Single-Valued Relations\ lemma single_valued_inter1: "single_valued R \ single_valued (R\S)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_inter2: "single_valued R \ single_valued (S\R)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_below_Id: "R\Id \ single_valued R" by (auto intro: single_valuedI) subsubsection \Bijective Relations\ definition "bijective R \ (\x y z. (x,y)\R \ (x,z)\R \ y=z) \ (\x y z. (x,z)\R \ (y,z)\R \ x=y)" lemma bijective_alt: "bijective R \ single_valued R \ single_valued (R\)" unfolding bijective_def single_valued_def by blast lemma bijective_Id[simp, intro!]: "bijective Id" by (auto simp: bijective_def) lemma bijective_Empty[simp, intro!]: "bijective {}" by (auto simp: bijective_def) subsubsection \Miscellaneous\ lemma pair_vimage_is_Image[simp]: "(Pair u -` E) = E``{u}" by auto lemma fst_in_Field: "fst ` R \ Field R" by (simp add: Field_def fst_eq_Domain) lemma snd_in_Field: "snd ` R \ Field R" by (simp add: Field_def snd_eq_Range) lemma ran_map_of: "ran (map_of xs) \ snd ` set (xs)" by (induct xs) (auto simp add: ran_def) lemma Image_subset_snd_image: "A `` B \ snd ` A" unfolding Image_def image_def by force lemma finite_Image_subset: "finite (A `` B) \ C \ A \ finite (C `` B)" by (metis Image_mono order_refl rev_finite_subset) lemma finite_Field_eq_finite[simp]: "finite (Field R) \ finite R" by (metis finite_cartesian_product finite_subset R_subset_Field finite_Field) definition "fun_of_rel R x \ SOME y. (x,y)\R" lemma for_in_RI(*[intro]*): "x\Domain R \ (x,fun_of_rel R x)\R" unfolding fun_of_rel_def by (auto intro: someI) lemma Field_not_elem: "v \ Field R \ \(x,y) \ R. x \ v \ y \ v" unfolding Field_def by auto lemma Sigma_UNIV_cancel[simp]: "(A \ X - A \ UNIV) = {}" by auto lemma same_fst_trancl[simp]: "(same_fst P R)\<^sup>+ = same_fst P (\x. (R x)\<^sup>+)" proof - { fix x y assume "(x,y)\(same_fst P R)\<^sup>+" hence "(x,y)\same_fst P (\x. (R x)\<^sup>+)" by induction (auto simp: same_fst_def) } moreover { fix f f' x y assume "((f,x),(f',y))\same_fst P (\x. (R x)\<^sup>+)" hence [simp]: "f'=f" "P f" and 1: "(x,y)\(R f)\<^sup>+" by (auto simp: same_fst_def) from 1 have "((f,x),(f',y))\(same_fst P R)\<^sup>+" apply induction subgoal by (rule r_into_trancl) auto subgoal by (erule trancl_into_trancl) auto done } ultimately show ?thesis by auto qed subsection \\option\ Datatype\ lemma le_some_optE: "\Some m\x; !!m'. \x=Some m'; m\m'\ \ P\ \ P" by (cases x) auto lemma not_Some_eq2[simp]: "(\x y. v \ Some (x,y)) = (v = None)" by (cases v) auto subsection "Maps" primrec the_default where "the_default _ (Some x) = x" | "the_default x None = x" declare map_add_dom_app_simps[simp] declare map_add_upd_left[simp] lemma ran_add[simp]: "dom f \ dom g = {} \ ran (f++g) = ran f \ ran g" by (fastforce simp add: ran_def map_add_def split: option.split_asm option.split) lemma nempty_dom: "\e\Map.empty; !!m. m\dom e \ P \ \ P" by (subgoal_tac "dom e \ {}") (blast, auto) lemma le_map_dom_mono: "m\m' \ dom m \ dom m'" apply (safe) apply (drule_tac x=x in le_funD) apply simp apply (erule le_some_optE) apply simp done lemma map_add_first_le: fixes m::"'a\('b::order)" shows "\ m\m' \ \ m++n \ m'++n" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split elim: le_funE) done lemma map_add_distinct_le: shows "\ m\m'; n\n'; dom m' \ dom n' = {} \ \ m++n \ m'++n'" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split) apply (fastforce elim: le_funE) apply (drule le_map_dom_mono) apply (drule le_map_dom_mono) apply (case_tac "m x") apply simp apply (force) apply (fastforce dest!: le_map_dom_mono) apply (erule le_funE) apply (erule_tac x=x in le_funE) apply simp done lemma map_add_left_comm: assumes A: "dom A \ dom B = {}" shows "A ++ (B ++ C) = B ++ (A ++ C)" proof - have "A ++ (B ++ C) = (A++B)++C" by simp also have "\ = (B++A)++C" by (simp add: map_add_comm[OF A]) also have "\ = B++(A++C)" by simp finally show ?thesis . qed lemmas map_add_ac = map_add_assoc map_add_comm map_add_left_comm lemma le_map_restrict[simp]: fixes m :: "'a \ ('b::order)" shows "m |` X \ m" by (rule le_funI) (simp add: restrict_map_def) lemma map_of_distinct_upd: "x \ set (map fst xs) \ [x \ y] ++ map_of xs = (map_of xs) (x \ y)" by (induct xs) (auto simp add: fun_upd_twist) lemma map_of_distinct_upd2: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd3: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd4: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)" using assms by (induct xs) (auto simp: map_of_eq_None_iff) lemma map_of_distinct_lookup: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) x = Some y" proof - have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x \ y)" using assms map_of_distinct_upd2 by simp thus ?thesis by simp qed lemma ran_distinct: assumes dist: "distinct (map fst al)" shows "ran (map_of al) = snd ` set al" using assms proof (induct al) case Nil then show ?case by simp next case (Cons kv al) then have "ran (map_of al) = snd ` set al" by simp moreover from Cons.prems have "map_of al (fst kv) = None" by (simp add: map_of_eq_None_iff) ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp qed lemma ran_is_image: "ran M = (the \ M) ` (dom M)" unfolding ran_def dom_def image_def by auto lemma map_card_eq_iff: assumes finite: "finite (dom M)" and card_eq: "card (dom M) = card (ran M)" and indom: "x \ dom M" shows "(M x = M y) \ (x = y)" proof - from ran_is_image finite card_eq have *: "inj_on (the \ M) (dom M)" using eq_card_imp_inj_on by metis thus ?thesis proof (cases "y \ dom M") case False with indom show ?thesis by auto next case True with indom have "the (M x) = the (M y) \ (x = y)" using inj_on_eq_iff[OF *] by auto thus ?thesis by auto qed qed lemma map_dom_ran_finite: "finite (dom M) \ finite (ran M)" by (simp add: ran_is_image) lemma map_update_eta_repair[simp]: (* An update operation may get simplified, if it happens to be eta-expanded. This lemma tries to repair some common expressions *) "dom (\x. if x=k then Some v else m x) = insert k (dom m)" "m k = None \ ran (\x. if x=k then Some v else m x) = insert v (ran m)" apply auto [] apply (force simp: ran_def) done lemma map_leI[intro?]: "\\x v. m1 x = Some v \ m2 x = Some v\ \ m1\\<^sub>mm2" unfolding map_le_def by force lemma map_leD: "m1\\<^sub>mm2 \ m1 k = Some v \ m2 k = Some v" unfolding map_le_def by force lemma map_restrict_insert_none_simp: "m x = None \ m|`(-insert x s) = m|`(-s)" by (auto intro!: ext simp:restrict_map_def) (* TODO: Move *) lemma eq_f_restr_conv: "s\dom (f A) \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" apply auto subgoal by (metis map_leI restrict_map_eq(2)) subgoal by (metis ComplD restrict_map_eq(2)) subgoal by (metis Compl_iff restrict_in) subgoal by (force simp: map_le_def restrict_map_def) done corollary eq_f_restr_ss_eq: "\ s\dom (f A) \ \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" using eq_f_restr_conv by blast subsubsection \Simultaneous Map Update\ definition "map_mmupd m K v k \ if k\K then Some v else m k" lemma map_mmupd_empty[simp]: "map_mmupd m {} v = m" by (auto simp: map_mmupd_def) lemma mmupd_in_upd[simp]: "k\K \ map_mmupd m K v k = Some v" by (auto simp: map_mmupd_def) lemma mmupd_notin_upd[simp]: "k\K \ map_mmupd m K v k = m k" by (auto simp: map_mmupd_def) lemma map_mmupdE: assumes "map_mmupd m K v k = Some x" obtains "k\K" "m k = Some x" | "k\K" "x=v" using assms by (auto simp: map_mmupd_def split: if_split_asm) lemma dom_mmupd[simp]: "dom (map_mmupd m K v) = dom m \ K" by (auto simp: map_mmupd_def split: if_split_asm) lemma le_map_mmupd_not_dom[simp, intro!]: "m \\<^sub>m map_mmupd m (K-dom m) v" by (auto simp: map_le_def) lemma map_mmupd_update_less: "K\K' \ map_mmupd m (K - dom m) v \\<^sub>m map_mmupd m (K'-dom m) v" by (auto simp: map_le_def map_mmupd_def) subsection\Connection between Maps and Sets of Key-Value Pairs\ definition map_to_set where "map_to_set m = {(k, v) . m k = Some v}" definition set_to_map where "set_to_map S k = Eps_Opt (\v. (k, v) \ S)" lemma set_to_map_simp : assumes inj_on_fst: "inj_on fst S" shows "(set_to_map S k = Some v) \ (k, v) \ S" proof (cases "\v. (k, v) \ S") case True note kv_ex = this then obtain v' where kv'_in: "(k, v') \ S" by blast with inj_on_fst have kv''_in: "\v''. (k, v'') \ S \ v' = v''" unfolding inj_on_def Ball_def by auto show ?thesis unfolding set_to_map_def by (simp add: kv_ex kv''_in) next case False hence kv''_nin: "\v''. (k, v'') \ S" by simp thus ?thesis by (simp add: set_to_map_def) qed lemma inj_on_fst_map_to_set : "inj_on fst (map_to_set m)" unfolding map_to_set_def inj_on_def by simp lemma map_to_set_inverse : "set_to_map (map_to_set m) = m" proof fix k show "set_to_map (map_to_set m) k = m k" proof (cases "m k") case None note mk_eq = this hence "\v. (k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k] show ?thesis unfolding mk_eq by auto next case (Some v) note mk_eq = this hence "(k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k v] show ?thesis unfolding mk_eq by auto qed qed lemma set_to_map_inverse : assumes inj_on_fst_S: "inj_on fst S" shows "map_to_set (set_to_map S) = S" proof (rule set_eqI) fix kv from set_to_map_simp [OF inj_on_fst_S, of "fst kv" "snd kv"] show "(kv \ map_to_set (set_to_map S)) = (kv \ S)" unfolding map_to_set_def by auto qed lemma map_to_set_empty[simp]: "map_to_set Map.empty = {}" unfolding map_to_set_def by simp lemma set_to_map_empty[simp]: "set_to_map {} = Map.empty" unfolding set_to_map_def[abs_def] by simp lemma map_to_set_empty_iff: "map_to_set m = {} \ m = Map.empty" "{} = map_to_set m \ m = Map.empty" unfolding map_to_set_def by auto lemma set_to_map_empty_iff: "set_to_map S = Map.empty \ S = {}" (is ?T1) "Map.empty = set_to_map S \ S = {}" (is ?T2) proof - show T1: ?T1 apply (simp only: set_eq_iff) apply (simp only: fun_eq_iff) apply (simp add: set_to_map_def) apply auto done from T1 show ?T2 by auto qed lemma map_to_set_upd[simp]: "map_to_set (m (k \ v)) = insert (k, v) (map_to_set m - {(k, v') |v'. True})" unfolding map_to_set_def apply (simp add: set_eq_iff) apply metis done lemma set_to_map_insert: assumes k_nin: "fst kv \ fst ` S" shows "set_to_map (insert kv S) = (set_to_map S) (fst kv \ snd kv)" proof fix k' obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from k_nin have k_nin': "\v'. (k, v') \ S" by (auto simp add: image_iff Ball_def) show "set_to_map (insert kv S) k' = (set_to_map S(fst kv \ snd kv)) k'" by (simp add: set_to_map_def k_nin') qed lemma map_to_set_dom : "dom m = fst ` (map_to_set m)" unfolding dom_def map_to_set_def by (auto simp add: image_iff) lemma map_to_set_ran : "ran m = snd ` (map_to_set m)" unfolding ran_def map_to_set_def by (auto simp add: image_iff) lemma set_to_map_dom : "dom (set_to_map S) = fst ` S" unfolding set_to_map_def[abs_def] dom_def by (auto simp add: image_iff Bex_def) lemma set_to_map_ran : "ran (set_to_map S) \ snd ` S" unfolding set_to_map_def[abs_def] ran_def subset_iff by (auto simp add: image_iff Bex_def) (metis Eps_Opt_eq_Some) lemma finite_map_to_set: "finite (map_to_set m) = finite (dom m)" unfolding map_to_set_def map_to_set_dom apply (intro iffI finite_imageI) apply assumption apply (rule finite_imageD[of fst]) apply assumption apply (simp add: inj_on_def) done lemma card_map_to_set : "card (map_to_set m) = card (dom m)" unfolding map_to_set_def map_to_set_dom apply (rule card_image[symmetric]) apply (simp add: inj_on_def) done lemma map_of_map_to_set : "distinct (map fst l) \ map_of l = m \ set l = map_to_set m" proof (induct l arbitrary: m) case Nil thus ?case by (simp add: map_to_set_empty_iff) blast next case (Cons kv l m) obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from Cons(2) have dist_l: "distinct (map fst l)" and kv'_nin: "\v'. (k, v') \ set l" by (auto simp add: image_iff) note ind_hyp = Cons(1)[OF dist_l] from kv'_nin have l_eq: "set (kv # l) = map_to_set m \ (set l = map_to_set (m (k := None))) \ m k = Some v" apply (simp add: map_to_set_def restrict_map_def set_eq_iff) apply (auto) apply (metis) apply (metis option.inject) done from kv'_nin have m_eq: "map_of (kv # l) = m \ map_of l = (m (k := None)) \ m k = Some v" apply (simp add: fun_eq_iff restrict_map_def map_of_eq_None_iff image_iff Ball_def) apply metis done show ?case unfolding m_eq l_eq using ind_hyp[of "m (k := None)"] by metis qed lemma map_to_set_map_of : "distinct (map fst l) \ map_to_set (map_of l) = set l" by (metis map_of_map_to_set) subsubsection \Mapping empty set to None\ definition "dflt_None_set S \ if S={} then None else Some S" lemma the_dflt_None_empty[simp]: "dflt_None_set {} = None" unfolding dflt_None_set_def by simp lemma the_dflt_None_nonempty[simp]: "S\{} \ dflt_None_set S = Some S" unfolding dflt_None_set_def by simp lemma the_dflt_None_set[simp]: "the_default {} (dflt_None_set x) = x" unfolding dflt_None_set_def by auto subsection \Orderings\ lemma (in order) min_arg_le[simp]: "n \ min m n \ min m n = n" "m \ min m n \ min m n = m" by (auto simp: min_def) lemma (in linorder) min_arg_not_ge[simp]: "\ min m n < m \ min m n = m" "\ min m n < n \ min m n = n" by (auto simp: min_def) lemma (in linorder) min_eq_arg[simp]: "min m n = m \ m\n" "min m n = n \ n\m" by (auto simp: min_def) lemma min_simps[simp]: "a<(b::'a::order) \ min a b = a" "b<(a::'a::order) \ min a b = b" by (auto simp add: min_def dest: less_imp_le) lemma (in -) min_less_self_conv[simp]: "min a b < a \ b < (a::_::linorder)" "min a b < b \ a < (b::_::linorder)" by (auto simp: min_def) lemma ord_eq_le_eq_trans: "\ a=b; b\c; c=d \ \ a\d" by auto lemma zero_comp_diff_simps[simp]: "(0::'a::linordered_idom) \ a - b \ b \ a" "(0::'a::linordered_idom) < a - b \ b < a" by auto subsubsection \Termination Measures\ text \Lexicographic measure, assuming upper bound for second component\ lemma mlex_fst_decrI: fixes a a' b b' N :: nat assumes "a a*N + N" using \b by linarith also have "\ \ a'*N" using \a by (metis Suc_leI ab_semigroup_add_class.add.commute ab_semigroup_mult_class.mult.commute mult_Suc_right mult_le_mono2) also have "\ \ a'*N + b'" by auto finally show ?thesis by auto qed lemma mlex_leI: fixes a a' b b' N :: nat assumes "a\a'" assumes "b\b'" shows "a*N + b \ a'*N + b'" using assms by (auto intro!: add_mono) lemma mlex_snd_decrI: fixes a a' b b' N :: nat assumes "a=a'" assumes "bCCPOs\ context ccpo begin lemma ccpo_Sup_mono: assumes C: "Complete_Partial_Order.chain (\) A" "Complete_Partial_Order.chain (\) B" assumes B: "\x\A. \y\B. x\y" shows "Sup A \ Sup B" proof (rule ccpo_Sup_least) fix x assume "x\A" with B obtain y where I: "y\B" and L: "x\y" by blast note L also from I ccpo_Sup_upper have "y\Sup B" by (blast intro: C) finally show "x\Sup B" . qed (rule C) lemma fixp_mono: assumes M: "monotone (\) (\) f" "monotone (\) (\) g" assumes LE: "\Z. f Z \ g Z" shows "ccpo_class.fixp f \ ccpo_class.fixp g" unfolding fixp_def[abs_def] apply (rule ccpo_Sup_mono) apply (rule chain_iterates M)+ proof rule fix x assume "x\ccpo_class.iterates f" thus "\y\ccpo_class.iterates g. x\y" proof (induct) case (step x) then obtain y where I: "y\ccpo_class.iterates g" and L: "x\y" by blast hence "g y \ ccpo_class.iterates g" and "f x \ g y" apply - apply (erule iterates.step) apply (rule order_trans) apply (erule monotoneD[OF M(1)]) apply (rule LE) done thus "\y\ccpo_class.iterates g. f x \ y" .. next case (Sup M) define N where "N = {SOME y. y\ccpo_class.iterates g \ x\y | x. x\M}" have N1: "\y\N. y\ccpo_class.iterates g \ (\x\M. x\y)" unfolding N_def apply auto apply (metis (lifting) Sup.hyps(2) tfl_some) by (metis (lifting) Sup.hyps(2) tfl_some) have N2: "\x\M. \y\N. x\y" unfolding N_def apply auto by (metis (lifting) Sup.hyps(2) tfl_some) have "N \ ccpo_class.iterates g" using Sup using N1 by auto hence C_chain: "Complete_Partial_Order.chain (\) N" using chain_iterates[OF M(2)] unfolding chain_def by auto have "Sup N \ ccpo_class.iterates g" and "Sup M \ Sup N" apply - apply (rule iterates.Sup[OF C_chain]) using N1 apply blast apply (rule ccpo_Sup_mono) apply (rule Sup.hyps) apply (rule C_chain) apply (rule N2) done thus ?case by blast qed qed end subsection \Code\ text \Constant for code-abort. If that gets executed, an abort-exception is raised.\ definition [simp]: "CODE_ABORT f = f ()" declare [[code abort: CODE_ABORT]] end diff --git a/thys/CakeML/generated/LemExtraDefs.thy b/thys/CakeML/generated/LemExtraDefs.thy --- a/thys/CakeML/generated/LemExtraDefs.thy +++ b/thys/CakeML/generated/LemExtraDefs.thy @@ -1,1257 +1,1258 @@ (*========================================================================*) (* Lem *) (* *) (* Dominic Mulligan, University of Cambridge *) (* Francesco Zappa Nardelli, INRIA Paris-Rocquencourt *) (* Gabriel Kerneis, University of Cambridge *) (* Kathy Gray, University of Cambridge *) (* Peter Boehm, University of Cambridge (while working on Lem) *) (* Peter Sewell, University of Cambridge *) (* Scott Owens, University of Kent *) (* Thomas Tuerk, University of Cambridge *) (* Brian Campbell, University of Edinburgh *) (* Shaked Flur, University of Cambridge *) (* Thomas Bauereiss, University of Cambridge *) (* Stephen Kell, University of Cambridge *) (* Thomas Williams *) (* Lars Hupel *) (* Basile Clement *) (* *) (* The Lem sources are copyright 2010-2018 *) (* by the authors above and Institut National de Recherche en *) (* Informatique et en Automatique (INRIA). *) (* *) (* All files except ocaml-lib/pmap.{ml,mli} and ocaml-libpset.{ml,mli} *) (* are distributed under the license below. The former are distributed *) (* under the LGPLv2, as in the LICENSE file. *) (* *) (* *) (* Redistribution and use in source and binary forms, with or without *) (* modification, are permitted provided that the following conditions *) (* are met: *) (* 1. Redistributions of source code must retain the above copyright *) (* notice, this list of conditions and the following disclaimer. *) (* 2. Redistributions in binary form must reproduce the above copyright *) (* notice, this list of conditions and the following disclaimer in the *) (* documentation and/or other materials provided with the distribution. *) (* 3. The names of the authors may not be used to endorse or promote *) (* products derived from this software without specific prior written *) (* permission. *) (* *) (* THIS SOFTWARE IS PROVIDED BY THE AUTHORS ``AS IS'' AND ANY EXPRESS *) (* OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED *) (* WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE *) (* ARE DISCLAIMED. IN NO EVENT SHALL THE AUTHORS BE LIABLE FOR ANY *) (* DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL *) (* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE *) (* GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS *) (* INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER *) (* IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR *) (* OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN *) (* IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) (*========================================================================*) chapter \Auxiliary Definitions needed by Lem\ theory "LemExtraDefs" imports Main "HOL-Combinatorics.List_Permutation" "HOL-Library.While_Combinator" begin hide_const (open) sign subsection \General\ consts failwith :: " 'a \ 'b" subsection \Lists\ fun index :: " 'a list \ nat \ 'a option " where "index [] n = None" | "index (x # xs) 0 = Some x" | "index (x # xs) (Suc n) = index xs n" lemma index_eq_some [simp]: "index l n = Some x \ (n < length l \ (x = l ! n))" proof (induct l arbitrary: n x) case Nil thus ?case by simp next case (Cons e es n x) note ind_hyp = this show ?case proof (cases n) case 0 thus ?thesis by auto next case (Suc n') with ind_hyp show ?thesis by simp qed qed lemma index_eq_none [simp]: "index l n = None \ length l \ n" by (rule iffD1[OF Not_eq_iff]) auto lemma index_simps [simp]: "length l \ n \ index l n = None" "n < length l \ index l n = Some (l ! n)" by (simp_all) fun find_indices :: "('a \ bool) \ 'a list \ nat list" where "find_indices P [] = []" | "find_indices P (x # xs) = (if P x then 0 # (map Suc (find_indices P xs)) else (map Suc (find_indices P xs)))" lemma length_find_indices : "length (find_indices P l) \ length l" by (induct l) auto lemma sorted_map_suc : "sorted l \ sorted (map Suc l)" by (induct l) (simp_all) lemma sorted_find_indices : "sorted (find_indices P xs)" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) from sorted_map_suc[OF this] show ?case by (simp) qed lemma find_indices_set [simp] : "set (find_indices P l) = {i. i < length l \ P (l ! i)}" proof (intro set_eqI) fix i show "i \ set (find_indices P l) \ (i \ {i. i < length l \ P (l ! i)})" proof (induct l arbitrary: i) case Nil thus ?case by simp next case (Cons x l' i) note ind_hyp = this show ?case proof (cases i) case 0 thus ?thesis by auto next case (Suc i') with ind_hyp[of i'] show ?thesis by auto qed qed qed definition find_index where "find_index P xs = (case find_indices P xs of [] \ None | i # _ \ Some i)" lemma find_index_eq_some [simp] : "(find_index P xs = Some ii) \ (ii < length xs \ P (xs ! ii) \ (\i' < ii. \(P (xs ! i'))))" (is "?lhs = ?rhs") proof (cases "find_indices P xs") case Nil with find_indices_set[of P xs] show ?thesis unfolding find_index_def by auto next case (Cons i il) note find_indices_eq = this from sorted_find_indices[of P xs] find_indices_eq have "sorted (i # il)" by simp hence i_leq: "\i'. i' \ set (i # il) \ i \ i'" by auto from find_indices_set[of P xs, unfolded find_indices_eq] have set_i_il_eq:"\i'. i' \ set (i # il) = (i' < length xs \ P (xs ! i'))" by simp have lhs_eq: "find_index P xs = Some i" unfolding find_index_def find_indices_eq by simp show ?thesis proof (intro iffI) assume ?lhs with lhs_eq have ii_eq[simp]: "ii = i" by simp from set_i_il_eq[of i] i_leq[unfolded set_i_il_eq] show ?rhs by auto (metis leD less_trans) next assume ?rhs with set_i_il_eq[of ii] have "ii \ set (i # il) \ (ii \ i)" by (metis leI length_pos_if_in_set nth_Cons_0 nth_mem set_i_il_eq) with i_leq [of ii] have "i = ii" by simp thus ?lhs unfolding lhs_eq by simp qed qed lemma find_index_eq_none [simp] : "(find_index P xs = None) \ (\x \ set xs. \(P x))" (is "?lhs = ?rhs") proof (rule iffD1[OF Not_eq_iff], intro iffI) assume "\ ?lhs" then obtain i where "find_index P xs = Some i" by auto hence "i < length xs \ P (xs ! i)" by simp thus "\ ?rhs" by auto next let ?p = "(\i. i < length xs \ P(xs ! i))" assume "\ ?rhs" then obtain i where "?p i" by (metis in_set_conv_nth) from LeastI [of ?p, OF \?p i\] have "?p (Least ?p)" . hence "find_index P xs = Some (Least ?p)" by (subst find_index_eq_some) (metis (mono_tags) less_trans not_less_Least) thus "\ ?lhs" by blast qed definition genlist where "genlist f n = map f (upt 0 n)" lemma genlist_length [simp] : "length (genlist f n) = n" unfolding genlist_def by simp lemma genlist_simps [simp]: "genlist f 0 = []" "genlist f (Suc n) = genlist f n @ [f n]" unfolding genlist_def by auto definition split_at where "split_at n l = (take n l, drop n l)" fun delete_first :: "('a \ bool) \ 'a list \ ('a list) option " where "delete_first P [] = None" | "delete_first P (x # xs) = (if (P x) then Some xs else map_option (\xs'. x # xs') (delete_first P xs))" declare delete_first.simps [simp del] lemma delete_first_simps [simp] : "delete_first P [] = None" "P x \ delete_first P (x # xs) = Some xs" "\(P x) \ delete_first P (x # xs) = map_option (\xs'. x # xs') (delete_first P xs)" unfolding delete_first.simps by auto lemmas delete_first_unroll = delete_first.simps(2) lemma delete_first_eq_none [simp] : "delete_first P l = None \ (\x \ set l. \ (P x))" by (induct l) (auto simp add: delete_first_unroll) lemma delete_first_eq_some : "delete_first P l = (Some l') \ (\l1 x l2. P x \ (\x \ set l1. \(P x)) \ (l = l1 @ (x # l2)) \ (l' = l1 @ l2))" (is "?lhs l l' = (\l1 x l2. ?rhs_body l1 x l2 l l')") proof (induct l arbitrary: l') case Nil thus ?case by simp next case (Cons e l l') note ind_hyp = this show ?case proof (cases "P e") case True show ?thesis proof (rule iffI) assume "?lhs (e # l) l'" with \P e\ have "l = l'" by simp with \P e\ have "?rhs_body [] e l' (e # l) l'" by simp thus "\l1 x l2. ?rhs_body l1 x l2 (e # l) l'" by blast next assume "\l1 x l2. ?rhs_body l1 x l2 (e # l) l'" then obtain l1 x l2 where body_ok: "?rhs_body l1 x l2 (e # l) l'" by blast from body_ok \P e\ have l1_eq[simp]: "l = l'" by (cases l1) (simp_all) with \P e\ show "?lhs (e # l) l'" by simp qed next case False define rhs_pred where "rhs_pred \ \l1 x l2 l l'. ?rhs_body l1 x l2 l l'" have rhs_fold: "\l1 x l2 l l'. ?rhs_body l1 x l2 l l' = rhs_pred l1 x l2 l l'" unfolding rhs_pred_def by simp have "(\z l1 x l2. rhs_pred l1 x l2 l z \ e # z = l') = (\l1 x l2. rhs_pred l1 x l2 (e # l) l')" proof (intro iffI) assume "\z l1 x l2. rhs_pred l1 x l2 l z \ e # z = l'" then obtain z l1 x l2 where "rhs_pred l1 x l2 l z" and l'_eq: "l' = e # z" by auto with \\(P e)\ have "rhs_pred (e # l1) x l2 (e # l) l'" unfolding rhs_pred_def by simp thus "\l1 x l2. rhs_pred l1 x l2 (e # l) l'" by blast next assume "\l1 x l2. rhs_pred l1 x l2 (e # l) l'" then obtain l1 x l2 where "rhs_pred l1 x l2 (e # l) l'" by blast with \\ (P e)\ obtain l1' where l1_eq[simp]: "l1 = e # l1'" unfolding rhs_pred_def by (cases l1) (auto) with \rhs_pred l1 x l2 (e # l) l'\ have "rhs_pred l1' x l2 l (l1' @ l2) \ e # (l1' @ l2) = l'" unfolding rhs_pred_def by (simp) thus "\z l1 x l2. rhs_pred l1 x l2 l z \ e # z = l'" by blast qed with \\ P e\ show ?thesis unfolding rhs_fold by (simp add: ind_hyp[unfolded rhs_fold]) qed qed lemma perm_eval [code] : "perm [] l \ l = []" (is ?g1) "perm (x # xs) l \ (case delete_first (\e. e = x) l of None => False | Some l' => perm xs l')" (is ?g2) proof - show ?g1 by auto next show ?g2 proof (cases "delete_first (\e. e = x) l") case None note del_eq = this hence "x \ set l" by auto with perm_set_eq [of "x # xs" l] have "\ perm (x # xs) l" by auto thus ?thesis unfolding del_eq by simp next case (Some l') note del_eq = this from del_eq[unfolded delete_first_eq_some] obtain l1 l2 where l_eq: "l = l1 @ [x] @ l2" and l'_eq: "l' = l1 @ l2" by auto have "(x # xs <~~> l1 @ x # l2) = (xs <~~> l1 @ l2)" proof - from perm_append_swap [of l1 "[x]"] perm_append2 [of "l1 @ [x]" "x # l1" l2] have "l1 @ x # l2 <~~> x # (l1 @ l2)" by simp hence "x # xs <~~> l1 @ x # l2 \ x # xs <~~> x # (l1 @ l2)" by (metis perm.trans perm_sym) thus ?thesis by simp qed with del_eq l_eq l'_eq show ?thesis by simp qed qed fun sorted_by :: "('a \ 'a \ bool)\ 'a list \ bool " where "sorted_by cmp [] = True" | "sorted_by cmp [_] = True" | "sorted_by cmp (x1 # x2 # xs) = ((cmp x1 x2) \ sorted_by cmp (x2 # xs))" lemma sorted_by_lesseq [simp] : "sorted_by ((\) :: ('a::{linorder}) => 'a => bool) = sorted" proof (rule ext) fix l :: "'a list" show "sorted_by (\) l = sorted l" proof (induct l) case Nil thus ?case by simp next case (Cons x xs) - thus ?case by (cases xs) (simp_all del: sorted.simps(2) add: sorted2_simps) + thus ?case + by (smt (verit, best) list.inject sorted1 sorted2 sorted_by.elims(1)) qed qed lemma sorted_by_cons_imp : "sorted_by cmp (x # xs) \ sorted_by cmp xs" by (cases xs) simp_all lemma sorted_by_cons_trans : assumes trans_cmp: "transp cmp" shows "sorted_by cmp (x # xs) = ((\x' \ set xs . cmp x x') \ sorted_by cmp xs)" proof (induct xs arbitrary: x) case Nil thus ?case by simp next case (Cons x2 xs x1) note ind_hyp = this from trans_cmp show ?case by (auto simp add: ind_hyp transp_def) qed fun insert_sort_insert_by :: "('a \ 'a \ bool)\ 'a \ 'a list \ 'a list " where "insert_sort_insert_by cmp e ([]) = ( [e])" | "insert_sort_insert_by cmp e (x # xs) = ( if cmp e x then (e # (x # xs)) else x # (insert_sort_insert_by cmp e xs))" lemma insert_sort_insert_by_length [simp] : "length (insert_sort_insert_by cmp e l) = Suc (length l)" by (induct l) auto lemma insert_sort_insert_by_set [simp] : "set (insert_sort_insert_by cmp e l) = insert e (set l)" by (induct l) auto lemma insert_sort_insert_by_perm : "(insert_sort_insert_by cmp e l) <~~> (e # l)" proof (induct l) case Nil thus ?case by simp next case (Cons e2 l') note ind_hyp = this have "e2 # e # l' <~~> e # e2 # l'" by (rule perm.swap) hence "e2 # insert_sort_insert_by cmp e l' <~~> e # e2 # l'" using ind_hyp by (metis cons_perm_eq perm.trans) thus ?case by simp qed lemma insert_sort_insert_by_sorted_by : assumes cmp_cases: "\y. y \ set l \ \ (cmp e y) \ cmp y e" assumes cmp_trans: "transp cmp" shows "sorted_by cmp l \ sorted_by cmp (insert_sort_insert_by cmp e l)" using cmp_cases proof (induct l) case Nil thus ?case by simp next case (Cons x1 l') note ind_hyp = Cons(1) note sorted_x1_l' = Cons(2) note cmp_cases = Cons(3) show ?case proof (cases l') case Nil with cmp_cases show ?thesis by simp next case (Cons x2 l'') note l'_eq = this from l'_eq sorted_x1_l' have "cmp x1 x2" "sorted_by cmp l'" by simp_all show ?thesis proof (cases "cmp e x1") case True with \cmp x1 x2\ \sorted_by cmp l'\ have "sorted_by cmp (x1 # l')" unfolding l'_eq by (simp) with \cmp e x1\ show ?thesis by simp next case False with cmp_cases have "cmp x1 e" by simp have "\x'. x' \ set l' \ cmp x1 x'" proof - fix x' assume "x' \ set l'" hence "x' = x2 \ cmp x2 x'" using \sorted_by cmp l'\ l'_eq sorted_by_cons_trans [OF cmp_trans, of x2 l''] by auto with transpD[OF cmp_trans, of x1 x2 x'] \cmp x1 x2\ show "cmp x1 x'" by blast qed hence "sorted_by cmp (x1 # insert_sort_insert_by cmp e l')" using ind_hyp [OF \sorted_by cmp l'\] \cmp x1 e\ cmp_cases unfolding sorted_by_cons_trans[OF cmp_trans] by simp with \\(cmp e x1)\ show ?thesis by simp qed qed qed fun insert_sort_by :: "('a \ 'a \ bool) \ 'a list \ 'a list " where "insert_sort_by cmp [] = []" | "insert_sort_by cmp (x # xs) = insert_sort_insert_by cmp x (insert_sort_by cmp xs)" lemma insert_sort_by_perm : "(insert_sort_by cmp l) <~~> l" proof (induct l) case Nil thus ?case by simp next case (Cons x l) thus ?case by simp (metis cons_perm_eq insert_sort_insert_by_perm perm.trans) qed lemma insert_sort_by_length [simp]: "length (insert_sort_by cmp l) = length l" by (induct l) auto lemma insert_sort_by_set [simp]: "set (insert_sort_by cmp l) = set l" by (induct l) auto definition sort_by where "sort_by = insert_sort_by" lemma sort_by_simps [simp]: "length (sort_by cmp l) = length l" "set (sort_by cmp l) = set l" unfolding sort_by_def by simp_all lemma sort_by_perm : "sort_by cmp l <~~> l" unfolding sort_by_def by (simp add: insert_sort_by_perm) subsection \Maps\ definition map_image :: "('v \ 'w) \ ('k, 'v) map \ ('k, 'w) map" where "map_image f m = (\k. map_option f (m k))" definition map_domain_image :: "('k \ 'v \ 'w) \ ('k, 'v) map \ ('k, 'w) map" where "map_domain_image f m = (\k. map_option (f k) (m k))" lemma map_image_simps [simp]: "(map_image f m) k = None \ m k = None" "(map_image f m) k = Some x \ (\x'. (m k = Some x') \ (x = f x'))" "(map_image f Map.empty) = Map.empty" "(map_image f (m (k \ v)) = (map_image f m) (k \ f v))" unfolding map_image_def by auto lemma map_image_dom_ran [simp]: "dom (map_image f m) = dom m" "ran (map_image f m) = f ` (ran m)" unfolding dom_def ran_def by auto definition map_to_set :: "('k, 'v) map \ ('k * 'v) set" where "map_to_set m = { (k, v) . m k = Some v }" lemma map_to_set_simps [simp] : "map_to_set Map.empty = {}" (is ?g1) "map_to_set (m ((k::'k) \ (v::'v))) = (insert (k, v) (map_to_set (m |` (- {k}))))" (is ?g2) proof - show ?g1 unfolding map_to_set_def by simp next show ?g2 proof (rule set_eqI) fix kv :: "('k * 'v)" obtain k' v' where kv_eq[simp]: "kv = (k', v')" by (rule prod.exhaust) show "(kv \ map_to_set (m(k \ v))) = (kv \ insert (k, v) (map_to_set (m |` (- {k}))))" by (auto simp add: map_to_set_def) qed qed subsection \Sets\ definition "set_choose s \ (SOME x. (x \ s))" definition without_trans_edges :: "('a \ 'a) set \ ('a \ 'a) set" where "without_trans_edges S \ let ts = trancl S in { (x, y) \ S. \z \ snd ` S. x \ z \ y \ z \ \ ((x, z) \ ts \ (z, y) \ ts)}" definition unbounded_lfp :: "'a set \ ('a set \ 'a set) \ 'a set" where "unbounded_lfp S f \ while (\x. x \ S) f S" definition unbounded_gfp :: "'a set \ ('a set \ 'a set) \ 'a set" where "unbounded_gfp S f \ while (\x. S \ x) f S" lemma set_choose_thm[simp]: "s \ {} \ (set_choose s) \ s" unfolding set_choose_def by (rule someI_ex) auto lemma set_choose_sing [simp]: "set_choose {x} = x" unfolding set_choose_def by auto lemma set_choose_code [code]: "set_choose (set [x]) = x" by auto lemma set_choose_in [intro] : assumes "s \ {}" shows "set_choose s \ s" proof - from \s \ {}\ obtain x where "x \ s" by auto thus ?thesis unfolding set_choose_def by (rule someI) qed definition set_case where "set_case s c_empty c_sing c_else = (if (s = {}) then c_empty else (if (card s = 1) then c_sing (set_choose s) else c_else))" lemma set_case_simps [simp] : "set_case {} c_empty c_sing c_else = c_empty" "set_case {x} c_empty c_sing c_else = c_sing x" "card s > 1 \ set_case s c_empty c_sing c_else = c_else" "\(finite s) \ set_case s c_empty c_sing c_else = c_else" unfolding set_case_def by auto lemma set_case_simp_insert2 [simp] : assumes x12_neq: "x1 \ x2" shows "set_case (insert x1 (insert x2 xs)) c_empty c_sing c_else = c_else" proof (cases "finite xs") case False thus ?thesis by (simp) next case True note fin_xs = this have "card {x1,x2} \ card (insert x1 (insert x2 xs))" by (rule card_mono) (auto simp add: fin_xs) with x12_neq have "1 < card (insert x1 (insert x2 xs))" by simp thus ?thesis by auto qed lemma set_case_code [code] : "set_case (set []) c_empty c_sing c_else = c_empty" "set_case (set [x]) c_empty c_sing c_else = c_sing x" "set_case (set (x1 # x2 # xs)) c_empty c_sing c_else = (if (x1 = x2) then set_case (set (x2 # xs)) c_empty c_sing c_else else c_else)" by auto definition set_lfp:: "'a set \ ('a set \ 'a set) \ 'a set" where "set_lfp s f = lfp (\s'. f s' \ s)" lemma set_lfp_tail_rec_def : assumes mono_f: "mono f" shows "set_lfp s f = (if (f s) \ s then s else (set_lfp (s \ f s) f))" (is "?ls = ?rs") proof (cases "f s \ s") case True note fs_sub_s = this from fs_sub_s have "\{u. f u \ s \ u} = s" by auto hence "?ls = s" unfolding set_lfp_def lfp_def . with fs_sub_s show "?ls = ?rs" by simp next case False note not_fs_sub_s = this from mono_f have mono_f': "mono (\s'. f s' \ s)" unfolding mono_def by auto have "\{u. f u \ s \ u} = \{u. f u \ (s \ f s) \ u}" (is "\?S1 = \?S2") proof have "?S2 \ ?S1" by auto thus "\?S1 \ \?S2" by (rule Inf_superset_mono) next { fix e assume "e \ \?S2" hence S2_prop: "\s'. f s' \ s' \ s \ s' \ f s \ s' \ e \ s'" by simp { fix s' assume "f s' \ s'" "s \ s'" from mono_f \s \ s'\ have "f s \ f s'" unfolding mono_def by simp with \f s' \ s'\ have "f s \ s'" by simp with \f s' \ s'\ \s \ s'\ S2_prop have "e \ s'" by simp } hence "e \ \?S1" by simp } thus "\?S2 \ \?S1" by auto qed hence "?ls = (set_lfp (s \ f s) f)" unfolding set_lfp_def lfp_def . with not_fs_sub_s show "?ls = ?rs" by simp qed lemma set_lfp_simps [simp] : "mono f \ f s \ s \ set_lfp s f = s" "mono f \ \(f s \ s) \ set_lfp s f = (set_lfp (s \ f s) f)" by (metis set_lfp_tail_rec_def)+ fun insert_in_list_at_arbitrary_pos where "insert_in_list_at_arbitrary_pos x [] = {[x]}" | "insert_in_list_at_arbitrary_pos x (y # ys) = insert (x # y # ys) ((\l. y # l) ` (insert_in_list_at_arbitrary_pos x ys))" lemma insert_in_list_at_arbitrary_pos_thm : "xl \ insert_in_list_at_arbitrary_pos x l \ (\l1 l2. l = l1 @ l2 \ xl = l1 @ [x] @ l2)" proof (induct l arbitrary: xl) case Nil thus ?case by simp next case (Cons y l xyl) note ind_hyp = this show ?case proof (rule iffI) assume xyl_in: "xyl \ insert_in_list_at_arbitrary_pos x (y # l)" show "\l1 l2. y # l = l1 @ l2 \ xyl = l1 @ [x] @ l2" proof (cases "xyl = x # y # l") case True hence "y # l = [] @ (y # l) \ xyl = [] @ [x] @ (y # l)" by simp thus ?thesis by blast next case False with xyl_in have "xyl \ (#) y ` insert_in_list_at_arbitrary_pos x l" by simp with ind_hyp obtain l1 l2 where "l = l1 @ l2 \ xyl = y # l1 @ x # l2" by (auto simp add: image_def Bex_def) hence "y # l = (y # l1) @ l2 \ xyl = (y # l1) @ [x] @ l2" by simp thus ?thesis by blast qed next assume "\l1 l2. y # l = l1 @ l2 \ xyl = l1 @ [x] @ l2" then obtain l1 l2 where yl_eq: "y # l = l1 @ l2" and xyl_eq: "xyl = l1 @ [x] @ l2" by blast show "xyl \ insert_in_list_at_arbitrary_pos x (y # l)" proof (cases l1) case Nil with yl_eq xyl_eq have "xyl = x # y # l" by simp thus ?thesis by simp next case (Cons y' l1') with yl_eq have l1_eq: "l1 = y # l1'" and l_eq: "l = l1' @ l2" by simp_all have "\l1'' l2''. l = l1'' @ l2'' \ l1' @ [x] @ l2 = l1'' @ [x] @ l2''" apply (rule_tac exI[where x = l1']) apply (rule_tac exI [where x = l2]) apply (simp add: l_eq) done hence "(l1' @ [x] @ l2) \ insert_in_list_at_arbitrary_pos x l" unfolding ind_hyp by blast hence "\l'. l' \ insert_in_list_at_arbitrary_pos x l \ l1 @ x # l2 = y # l'" by (rule_tac exI [where x = "l1' @ [x] @ l2"]) (simp add: l1_eq) thus ?thesis by (simp add: image_def Bex_def xyl_eq) qed qed qed definition list_of_set_set :: "'a set \ ('a list) set" where "list_of_set_set s = { l . (set l = s) \ distinct l }" lemma list_of_set_set_empty [simp]: "list_of_set_set {} = {[]}" unfolding list_of_set_set_def by auto lemma list_of_set_set_insert [simp] : "list_of_set_set (insert x s) = \ ((insert_in_list_at_arbitrary_pos x) ` (list_of_set_set (s - {x})))" (is "?lhs = ?rhs") proof (intro set_eqI) fix l have "(set l = insert x s \ distinct l) \ (\l1 l2. set (l1 @ l2) = s - {x} \ distinct (l1 @ l2) \ l = l1 @ x # l2)" proof (intro iffI) assume "set l = insert x s \ distinct l" hence set_l_eq: "set l = insert x s" and "distinct l" by simp_all from \set l = insert x s\ have "x \ set l" by simp then obtain l1 l2 where l_eq: "l = l1 @ x # l2" unfolding in_set_conv_decomp by blast from \distinct l\ l_eq have "distinct (l1 @ l2)" and x_nin: "x \ set (l1 @ l2)" by auto from x_nin set_l_eq[unfolded l_eq] have set_l12_eq: "set (l1 @ l2) = s - {x}" by auto from \distinct (l1 @ l2)\ l_eq set_l12_eq show "\l1 l2. set (l1 @ l2) = s - {x} \ distinct (l1 @ l2) \ l = l1 @ x # l2" by blast next assume "\l1 l2. set (l1 @ l2) = s - {x} \ distinct (l1 @ l2) \ l = l1 @ x # l2" then obtain l1 l2 where "set (l1 @ l2) = s - {x}" "distinct (l1 @ l2)" "l = l1 @ x # l2" by blast thus "set l = insert x s \ distinct l" by auto qed thus "l \ list_of_set_set (insert x s) \ l \ (\ ((insert_in_list_at_arbitrary_pos x) ` (list_of_set_set (s - {x}))))" unfolding list_of_set_set_def by (simp add: insert_in_list_at_arbitrary_pos_thm ex_simps[symmetric] del: ex_simps) qed lemma list_of_set_set_code [code]: "list_of_set_set (set []) = {[]}" "list_of_set_set (set (x # xs)) = \ ((insert_in_list_at_arbitrary_pos x) ` (list_of_set_set ((set xs) - {x})))" by simp_all lemma list_of_set_set_is_empty : "list_of_set_set s = {} \ \ (finite s)" proof - have "finite s \ (\l. set l = s \ distinct l)" proof (rule iffI) assume "\l. set l = s \ distinct l" then obtain l where "s = set l" by blast thus "finite s" by simp next assume "finite s" thus "\l. set l = s \ distinct l" proof (induct s) case empty show ?case by auto next case (insert e s) note e_nin_s = insert(2) from insert(3) obtain l where set_l: "set l = s" and dist_l: "distinct l" by blast from set_l have set_el: "set (e # l) = insert e s" by auto from dist_l set_l e_nin_s have dist_el: "distinct (e # l)" by simp from set_el dist_el show ?case by blast qed qed thus ?thesis unfolding list_of_set_set_def by simp qed definition list_of_set :: "'a set \ 'a list" where "list_of_set s = set_choose (list_of_set_set s)" lemma list_of_set [simp] : assumes fin_s: "finite s" shows "set (list_of_set s) = s" "distinct (list_of_set s)" proof - from fin_s list_of_set_set_is_empty[of s] have "\ (list_of_set_set s = {})" by simp hence "list_of_set s \ list_of_set_set s" unfolding list_of_set_def by (rule set_choose_thm) thus "set (list_of_set s) = s" "distinct (list_of_set s)" unfolding list_of_set_set_def by simp_all qed lemma list_of_set_in: "finite s \ list_of_set s \ list_of_set_set s" unfolding list_of_set_def by (metis list_of_set_set_is_empty set_choose_thm) definition ordered_list_of_set where "ordered_list_of_set cmp s = set_choose (sort_by cmp ` list_of_set_set s)" subsection \sum\ find_consts "'a list => ('a list * _)" fun sum_partition :: "('a + 'b) list \ 'a list * 'b list" where "sum_partition [] = ([], [])" | "sum_partition ((Inl l) # lrs) = (let (ll, rl) = sum_partition lrs in (l # ll, rl))" | "sum_partition ((Inr r) # lrs) = (let (ll, rl) = sum_partition lrs in (ll, r # rl))" lemma sum_partition_length : "List.length lrs = List.length (fst (sum_partition lrs)) + List.length (snd (sum_partition lrs))" proof (induct lrs) case Nil thus ?case by simp next case (Cons lr lrs) thus ?case by (cases lr) (auto split: prod.split) qed subsection \sorting\ subsection \Strings\ declare String.literal.explode_inverse [simp] subsection \num to string conversions\ definition nat_list_to_string :: "nat list \ string" where "nat_list_to_string nl = map char_of nl" definition is_digit where "is_digit (n::nat) = (n < 10)" lemma is_digit_simps[simp] : "n < 10 \ is_digit n" "\(n < 10) \ \(is_digit n)" unfolding is_digit_def by simp_all lemma is_digit_expand : "is_digit n \ (n = 0) \ (n = 1) \ (n = 2) \ (n = 3) \ (n = 4) \ (n = 5) \ (n = 6) \ (n = 7) \ (n = 8) \ (n = 9)" unfolding is_digit_def by auto lemmas is_digitE = is_digit_expand[THEN iffD1,elim_format] lemmas is_digitI = is_digit_expand[THEN iffD2,rule_format] definition is_digit_char where "is_digit_char c \ (c = CHR ''0'') \ (c = CHR ''5'') \ (c = CHR ''1'') \ (c = CHR ''6'') \ (c = CHR ''2'') \ (c = CHR ''7'') \ (c = CHR ''3'') \ (c = CHR ''8'') \ (c = CHR ''4'') \ (c = CHR ''9'')" lemma is_digit_char_simps[simp] : "is_digit_char (CHR ''0'')" "is_digit_char (CHR ''1'')" "is_digit_char (CHR ''2'')" "is_digit_char (CHR ''3'')" "is_digit_char (CHR ''4'')" "is_digit_char (CHR ''5'')" "is_digit_char (CHR ''6'')" "is_digit_char (CHR ''7'')" "is_digit_char (CHR ''8'')" "is_digit_char (CHR ''9'')" unfolding is_digit_char_def by simp_all lemmas is_digit_charE = is_digit_char_def[THEN iffD1,elim_format] lemmas is_digit_charI = is_digit_char_def[THEN iffD2,rule_format] definition digit_to_char :: "nat \ char" where "digit_to_char n = ( if n = 0 then CHR ''0'' else if n = 1 then CHR ''1'' else if n = 2 then CHR ''2'' else if n = 3 then CHR ''3'' else if n = 4 then CHR ''4'' else if n = 5 then CHR ''5'' else if n = 6 then CHR ''6'' else if n = 7 then CHR ''7'' else if n = 8 then CHR ''8'' else if n = 9 then CHR ''9'' else CHR ''X'')" lemma digit_to_char_simps [simp]: "digit_to_char 0 = CHR ''0''" "digit_to_char (Suc 0) = CHR ''1''" "digit_to_char 2 = CHR ''2''" "digit_to_char 3 = CHR ''3''" "digit_to_char 4 = CHR ''4''" "digit_to_char 5 = CHR ''5''" "digit_to_char 6 = CHR ''6''" "digit_to_char 7 = CHR ''7''" "digit_to_char 8 = CHR ''8''" "digit_to_char 9 = CHR ''9''" "n > 9 \ digit_to_char n = CHR ''X''" unfolding digit_to_char_def by simp_all definition char_to_digit :: "char \ nat" where "char_to_digit c = ( if c = CHR ''0'' then 0 else if c = CHR ''1'' then 1 else if c = CHR ''2'' then 2 else if c = CHR ''3'' then 3 else if c = CHR ''4'' then 4 else if c = CHR ''5'' then 5 else if c = CHR ''6'' then 6 else if c = CHR ''7'' then 7 else if c = CHR ''8'' then 8 else if c = CHR ''9'' then 9 else 10)" lemma char_to_digit_simps [simp]: "char_to_digit (CHR ''0'') = 0" "char_to_digit (CHR ''1'') = 1" "char_to_digit (CHR ''2'') = 2" "char_to_digit (CHR ''3'') = 3" "char_to_digit (CHR ''4'') = 4" "char_to_digit (CHR ''5'') = 5" "char_to_digit (CHR ''6'') = 6" "char_to_digit (CHR ''7'') = 7" "char_to_digit (CHR ''8'') = 8" "char_to_digit (CHR ''9'') = 9" unfolding char_to_digit_def by simp_all lemma diget_to_char_inv[simp]: assumes is_digit: "is_digit n" shows "char_to_digit (digit_to_char n) = n" using is_digit unfolding is_digit_expand by auto lemma char_to_diget_inv[simp]: assumes is_digit: "is_digit_char c" shows "digit_to_char (char_to_digit c) = c" using is_digit unfolding char_to_digit_def is_digit_char_def by auto lemma char_to_digit_div_mod [simp]: assumes is_digit: "is_digit_char c" shows "char_to_digit c < 10" using is_digit unfolding char_to_digit_def is_digit_char_def by auto lemma is_digit_char_intro[simp]: "is_digit (char_to_digit c) = is_digit_char c" unfolding char_to_digit_def is_digit_char_def is_digit_expand by auto lemma is_digit_intro[simp]: "is_digit_char (digit_to_char n) = is_digit n" unfolding digit_to_char_def is_digit_char_def is_digit_expand by auto lemma digit_to_char_11: "digit_to_char n1 = digit_to_char n2 \ (is_digit n1 = is_digit n2) \ (is_digit n1 \ (n1 = n2))" by (metis diget_to_char_inv is_digit_intro) lemma char_to_digit_11: "char_to_digit c1 = char_to_digit c2 \ (is_digit_char c1 = is_digit_char c2) \ (is_digit_char c1 \ (c1 = c2))" by (metis char_to_diget_inv is_digit_char_intro) function nat_to_string :: "nat \ string" where "nat_to_string n = (if (is_digit n) then [digit_to_char n] else nat_to_string (n div 10) @ [digit_to_char (n mod 10)])" by auto termination by (relation "measure id") (auto simp add: is_digit_def) definition int_to_string :: "int \ string" where "int_to_string i \ if i < 0 then ''-'' @ nat_to_string (nat (abs i)) else nat_to_string (nat i)" lemma nat_to_string_simps[simp]: "is_digit n \ nat_to_string n = [digit_to_char n]" "\(is_digit n) \ nat_to_string n = nat_to_string (n div 10) @ [digit_to_char (n mod 10)]" by simp_all declare nat_to_string.simps[simp del] lemma nat_to_string_neq_nil[simp]: "nat_to_string n \ []" by (cases "is_digit n") simp_all lemmas nat_to_string_neq_nil2[simp] = nat_to_string_neq_nil[symmetric] lemma nat_to_string_char_to_digit [simp]: "is_digit_char c \ nat_to_string (char_to_digit c) = [c]" by auto lemma nat_to_string_11[simp] : "(nat_to_string n1 = nat_to_string n2) \ n1 = n2" proof (rule iffI) assume "n1 = n2" thus "nat_to_string n1 = nat_to_string n2" by simp next assume "nat_to_string n1 = nat_to_string n2" thus "n1 = n2" proof (induct n2 arbitrary: n1 rule: less_induct) case (less n2') note ind_hyp = this(1) note n2s_eq = less(2) have is_dig_eq: "is_digit n2' = is_digit n1" using n2s_eq apply (cases "is_digit n2'") apply (case_tac [!] "is_digit n1") apply (simp_all) done show ?case proof (cases "is_digit n2'") case True with n2s_eq is_dig_eq show ?thesis by simp (metis digit_to_char_11) next case False with is_dig_eq have not_digs : "\ (is_digit n1)" "\ (is_digit n2')" by simp_all from not_digs(2) have "n2' div 10 < n2'" unfolding is_digit_def by auto note ind_hyp' = ind_hyp [OF this, of "n1 div 10"] from not_digs n2s_eq ind_hyp' digit_to_char_11[of "n1 mod 10" "n2' mod 10"] have "(n1 mod 10) = (n2' mod 10)" "n1 div 10 = n2' div 10" by simp_all thus "n1 = n2'" by (metis div_mult_mod_eq) qed qed qed definition "is_nat_string s \ (\c\set s. is_digit_char c)" definition "is_strong_nat_string s \ is_nat_string s \ (s \ []) \ (hd s = CHR ''0'' \ length s = 1)" lemma is_nat_string_simps[simp]: "is_nat_string []" "is_nat_string (c # s) \ is_digit_char c \ is_nat_string s" unfolding is_nat_string_def by simp_all lemma is_strong_nat_string_simps[simp]: "\(is_strong_nat_string [])" "is_strong_nat_string (c # s) \ is_digit_char c \ is_nat_string s \ (c = CHR ''0'' \ s = [])" unfolding is_strong_nat_string_def by simp_all fun string_to_nat_aux :: "nat \ string \ nat" where "string_to_nat_aux n [] = n" | "string_to_nat_aux n (d#ds) = string_to_nat_aux (n*10 + char_to_digit d) ds" definition string_to_nat :: "string \ nat option" where "string_to_nat s \ (if is_nat_string s then Some (string_to_nat_aux 0 s) else None)" definition string_to_nat' :: "string \ nat" where "string_to_nat' s \ the (string_to_nat s)" lemma string_to_nat_aux_inv : assumes "is_nat_string s" assumes "n > 0 \ is_strong_nat_string s" shows "nat_to_string (string_to_nat_aux n s) = (if n = 0 then '''' else nat_to_string n) @ s" using assms proof (induct s arbitrary: n) case Nil thus ?case by (simp add: is_strong_nat_string_def) next case (Cons c s n) from Cons(2) have "is_digit_char c" "is_nat_string s" by simp_all note cs_ok = Cons(3) let ?m = "n*10 + char_to_digit c" note ind_hyp = Cons(1)[OF \is_nat_string s\, of ?m] from \is_digit_char c\ have m_div: "?m div 10 = n" and m_mod: "?m mod 10 = char_to_digit c" unfolding is_digit_char_def by auto show ?case proof (cases "?m = 0") case True with \is_digit_char c\ have "n = 0" "c = CHR ''0''" unfolding is_digit_char_def by auto moreover with cs_ok have "s = []" by simp ultimately show ?thesis by simp next case False note m_neq_0 = this with ind_hyp have ind_hyp': "nat_to_string (string_to_nat_aux ?m s) = (nat_to_string ?m) @ s" by auto hence "nat_to_string (string_to_nat_aux n (c # s)) = (nat_to_string ?m) @ s" by simp with \is_digit_char c\ m_div show ?thesis by simp qed qed lemma string_to_nat_inv : assumes strong_nat_s: "is_strong_nat_string s" assumes s2n_s: "string_to_nat s = Some n" shows "nat_to_string n = s" proof - from strong_nat_s have nat_s: "is_nat_string s" unfolding is_strong_nat_string_def by simp with s2n_s have n_eq: "n = string_to_nat_aux 0 s" unfolding string_to_nat_def by simp from string_to_nat_aux_inv[of s 0, folded n_eq] nat_s strong_nat_s show ?thesis by simp qed lemma nat_to_string_induct [case_names "digit" "non_digit"]: assumes digit: "\d. is_digit d \ P d" assumes not_digit: "\n. \(is_digit n) \ P (n div 10) \ P (n mod 10) \ P n" shows "P n" proof (induct n rule: less_induct) case (less n) note ind_hyp = this show ?case proof (cases "is_digit n") case True with digit show ?thesis by simp next case False note not_dig = this hence "n div 10 < n" "n mod 10 < n" unfolding is_digit_def by auto with not_dig ind_hyp not_digit show ?thesis by simp qed qed lemma nat_to_string___is_nat_string [simp]: "is_nat_string (nat_to_string n)" unfolding is_nat_string_def proof (induct n rule: nat_to_string_induct) case (digit d) thus ?case by simp next case (non_digit n) thus ?case by simp qed lemma nat_to_string___eq_0 [simp]: "(nat_to_string n = (CHR ''0'')#s) \ (n = 0 \ s = [])" unfolding is_nat_string_def proof (induct n arbitrary: s rule: nat_to_string_induct) case (digit d) thus ?case unfolding is_digit_expand by auto next case (non_digit n) obtain c s' where ns_eq: "nat_to_string (n div 10) = c # s'" by (cases "nat_to_string (n div 10)") auto from non_digit(1) have "n div 10 \ 0" unfolding is_digit_def by auto with non_digit(2) ns_eq have c_neq: "c \ CHR ''0''" by auto from \\ (is_digit n)\ c_neq ns_eq show ?case by auto qed lemma nat_to_string___is_strong_nat_string: "is_strong_nat_string (nat_to_string n)" proof (cases "is_digit n") case True thus ?thesis by simp next case False note not_digit = this obtain c s' where ns_eq: "nat_to_string n = c # s'" by (cases "nat_to_string n") auto from not_digit have "0 < n" unfolding is_digit_def by simp with ns_eq have c_neq: "c \ CHR ''0''" by auto hence "hd (nat_to_string n) \ CHR ''0''" unfolding ns_eq by simp thus ?thesis unfolding is_strong_nat_string_def by simp qed lemma nat_to_string_inv : "string_to_nat (nat_to_string n) = Some n" by (metis nat_to_string_11 nat_to_string___is_nat_string nat_to_string___is_strong_nat_string string_to_nat_def string_to_nat_inv) definition The_opt where "The_opt p = (if (\!x. p x) then Some (The p) else None)" lemma The_opt_eq_some [simp] : "((The_opt p) = (Some x)) \ ((p x) \ ((\ y. p y \ (x = y))))" (is "?lhs = ?rhs") proof (cases "\!x. p x") case True note exists_unique = this then obtain x where p_x: "p x" by auto from the1_equality[of p x] exists_unique p_x have the_opt_eq: "The_opt p = Some x" unfolding The_opt_def by simp from exists_unique the_opt_eq p_x show ?thesis by auto next case False note not_unique = this hence "The_opt p = None" unfolding The_opt_def by simp with not_unique show ?thesis by auto qed lemma The_opt_eq_none [simp] : "((The_opt p) = None) \ \(\!x. p x)" unfolding The_opt_def by auto end diff --git a/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy b/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy --- a/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy +++ b/thys/Collections/ICF/gen_algo/PrioUniqueByAnnotatedList.thy @@ -1,1114 +1,1114 @@ section \\isaheader{Implementing Unique Priority Queues by Annotated Lists}\ theory PrioUniqueByAnnotatedList imports "../spec/AnnotatedListSpec" "../spec/PrioUniqueSpec" begin text \ In this theory we use annotated lists to implement unique priority queues with totally ordered elements. This theory is written as a generic adapter from the AnnotatedList interface to the unique priority queue interface. The annotated list stores a sequence of elements annotated with priorities\footnote{Technically, the annotated list elements are of unit-type, and the annotations hold both, the priority queue elements and the priorities. This is required as we defined annotated lists to only sum up the elements annotations.} The monoids operations forms the maximum over the elements and the minimum over the priorities. The sequence of pairs is ordered by ascending elements' order. The insertion point for a new element, or the priority of an existing element can be found by splitting the sequence at the point where the maximum of the elements read so far gets bigger than the element to be inserted. The minimum priority can be read out as the sum over the whole sequence. Finding the element with minimum priority is done by splitting the sequence at the point where the minimum priority of the elements read so far becomes equal to the minimum priority of the whole sequence. \ subsection "Definitions" subsubsection "Monoid" datatype ('e, 'a) LP = Infty | LP 'e 'a fun p_unwrap :: "('e,'a) LP \ ('e \ 'a)" where "p_unwrap (LP e a) = (e , a)" fun p_min :: "('e::linorder, 'a::linorder) LP \ ('e, 'a) LP \ ('e, 'a) LP" where "p_min Infty Infty = Infty"| "p_min Infty (LP e a) = LP e a"| "p_min (LP e a) Infty = LP e a"| "p_min (LP e1 a) (LP e2 b) = (LP (max e1 e2) (min a b))" fun e_less_eq :: "'e \ ('e::linorder, 'a::linorder) LP \ bool" where "e_less_eq e Infty = False"| "e_less_eq e (LP e' _) = (e \ e')" text_raw\\paragraph{Instantiation of classes}\ \\\ lemma p_min_re_neut[simp]: "p_min a Infty = a" by (induct a) auto lemma p_min_le_neut[simp]: "p_min Infty a = a" by (induct a) auto lemma p_min_asso: "p_min (p_min a b) c = p_min a (p_min b c)" apply(induct a b rule: p_min.induct ) apply (auto) apply (induct c) apply (auto) apply (metis max.assoc) apply (metis min.assoc) done lemma lp_mono: "class.monoid_add p_min Infty" by unfold_locales (auto simp add: p_min_asso) instantiation LP :: (linorder,linorder) monoid_add begin definition zero_def: "0 == Infty" definition plus_def: "a+b == p_min a b" instance by intro_classes (auto simp add: p_min_asso zero_def plus_def) end fun p_less_eq :: "('e, 'a::linorder) LP \ ('e, 'a) LP \ bool" where "p_less_eq (LP e a) (LP f b) = (a \ b)"| "p_less_eq _ Infty = True"| "p_less_eq Infty (LP e a) = False" fun p_less :: "('e, 'a::linorder) LP \ ('e, 'a) LP \ bool" where "p_less (LP e a) (LP f b) = (a < b)"| "p_less (LP e a) Infty = True"| "p_less Infty _ = False" lemma p_less_le_not_le : "p_less x y \ p_less_eq x y \ \ (p_less_eq y x)" by (induct x y rule: p_less.induct) auto lemma p_order_refl : "p_less_eq x x" by (induct x) auto lemma p_le_inf : "p_less_eq Infty x \ x = Infty" by (induct x) auto lemma p_order_trans : "\p_less_eq x y; p_less_eq y z\ \ p_less_eq x z" apply (induct y z rule: p_less.induct) apply auto apply (induct x) apply auto apply (cases x) apply auto apply(induct x) apply (auto simp add: p_le_inf) apply (metis p_le_inf p_less_eq.simps(2)) done lemma p_linear2 : "p_less_eq x y \ p_less_eq y x" apply (induct x y rule: p_less_eq.induct) apply auto done instantiation LP :: (type, linorder) preorder begin definition plesseq_def: "less_eq = p_less_eq" definition pless_def: "less = p_less" instance apply (intro_classes) apply (simp only: p_less_le_not_le pless_def plesseq_def) apply (simp only: p_order_refl plesseq_def pless_def) apply (simp only: plesseq_def) apply (metis p_order_trans) done end subsubsection "Operations" definition aluprio_\ :: "('s \ (unit \ ('e::linorder,'a::linorder) LP) list) \ 's \ ('e::linorder \ 'a::linorder)" where "aluprio_\ \ ft == (map_of (map p_unwrap (map snd (\ ft))))" definition aluprio_invar :: "('s \ (unit \ ('c::linorder, 'd::linorder) LP) list) \ ('s \ bool) \ 's \ bool" where "aluprio_invar \ invar ft == invar ft \ (\ x\set (\ ft). snd x\Infty) \ sorted (map fst (map p_unwrap (map snd (\ ft)))) \ distinct (map fst (map p_unwrap (map snd (\ ft)))) " definition aluprio_empty where "aluprio_empty empt = empt" definition aluprio_isEmpty where "aluprio_isEmpty isEmpty = isEmpty" definition aluprio_insert :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ bool) \ ('s \ 's \ 's) \ ('s \ unit \ ('e,'a) LP \ 's) \ 's \ 'e \ 'a \ 's" where " aluprio_insert splits annot isEmpty app consr s e a = (if e_less_eq e (annot s) \ \ isEmpty s then (let (l, (_,lp) , r) = splits (e_less_eq e) Infty s in (if e < fst (p_unwrap lp) then app (consr (consr l () (LP e a)) () lp) r else app (consr l () (LP e a)) r )) else consr s () (LP e a)) " definition aluprio_pop :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ 's \ 's) \ 's \ 'e \'a \'s" where "aluprio_pop splits annot app s = (let (l, (_,lp) , r) = splits (\ x. x \ (annot s)) Infty s in (case lp of (LP e a) \ (e, a, app l r) ))" definition aluprio_prio :: "((('e::linorder,'a::linorder) LP \ bool) \ ('e,'a) LP \ 's \ ('s \ (unit \ ('e,'a) LP) \ 's)) \ ('s \ ('e,'a) LP) \ ('s \ bool) \ 's \ 'e \ 'a option" where " aluprio_prio splits annot isEmpty s e = (if e_less_eq e (annot s) \ \ isEmpty s then (let (l, (_,lp) , r) = splits (e_less_eq e) Infty s in (if e = fst (p_unwrap lp) then Some (snd (p_unwrap lp)) else None)) else None) " lemmas aluprio_defs = aluprio_invar_def aluprio_\_def aluprio_empty_def aluprio_isEmpty_def aluprio_insert_def aluprio_pop_def aluprio_prio_def subsection "Correctness" subsubsection "Auxiliary Lemmas" lemma p_linear: "(x::('e, 'a::linorder) LP) \ y \ y \ x" by (unfold plesseq_def) (simp only: p_linear2) lemma e_less_eq_mon1: "e_less_eq e x \ e_less_eq e (x + y)" apply (cases x) apply (auto simp add: plus_def) apply (cases y) apply (auto simp add: max.coboundedI1) done lemma e_less_eq_mon2: "e_less_eq e y \ e_less_eq e (x + y)" apply (cases x) apply (auto simp add: plus_def) apply (cases y) apply (auto simp add: max.coboundedI2) done lemmas e_less_eq_mon = e_less_eq_mon1 e_less_eq_mon2 lemma p_less_eq_mon: "(x::('e::linorder,'a::linorder) LP) \ z \ (x + y) \ z" apply(cases y) apply(auto simp add: plus_def) apply (cases x) apply (cases z) apply (auto simp add: plesseq_def) apply (cases z) apply (auto simp add: min.coboundedI1) done lemma p_less_eq_lem1: "\\ (x::('e::linorder,'a::linorder) LP) \ z; (x + y) \ z\ \ y \ z " apply (cases x,auto simp add: plus_def) apply (cases y, auto) apply (cases z, auto simp add: plesseq_def) apply (metis min_le_iff_disj) done lemma infadd: "x \ Infty \x + y \ Infty" apply (unfold plus_def) apply (induct x y rule: p_min.induct) apply auto done lemma e_less_eq_sum_list: "\\ e_less_eq e (sum_list xs)\ \ \x \ set xs. \ e_less_eq e x" proof (induct xs) case Nil thus ?case by simp next case (Cons a xs) hence "\ e_less_eq e (sum_list xs)" by (auto simp add: e_less_eq_mon) hence v1: "\x\set xs. \ e_less_eq e x" using Cons.hyps by simp from Cons.prems have "\ e_less_eq e a" by (auto simp add: e_less_eq_mon) with v1 show "\x\set (a#xs). \ e_less_eq e x" by simp qed lemma e_less_eq_p_unwrap: "\x \ Infty;\ e_less_eq e x\ \ fst (p_unwrap x) < e" by (cases x) auto lemma e_less_eq_refl : "b \ Infty \ e_less_eq (fst (p_unwrap b)) b" by (cases b) auto lemma e_less_eq_sum_list2: assumes "\x\set (\s). snd x \ Infty" "((), b) \ set (\s)" shows "e_less_eq (fst (p_unwrap b)) (sum_list (map snd (\s)))" apply(insert assms) apply (induct "\s") apply (auto simp add: zero_def e_less_eq_mon e_less_eq_refl) done lemma e_less_eq_lem1: "\\ e_less_eq e a;e_less_eq e (a + b)\ \ e_less_eq e b" apply (auto simp add: plus_def) apply (cases a) apply auto apply (cases b) apply auto apply (metis le_max_iff_disj) done lemma p_unwrap_less_sum: "snd (p_unwrap ((LP e aa) + b)) \ aa" apply (cases b) apply (auto simp add: plus_def) done lemma sum_list_less_elems: "\x\set xs. snd x \ Infty \ \y\set (map snd (map p_unwrap (map snd xs))). snd (p_unwrap (sum_list (map snd xs))) \ y" proof (induct xs) case Nil thus ?case by simp next case (Cons a as) thus ?case apply auto apply (cases "(snd a)" rule: p_unwrap.cases) apply auto apply (cases "sum_list (map snd as)") apply auto apply (metis linorder_linear p_min_re_neut p_unwrap.simps plus_def[abs_def] snd_eqD) apply (auto simp add: p_unwrap_less_sum) apply (unfold plus_def) apply (cases "(snd a, sum_list (map snd as))" rule: p_min.cases) apply auto apply (cases "map snd as") apply (auto simp add: infadd) apply (metis min.coboundedI2 snd_conv) done qed lemma distinct_sortet_list_app: "\sorted xs; distinct xs; xs = as @ b # cs\ \ \ x\ set cs. b < x" by (metis distinct.simps(2) distinct_append - antisym_conv2 sorted.simps(2) sorted_append) + antisym_conv2 sorted_wrt.simps(2) sorted_append) lemma distinct_sorted_list_lem1: assumes "sorted xs" "sorted ys" "distinct xs" "distinct ys" " \ x \ set xs. x < e" " \ y \ set ys. e < y" shows "sorted (xs @ e # ys)" "distinct (xs @ e # ys)" proof - from assms (5,6) have "\x\set xs. \y\set ys. x \ y" by force thus "sorted (xs @ e # ys)" using assms by (auto simp add: sorted_append) have "set xs \ set ys = {}" using assms (5,6) by force thus "distinct (xs @ e # ys)" using assms by (auto) qed lemma distinct_sorted_list_lem2: assumes "sorted xs" "sorted ys" "distinct xs" "distinct ys" "e < e'" " \ x \ set xs. x < e" " \ y \ set ys. e' < y" shows "sorted (xs @ e # e' # ys)" "distinct (xs @ e # e' # ys)" proof - have "sorted (e' # ys)" "distinct (e' # ys)" "\ y \ set (e' # ys). e < y" using assms(2,4,5,7) by (auto) thus "sorted (xs @ e # e' # ys)" "distinct (xs @ e # e' # ys)" using assms(1,3,6) distinct_sorted_list_lem1[of xs "e' # ys" e] by auto qed lemma map_of_distinct_upd: "x \ set (map fst xs) \ [x \ y] ++ map_of xs = (map_of xs) (x \ y)" by (induct xs) (auto simp add: fun_upd_twist) lemma map_of_distinct_upd2: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd3: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd4: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)" apply(insert assms) apply(induct xs) apply clarsimp apply (metis dom_map_of_conv_image_fst fun_upd_None_restrict restrict_complement_singleton_eq restrict_map_self) apply (auto simp add: map_of_eq_None_iff) [] done lemma map_of_distinct_lookup: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) x = Some y" proof - have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x \ y)" using assms map_of_distinct_upd2 by simp thus ?thesis by simp qed lemma ran_distinct: assumes dist: "distinct (map fst al)" shows "ran (map_of al) = snd ` set al" using assms proof (induct al) case Nil then show ?case by simp next case (Cons kv al) then have "ran (map_of al) = snd ` set al" by simp moreover from Cons.prems have "map_of al (fst kv) = None" by (simp add: map_of_eq_None_iff) ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp qed subsubsection "Finite" lemma aluprio_finite_correct: "uprio_finite (aluprio_\ \) (aluprio_invar \ invar)" by(unfold_locales) (simp add: aluprio_defs finite_dom_map_of) subsubsection "Empty" lemma aluprio_empty_correct: assumes "al_empty \ invar empt" shows "uprio_empty (aluprio_\ \) (aluprio_invar \ invar) (aluprio_empty empt)" proof - interpret al_empty \ invar empt by fact show ?thesis apply (unfold_locales) apply (auto simp add: empty_correct aluprio_defs) done qed subsubsection "Is Empty" lemma aluprio_isEmpty_correct: assumes "al_isEmpty \ invar isEmpty" shows "uprio_isEmpty (aluprio_\ \) (aluprio_invar \ invar) (aluprio_isEmpty isEmpty)" proof - interpret al_isEmpty \ invar isEmpty by fact show ?thesis apply (unfold_locales) apply (auto simp add: aluprio_defs isEmpty_correct) done qed subsubsection "Insert" lemma annot_inf: assumes A: "invar s" "\x\set (\ s). snd x \ Infty" "al_annot \ invar annot" shows "annot s = Infty \ \ s = [] " proof - from A have invs: "invar s" by (simp add: aluprio_defs) interpret al_annot \ invar annot by fact show "annot s = Infty \ \ s = []" proof (cases "\ s = []") case True hence "map snd (\ s) = []" by simp hence "sum_list (map snd (\ s)) = Infty" by (auto simp add: zero_def) with invs have "annot s = Infty" by (auto simp add: annot_correct) with True show ?thesis by simp next case False hence " \x xs. (\ s) = x # xs" by (cases "\ s") auto from this obtain x xs where [simp]: "(\ s) = x # xs" by blast from this assms(2) have "snd x \ Infty" by (auto simp add: aluprio_defs) hence "sum_list (map snd (\ s)) \ Infty" by (auto simp add: infadd) thus ?thesis using annot_correct invs False by simp qed qed lemma e_less_eq_annot: assumes "al_annot \ invar annot" "invar s" "\x\set (\ s). snd x \ Infty" "\ e_less_eq e (annot s)" shows "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" proof - interpret al_annot \ invar annot by fact from assms(2) have "annot s = sum_list (map snd (\ s))" by (auto simp add: annot_correct) with assms(4) have "\x \ set (map snd (\ s)). \ e_less_eq e x" by (metis e_less_eq_sum_list) with assms(3) show ?thesis by (auto simp add: e_less_eq_p_unwrap) qed lemma aluprio_insert_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_isEmpty \ invar isEmpty" "al_app \ invar app" "al_consr \ invar consr" shows "uprio_insert (aluprio_\ \) (aluprio_invar \ invar) (aluprio_insert splits annot isEmpty app consr)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_isEmpty \ invar isEmpty by fact interpret al_app \ invar app by fact interpret al_consr \ invar consr by fact show ?thesis proof (unfold_locales, unfold aluprio_defs, goal_cases) case g1asms: (1 s e a) thus ?case proof (cases "e_less_eq e (annot s) \ \ isEmpty s") case False with g1asms show ?thesis apply (auto simp add: consr_correct ) proof goal_cases case prems: 1 with assms(2) have "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" by (simp add: e_less_eq_annot) with prems(3) show ?case by(auto simp add: sorted_append) next case prems: 2 hence "annot s = sum_list (map snd (\ s))" by (simp add: annot_correct) with prems show ?case by (auto simp add: e_less_eq_sum_list2) next case prems: 3 hence "\ s = []" by (auto simp add: isEmpty_correct) thus ?case by simp next case prems: 4 hence "\ s = []" by (auto simp add: isEmpty_correct) with prems show ?case by simp qed next case True note T1 = this obtain l uu lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 g1asms annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using g1asms v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: "\ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) g1asms v8 by auto have v10: "sorted (map fst (map p_unwrap (map snd (\ l))))" "distinct (map fst (map p_unwrap (map snd (\ l))))" "sorted (map fst (map p_unwrap (map snd (\ r))))" "distinct (map fst (map p_unwrap (map snd (\ l))))" using g1asms v8 by (auto simp add: sorted_append) from l_lp_r T1 g1asms show ?thesis proof (fold aluprio_insert_def, cases "e < fst (p_unwrap lp)") case True hence v11: "aluprio_insert splits annot isEmpty app consr s e a = app (consr (consr l () (LP e a)) () lp) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v12: "invar (app (consr (consr l () (LP e a)) () lp) r)" using v4(4,5) by (auto simp add: app_correct consr_correct) have v13: "\ (app (consr (consr l () (LP e a)) () lp) r) = \ l @ ((),(LP e a)) # ((), lp) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) hence v14: "(\x\set (\ (app (consr (consr l () (LP e a)) () lp) r)). snd x \ Infty)" using g1asms v4(1) by auto have v15: "e = fst(p_unwrap (LP e a))" by simp hence v16: "sorted (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # ((), lp) # \ r))))" "distinct (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # ((), lp) # \ r))))" using v10(1,3) v7 True v9 v4(1) g1asms distinct_sorted_list_lem2 by (auto simp add: sorted_append) thus "invar (aluprio_insert splits annot isEmpty app consr s e a) \ (\x\set (\ (aluprio_insert splits annot isEmpty app consr s e a)). snd x \ Infty) \ sorted (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a))))) \ distinct (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))))" using v11 v12 v13 v14 by simp next case False hence v11: "aluprio_insert splits annot isEmpty app consr s e a = app (consr l () (LP e a)) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v12: "invar (app (consr l () (LP e a)) r)" using v4(4,5) by (auto simp add: app_correct consr_correct) have v13: "\ (app (consr l () (LP e a)) r) = \ l @ ((),(LP e a)) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) hence v14: "(\x\set (\ (app (consr l () (LP e a)) r)). snd x \ Infty)" using g1asms v4(1) by auto have v15: "e = fst(p_unwrap (LP e a))" by simp have v16: "e = fst(p_unwrap lp)" using False v5 by (cases lp) auto hence v17: "sorted (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # \ r))))" "distinct (map fst (map p_unwrap (map snd (\ l @ ((),(LP e a)) # \ r))))" using v16 v15 v10(1,3) v7 True v9 v4(1) g1asms distinct_sorted_list_lem1 by (auto simp add: sorted_append) thus "invar (aluprio_insert splits annot isEmpty app consr s e a) \ (\x\set (\ (aluprio_insert splits annot isEmpty app consr s e a)). snd x \ Infty) \ sorted (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a))))) \ distinct (map fst (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))))" using v11 v12 v13 v14 by simp qed qed next case g1asms: (2 s e a) thus ?case proof (cases "e_less_eq e (annot s) \ \ isEmpty s") case False with g1asms show ?thesis apply (auto simp add: consr_correct) proof goal_cases case prems: 1 with assms(2) have "\x \ set (map (fst \ (p_unwrap \ snd)) (\ s)). x < e" by (simp add: e_less_eq_annot) hence "e \ set (map fst ((map (p_unwrap \ snd)) (\ s)))" by auto thus ?case by (auto simp add: map_of_distinct_upd) next case prems: 2 hence "\ s = []" by (auto simp add: isEmpty_correct) thus ?case by simp qed next case True note T1 = this obtain l lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 g1asms annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using g1asms v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) g1asms v8 by auto hence v10: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). e < x" using v6 by auto have v11: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using v7 v10 v8 g1asms by auto from l_lp_r T1 g1asms show ?thesis proof (fold aluprio_insert_def, cases "e < fst (p_unwrap lp)") case True hence v12: "aluprio_insert splits annot isEmpty app consr s e a = app (consr (consr l () (LP e a)) () lp) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v13: "\ (app (consr (consr l () (LP e a)) () lp) r) = \ l @ ((),(LP e a)) # ((), lp) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) have v14: "e = fst(p_unwrap (LP e a))" by simp have v15: "e \ set (map fst (map p_unwrap (map snd(((),lp)#\ r))))" using v11(2) True by auto note map_of_distinct_upd2[OF v11(1) v15] thus "map_of (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))) = map_of (map p_unwrap (map snd (\ s)))(e \ a)" using v12 v13 v4(1) by simp next case False hence v12: "aluprio_insert splits annot isEmpty app consr s e a = app (consr l () (LP e a)) r" using l_lp_r T1 by (auto simp add: aluprio_defs) have v13: "\ (app (consr l () (LP e a)) r) = \ l @ ((),(LP e a)) # \ r" using v4(4,5) by (auto simp add: app_correct consr_correct) have v14: "e = fst(p_unwrap lp)" using False v5 by (cases lp) auto note v15 = map_of_distinct_upd3[OF v11(1) v11(2)] have v16:"(map p_unwrap (map snd (\ s))) = (map p_unwrap (map snd (\ l))) @ (e,snd(p_unwrap lp)) # (map p_unwrap (map snd (\ r)))" using v4(1) v14 by simp note v15[of a "snd(p_unwrap lp)"] thus "map_of (map p_unwrap (map snd (\ (aluprio_insert splits annot isEmpty app consr s e a)))) = map_of (map p_unwrap (map snd (\ s)))(e \ a)" using v12 v13 v16 by simp qed qed qed qed subsubsection "Prio" lemma aluprio_prio_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_isEmpty \ invar isEmpty" shows "uprio_prio (aluprio_\ \) (aluprio_invar \ invar) (aluprio_prio splits annot isEmpty)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_isEmpty \ invar isEmpty by fact show ?thesis proof (unfold_locales) fix s e assume inv1: "aluprio_invar \ invar s" hence sinv: "invar s" "(\ x\set (\ s). snd x\Infty)" "sorted (map fst (map p_unwrap (map snd (\ s))))" "distinct (map fst (map p_unwrap (map snd (\ s))))" by (auto simp add: aluprio_defs) show "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" proof(cases "e_less_eq e (annot s) \ \ isEmpty s") case False note F1 = this thus ?thesis proof(cases "isEmpty s") case True hence "\ s = []" using sinv isEmpty_correct by simp hence "aluprio_\ \ s = Map.empty" by (simp add:aluprio_defs) hence "aluprio_\ \ s e = None" by simp thus "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" using F1 by (auto simp add: aluprio_defs) next case False hence v3:"\ e_less_eq e (annot s)" using F1 by simp note v4=e_less_eq_annot[OF assms(2)] note v4[OF sinv(1) sinv(2) v3] hence v5:"e\set (map (fst \ (p_unwrap \ snd)) (\ s))" by auto hence "map_of (map (p_unwrap \ snd) (\ s)) e = None" using map_of_eq_None_iff by (metis map_map map_of_eq_None_iff set_map v5) thus "aluprio_prio splits annot isEmpty s e = aluprio_\ \ s e" using F1 by (auto simp add: aluprio_defs) qed next case True note T1 = this obtain l uu lp r where l_lp_r: "(splits (e_less_eq e) Infty s) = (l, ((), lp), r) " by (cases "splits (e_less_eq e) Infty s", auto) note v2 = splits_correct[of s "e_less_eq e" Infty l "()" lp r] have v3: "invar s" "\ e_less_eq e Infty" "e_less_eq e (Infty + sum_list (map snd (\ s)))" using T1 sinv annot_correct by (auto simp add: plus_def) have v4: "\ s = \ l @ ((), lp) # \ r" "\ e_less_eq e (Infty + sum_list (map snd (\ l)))" "e_less_eq e (Infty + sum_list (map snd (\ l)) + lp)" "invar l" "invar r" using v2[OF v3(1) _ v3(2) v3(3) l_lp_r] e_less_eq_mon(1) by auto hence v5: "e_less_eq e lp" by (metis e_less_eq_lem1) hence v6: "e \ (fst (p_unwrap lp))" by (cases lp) auto have "(Infty + sum_list (map snd (\ l))) = (annot l)" by (metis add_0_left annot_correct v4(4) zero_def) hence v7:"\ e_less_eq e (annot l)" using v4(2) by simp have "\x\set (\ l). snd x \ Infty" using sinv v4(1) by simp hence v7: "\x \ set (map (fst \ (p_unwrap \ snd)) (\ l)). x < e" using v4(4) v7 assms(2) by(simp add: e_less_eq_annot) have v8:"map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ fst(p_unwrap lp) # map fst (map p_unwrap (map snd (\ r)))" using v4(1) by simp note distinct_sortet_list_app[of "map fst (map p_unwrap (map snd (\ s)))" "map fst (map p_unwrap (map snd (\ l)))" "fst(p_unwrap lp)" "map fst (map p_unwrap (map snd (\ r)))"] hence v9: "\ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). fst(p_unwrap lp) < x" using v4(1) sinv v8 by auto hence v10: " \ x\set (map (fst \ (p_unwrap \ snd)) (\ r)). e < x" using v6 by auto have v11: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using v7 v10 v8 sinv by auto from l_lp_r T1 sinv show ?thesis proof (cases "e = fst (p_unwrap lp)") case False have v12: "e \ set (map fst (map p_unwrap (map snd(\ s))))" using v11 False v4(1) by auto hence "map_of (map (p_unwrap \ snd) (\ s)) e = None" using map_of_eq_None_iff by (metis map_map map_of_eq_None_iff set_map v12) thus ?thesis using T1 False l_lp_r by (auto simp add: aluprio_defs) next case True have v12: "map (p_unwrap \ snd) (\ s) = map p_unwrap (map snd (\ l)) @ (e,snd (p_unwrap lp)) # map p_unwrap (map snd (\ r))" using v4(1) True by simp note map_of_distinct_lookup[OF v11] hence "map_of (map (p_unwrap \ snd) (\ s)) e = Some (snd (p_unwrap lp))" using v12 by simp thus ?thesis using T1 True l_lp_r by (auto simp add: aluprio_defs) qed qed qed qed subsubsection "Pop" lemma aluprio_pop_correct: assumes "al_splits \ invar splits" "al_annot \ invar annot" "al_app \ invar app" shows "uprio_pop (aluprio_\ \) (aluprio_invar \ invar) (aluprio_pop splits annot app)" proof - interpret al_splits \ invar splits by fact interpret al_annot \ invar annot by fact interpret al_app \ invar app by fact show ?thesis proof (unfold_locales) fix s e a s' assume A: "aluprio_invar \ invar s" "aluprio_\ \ s \ Map.empty" "aluprio_pop splits annot app s = (e, a, s')" hence v1: "\ s \ []" by (auto simp add: aluprio_defs) obtain l lp r where l_lp_r: "splits (\ x. x\annot s) Infty s = (l,((),lp),r)" by (cases "splits (\ x. x\annot s) Infty s", auto) have invs: "invar s" "(\x\set (\ s). snd x \ Infty)" "sorted (map fst (map p_unwrap (map snd (\ s))))" "distinct (map fst (map p_unwrap (map snd (\ s))))" using A by (auto simp add:aluprio_defs) note a1 = annot_inf[of invar s \ annot] note a1[OF invs(1) invs(2) assms(2)] hence v2: "annot s \ Infty" using v1 by simp hence v3: "\ Infty \ annot s" by(cases "annot s") (auto simp add: plesseq_def) have v4: "annot s = sum_list (map snd (\ s))" by (auto simp add: annot_correct invs(1)) hence v5: "(Infty + sum_list (map snd (\ s))) \ annot s" by (auto simp add: plus_def) note p_mon = p_less_eq_mon[of _ "annot s"] note v6 = splits_correct[OF invs(1)] note v7 = v6[of "\ x. x \ annot s"] note v7[OF _ v3 v5 l_lp_r] p_mon hence v8: " \ s = \ l @ ((), lp) # \ r" "\ Infty + sum_list (map snd (\ l)) \ annot s" "Infty + sum_list (map snd (\ l)) + lp \ annot s" "invar l" "invar r" by auto hence v9: "lp \ Infty" using invs(2) by auto hence v10: "s' = app l r" "(e,a) = p_unwrap lp" using l_lp_r A(3) apply (auto simp add: aluprio_defs) apply (cases lp) apply auto apply (cases lp) apply auto done have "lp \ annot s" using v8(2,3) p_less_eq_lem1 by auto hence v11: "a \ snd (p_unwrap (annot s))" using v10(2) v2 v9 apply (cases "annot s") apply auto apply (cases lp) apply (auto simp add: plesseq_def) done note sum_list_less_elems[OF invs(2)] hence v12: "\y\set (map snd (map p_unwrap (map snd (\ s)))). a \ y" using v4 v11 by auto have "ran (aluprio_\ \ s) = set (map snd (map p_unwrap (map snd (\ s))))" using ran_distinct[OF invs(4)] apply (unfold aluprio_defs) apply (simp only: set_map) done hence ziel1: "\y\ran (aluprio_\ \ s). a \ y" using v12 by simp have v13: "map p_unwrap (map snd (\ s)) = map p_unwrap (map snd (\ l)) @ (e,a) # map p_unwrap (map snd (\ r))" using v8(1) v10 by auto hence v14: "map fst (map p_unwrap (map snd (\ s))) = map fst (map p_unwrap (map snd (\ l))) @ e # map fst (map p_unwrap (map snd (\ r)))" by auto hence v15: "e \ set (map fst (map p_unwrap (map snd (\ l))))" "e \ set (map fst (map p_unwrap (map snd (\ r))))" using invs(4) by auto note map_of_distinct_lookup[OF v15] note this[of a] hence ziel2: "aluprio_\ \ s e = Some a" using v13 by (unfold aluprio_defs, auto) have v16: "\ s' = \ l @ \ r" "invar s'" using v8(4,5) app_correct v10 by auto note map_of_distinct_upd4[OF v15] note this[of a] hence ziel3: "aluprio_\ \ s' = (aluprio_\ \ s)(e := None)" unfolding aluprio_defs using v16(1) v13 by auto have ziel4: "aluprio_invar \ invar s'" using v16 v8(1) invs(2,3,4) unfolding aluprio_defs by (auto simp add: sorted_append) show "aluprio_invar \ invar s' \ aluprio_\ \ s' = (aluprio_\ \ s)(e := None) \ aluprio_\ \ s e = Some a \ (\y\ran (aluprio_\ \ s). a \ y)" using ziel1 ziel2 ziel3 ziel4 by simp qed qed lemmas aluprio_correct = aluprio_finite_correct aluprio_empty_correct aluprio_isEmpty_correct aluprio_insert_correct aluprio_pop_correct aluprio_prio_correct locale aluprio_defs = StdALDefs ops for ops :: "(unit,('e::linorder,'a::linorder) LP,'s) alist_ops" begin definition [icf_rec_def]: "aluprio_ops \ \ upr_\ = aluprio_\ \, upr_invar = aluprio_invar \ invar, upr_empty = aluprio_empty empty, upr_isEmpty = aluprio_isEmpty isEmpty, upr_insert = aluprio_insert splits annot isEmpty app consr, upr_pop = aluprio_pop splits annot app, upr_prio = aluprio_prio splits annot isEmpty \" end locale aluprio = aluprio_defs ops + StdAL ops for ops :: "(unit,('e::linorder,'a::linorder) LP,'s) alist_ops" begin lemma aluprio_ops_impl: "StdUprio aluprio_ops" apply (rule StdUprio.intro) apply (simp_all add: icf_rec_unf) apply (rule aluprio_correct) apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] apply (rule aluprio_correct, unfold_locales) [] done end end diff --git a/thys/Collections/Iterator/Proper_Iterator.thy b/thys/Collections/Iterator/Proper_Iterator.thy --- a/thys/Collections/Iterator/Proper_Iterator.thy +++ b/thys/Collections/Iterator/Proper_Iterator.thy @@ -1,269 +1,268 @@ section \Proper Iterators\ theory Proper_Iterator imports SetIteratorOperations Automatic_Refinement.Refine_Lib begin text \ Proper iterators provide a way to obtain polymorphic iterators even inside locale contexts. For this purpose, an iterator that converts the set to a list is fixed inside the locale, and polymorphic iterators are described by folding over the generated list. In order to ensure efficiency, it is shown that folding over the generated list is equivalent to directly iterating over the set, and this equivalence is set up as a code preprocessing rule. \ subsection \Proper Iterators\ text \A proper iterator can be expressed as a fold over a list, where the list does only depend on the set. In particular, it does not depend on the type of the state. We express this by the following definition, using two iterators with different types:\ definition proper_it :: "('x,'\1) set_iterator \ ('x,'\2) set_iterator \ bool" where "proper_it it it' \ (\l. it=foldli l \ it'=foldli l)" lemma proper_itI[intro?]: fixes it :: "('x,'\1) set_iterator" and it' :: "('x,'\2) set_iterator" assumes "it=foldli l \ it'=foldli l" shows "proper_it it it'" using assms unfolding proper_it_def by auto lemma proper_itE: fixes it :: "('x,'\1) set_iterator" and it' :: "('x,'\2) set_iterator" assumes "proper_it it it'" obtains l where "it=foldli l" and "it'=foldli l" using assms unfolding proper_it_def by auto lemma proper_it_parE: fixes it :: "'a \ ('x,'\1) set_iterator" and it' :: "'a \ ('x,'\2) set_iterator" assumes "\x. proper_it (it x) (it' x)" obtains f where "it = (\x. foldli (f x))" and "it' = (\x. foldli (f x))" using assms unfolding proper_it_def by metis definition proper_it' where "proper_it' it it' \ \s. proper_it (it s) (it' s)" lemma proper_it'I: "\\s. proper_it (it s) (it' s)\ \ proper_it' it it'" unfolding proper_it'_def by blast lemma proper_it'D: "proper_it' it it' \ proper_it (it s) (it' s)" unfolding proper_it'_def by blast subsubsection \Properness Preservation\ ML \ structure Icf_Proper_Iterator = struct structure icf_proper_iteratorI = Named_Thms ( val name = @{binding icf_proper_iteratorI_raw} val description = "ICF (internal): Rules to show properness of iterators" ) val get = icf_proper_iteratorI.get fun add_thm thm = icf_proper_iteratorI.add_thm thm val add = Thm.declaration_attribute add_thm fun del_thm thm = icf_proper_iteratorI.del_thm thm val del = Thm.declaration_attribute del_thm val setup = I #> icf_proper_iteratorI.setup #> Attrib.setup @{binding icf_proper_iteratorI} (Attrib.add_del add del) ("ICF: Rules to show properness of iterators") #> Global_Theory.add_thms_dynamic (@{binding icf_proper_iteratorI}, get o Context.proof_of ) end \ setup \Icf_Proper_Iterator.setup\ lemma proper_iterator_trigger: "proper_it it it' \ proper_it it it'" "proper_it' itf itf' \ proper_it' itf itf'" . declaration \ Tagged_Solver.declare_solver @{thms proper_iterator_trigger} @{binding proper_iterator} "Proper iterator solver" (fn ctxt => REPEAT_ALL_NEW (resolve_tac ctxt (Icf_Proper_Iterator.get ctxt))) \ lemma pi_foldli[icf_proper_iteratorI]: "proper_it (foldli l :: ('a,'\) set_iterator) (foldli l)" unfolding proper_it_def by auto lemma pi_foldri[icf_proper_iteratorI]: "proper_it (foldri l :: ('a,'\) set_iterator) (foldri l)" unfolding proper_it_def foldri_def by auto lemma pi'_foldli[icf_proper_iteratorI]: "proper_it' (foldli o tsl) (foldli o tsl)" apply (clarsimp simp add: proper_it'_def) apply (tagged_solver) done lemma pi'_foldri[icf_proper_iteratorI]: "proper_it' (foldri o tsl) (foldri o tsl)" apply (clarsimp simp add: proper_it'_def) apply (tagged_solver) done text \Iterator combinators preserve properness\ lemma pi_emp[icf_proper_iteratorI]: "proper_it set_iterator_emp set_iterator_emp" unfolding proper_it_def set_iterator_emp_def[abs_def] by (auto intro!: ext exI[where x="[]"]) lemma pi_sng[icf_proper_iteratorI]: "proper_it (set_iterator_sng x) (set_iterator_sng x)" unfolding proper_it_def set_iterator_sng_def[abs_def] by (auto intro!: ext exI[where x="[x]"]) lemma pi_union[icf_proper_iteratorI]: assumes PA: "proper_it it_a it_a'" assumes PB: "proper_it it_b it_b'" shows "proper_it (set_iterator_union it_a it_b) (set_iterator_union it_a' it_b')" unfolding set_iterator_union_def apply (rule proper_itE[OF PA]) apply (rule proper_itE[OF PB]) apply (rule_tac l="l@la" in proper_itI) apply simp apply (intro conjI ext) apply (simp_all add: foldli_append) done lemma pi_product[icf_proper_iteratorI]: fixes it_a :: "('a,'\a) set_iterator" fixes it_b :: "'a \ ('b,'\a) set_iterator" assumes PA: "proper_it it_a it_a'" and PB: "\x. proper_it (it_b x) (it_b' x)" shows "proper_it (set_iterator_product it_a it_b) (set_iterator_product it_a' it_b')" proof - from PB have PB': "\x. proper_it (it_b x) (it_b' x)" .. show ?thesis unfolding proper_it_def apply (rule proper_itE[OF PA]) apply (rule proper_it_parE[OF PB']) apply (auto simp add: set_iterator_product_foldli_conv) done qed lemma pi_image_filter[icf_proper_iteratorI]: fixes it :: "('x,'\1) set_iterator" and it' :: "('x,'\2) set_iterator" and g :: "'x \ 'y option" assumes P: "proper_it it it'" shows "proper_it (set_iterator_image_filter g it) (set_iterator_image_filter g it')" unfolding proper_it_def apply (rule proper_itE[OF P]) apply (auto simp: set_iterator_image_filter_foldli_conv) done lemma pi_filter[icf_proper_iteratorI]: assumes P: "proper_it it it'" shows "proper_it (set_iterator_filter P it) (set_iterator_filter P it')" unfolding proper_it_def apply (rule proper_itE[OF P]) by (auto simp: set_iterator_filter_foldli_conv) lemma pi_image[icf_proper_iteratorI]: assumes P: "proper_it it it'" shows "proper_it (set_iterator_image g it) (set_iterator_image g it')" unfolding proper_it_def apply (rule proper_itE[OF P]) by (auto simp: set_iterator_image_foldli_conv) lemma pi_dom[icf_proper_iteratorI]: assumes P: "proper_it it it'" shows "proper_it (map_iterator_dom it) (map_iterator_dom it')" unfolding proper_it_def apply (rule proper_itE[OF P]) by (auto simp: map_iterator_dom_foldli_conv) lemma set_iterator_product_eq2: assumes "\a\set la. itb a = itb' a" shows "set_iterator_product (foldli la) itb = set_iterator_product (foldli la) itb'" proof (intro ext) fix c f \ show "set_iterator_product (foldli la) itb c f \ = set_iterator_product (foldli la) itb' c f \" using assms unfolding set_iterator_product_def apply (induct la arbitrary: \) apply (auto) done qed subsubsection \Optimizing Folds\ text \ Using an iterator to create a list. The optimizations will match the pattern \foldli (it_to_list it s)\ \ definition "it_to_list it s \ (it s) (\_. True) (\x l. l@[x]) []" lemma map_it_to_list_genord_correct: assumes A: "map_iterator_genord (it s) m (\(k,_) (k',_). R k k')" shows "map_of (it_to_list it s) = m \ distinct (map fst (it_to_list it s)) \ sorted_wrt R ((map fst (it_to_list it s)))" unfolding it_to_list_def apply (rule map_iterator_genord_rule_insert_P[OF A, where I=" \it l. map_of l = m |` it \ distinct (map fst l) \ sorted_wrt R ((map fst l)) "]) apply auto apply (auto simp: restrict_map_def) [] apply (metis Some_eq_map_of_iff restrict_map_eq(2)) apply (auto simp add: sorted_wrt_append) by (metis (lifting) restrict_map_eq(2) weak_map_of_SomeI) lemma (in linorder) map_it_to_list_linord_correct: assumes A: "map_iterator_linord (it s) m" shows "map_of (it_to_list it s) = m \ distinct (map fst (it_to_list it s)) \ sorted ((map fst (it_to_list it s)))" using map_it_to_list_genord_correct[where it=it, - OF A[unfolded set_iterator_map_linord_def]] - by (simp add: sorted_sorted_wrt) + OF A[unfolded set_iterator_map_linord_def]] . lemma (in linorder) map_it_to_list_rev_linord_correct: assumes A: "map_iterator_rev_linord (it s) m" shows "map_of (it_to_list it s) = m \ distinct (map fst (it_to_list it s)) \ sorted (rev (map fst (it_to_list it s)))" using map_it_to_list_genord_correct[where it=it, OF A[unfolded set_iterator_map_rev_linord_def]] by simp end diff --git a/thys/Collections/Iterator/SetIterator.thy b/thys/Collections/Iterator/SetIterator.thy --- a/thys/Collections/Iterator/SetIterator.thy +++ b/thys/Collections/Iterator/SetIterator.thy @@ -1,795 +1,795 @@ (* Title: Iterators over Finite Sets Author: Thomas Tuerk Maintainer: Thomas Tuerk *) section \\isaheader{Iterators over Finite Sets}\ theory SetIterator imports Automatic_Refinement.Misc Automatic_Refinement.Foldi (*"../../Refine_Monadic/Refine_Monadic"*) begin text \When reasoning about finite sets, it is often handy to be able to iterate over the elements of the set. If the finite set is represented by a list, the @{term fold} operation can be used. For general finite sets, the order of elements is not fixed though. Therefore, nondeterministic iterators are introduced in this theory.\ subsection \General Iterators\ type_synonym ('x,'\) set_iterator = "('\\bool) \ ('x\'\\'\) \ '\ \ '\" locale set_iterator_genord = fixes iti::"('x,'\) set_iterator" and S0::"'x set" and R::"'x \ 'x \ bool" assumes foldli_transform: "\l0. distinct l0 \ S0 = set l0 \ sorted_wrt R l0 \ iti = foldli l0" begin text \Let's prove some trivial lemmata to show that the formalisation agrees with the view of iterators described above.\ lemma set_iterator_weaken_R : "(\x y. \x \ S0; y \ S0; R x y\ \ R' x y) \ set_iterator_genord iti S0 R'" by (metis set_iterator_genord.intro foldli_transform sorted_wrt_mono_rel) lemma finite_S0 : shows "finite S0" by (metis finite_set foldli_transform) lemma iti_stop_rule_cond : assumes not_c: "\(c \)" shows "iti c f \ = \" proof - from foldli_transform obtain l0 where l0_props: "iti c = foldli l0 c" by blast with foldli_not_cond [of c \ l0 f, OF not_c] show ?thesis by simp qed lemma iti_stop_rule_emp : assumes S0_emp: "S0 = {}" shows "iti c f \ = \" proof - from foldli_transform obtain l0 where l0_props: "S0 = set l0" "iti c = foldli l0 c" by blast with foldli.simps(1) [of c f \] S0_emp show ?thesis by simp qed text \Reducing everything to folding is cumbersome. Let's define a few high-level inference rules.\ lemma iteratei_rule_P: assumes "I S0 \0" "\S \ x. \ c \; x \ S; I S \; S \ S0; \y\S - {x}. R x y; \y\S0 - S. R y x\ \ I (S - {x}) (f x \)" "\\. I {} \ \ P \" "\\ S. \ S \ S0; S \ {}; \ c \; I S \; \x\S. \y\S0-S. R y x \ \ P \" shows "P (iti c f \0)" proof - { fix P I \0 f assume I: "I S0 \0" and R: "\S \ x. \c \; x \ S; I S \; S \ S0; \y\S-{x}. R x y\ \ I (S - {x}) (f x \)" and C1: "I {} (iti c f \0) \ P" and C2:"\S. \S \ S0; S \ {}; \ c (iti c f \0); I S (iti c f \0)\ \ P" from foldli_transform obtain l0 where l0_props: "distinct l0" "S0 = set l0" "sorted_wrt R l0" "iti c = foldli l0 c" by auto from I R have "I {} (iti c f \0) \ (\S. S \ S0 \ S \ {} \ \ (c (iti c f \0)) \ I S (iti c f \0))" unfolding l0_props using l0_props(1,3) proof (induct l0 arbitrary: \0) case Nil thus ?case by simp next case (Cons x l0) note ind_hyp = Cons(1) note I_x_l0 = Cons(2) note step = Cons(3) from Cons(4) have dist_l0: "distinct l0" and x_nin_l0: "x \ set l0" by simp_all from Cons(5) have R_l0: "\y\set l0. R x y" and sort_l0: "sorted_wrt R l0" by simp_all show ?case proof (cases "c \0") case False with I_x_l0 show ?thesis apply (rule_tac disjI2) apply (rule_tac exI[where x="set (x # l0)"]) apply (simp) done next case True note c_\0 = this from step[of \0 x "set (x # l0)"] I_x_l0 R_l0 c_\0 x_nin_l0 have step': "I (set l0) (f x \0)" by (simp_all add: Ball_def) from ind_hyp [OF step' step dist_l0 sort_l0] have "I {} (foldli l0 c f (f x \0)) \ (\S. S \ set l0 \ S \ {} \ \ c (foldli l0 c f (f x \0)) \ I S (foldli l0 c f (f x \0)))" by (fastforce) thus ?thesis by (simp add: c_\0 subset_iff) metis qed qed with C1 C2 have "P" by blast } note aux = this from assms show ?thesis apply (rule_tac aux [of "\S \. I S \ \ (\x\S. \y\S0-S. R y x)" \0 f]) apply auto done qed text \Instead of removing elements one by one from the invariant, adding them is sometimes more natural.\ lemma iteratei_rule_insert_P: assumes pre : "I {} \0" "\S \ x. \ c \; x \ S0 - S; I S \; S \ S0; \y\(S0 - S) - {x}. R x y; \y\S. R y x\ \ I (insert x S) (f x \)" "\\. I S0 \ \ P \" "\\ S. \ S \ S0; S \ S0; \ (c \); I S \; \x\S0-S. \y\S. R y x \ \ P \" shows "P (iti c f \0)" proof - let ?I' = "\S \. I (S0 - S) \" have pre1: "!!\ S. \ S \ S0; S \ {}; \ (c \); ?I' S \; \x\S. \y\S0-S. R y x\ \ P \" proof - fix S \ assume AA: "S \ S0" "S \ {}" "\ (c \)" "?I' S \" "\x\S. \y\S0-S. R y x" with pre(4) [of "S0 - S"] show "P \" by auto qed have pre2 :"\x S \. \c \; x \ S; S \ S0; ?I' S \; \y\S - {x}. R x y; \y\S0-S. R y x \ \ ?I' (S - {x}) (f x \)" proof - fix x S \ assume AA : "c \" "x \ S" "S \ S0" "?I' S \" "\y\S - {x}. R x y" "\y\S0 - S. R y x" from AA(2) AA(3) have "S0 - (S - {x}) = insert x (S0 - S)" "S0 - (S0 - S) = S" by auto moreover note pre(2) [of \ x "S0 - S"] AA ultimately show "?I' (S - {x}) (f x \)" by auto qed show "P (iti c f \0)" apply (rule iteratei_rule_P [of ?I' \0 c f P]) apply (simp add: pre) apply (rule pre2) apply simp_all apply (simp add: pre(3)) apply (simp add: pre1) done qed text \Show that iti without interruption is related to fold\ lemma iti_fold: assumes lc_f: "comp_fun_commute f" shows "iti (\_. True) f \0 = Finite_Set.fold f \0 S0" proof (rule iteratei_rule_insert_P [where I = "\X \'. \' = Finite_Set.fold f \0 X"]) fix S \ x assume "x \ S0 - S" "S \ S0" and \_eq: "\ = Finite_Set.fold f \0 S" from finite_S0 \S \ S0\ have fin_S: "finite S" by (metis finite_subset) from \x \ S0 - S\ have x_nin_S: "x \ S" by simp note fold_eq = comp_fun_commute.fold_insert [OF lc_f fin_S x_nin_S] show "f x \ = Finite_Set.fold f \0 (insert x S)" by (simp add: fold_eq \_eq) qed simp_all end subsection \Iterators over Maps\ type_synonym ('k,'v,'\) map_iterator = "('k\'v,'\) set_iterator" text \Iterator over the key-value pairs of a finite map are called iterators over maps.\ abbreviation "map_iterator_genord it m R \ set_iterator_genord it (map_to_set m) R" subsection \Unordered Iterators\ text \Often one does not care about the order in which the elements are processed. Therefore, the selection function can be set to not impose any further restrictings. This leads to considerably simpler theorems.\ definition "set_iterator it S0 \ set_iterator_genord it S0 (\_ _. True)" abbreviation "map_iterator it m \ set_iterator it (map_to_set m)" lemma set_iterator_intro : "set_iterator_genord it S0 R \ set_iterator it S0" unfolding set_iterator_def apply (rule set_iterator_genord.set_iterator_weaken_R [where R = R]) apply simp_all done lemma set_iterator_no_cond_rule_P: "\ set_iterator it S0; I S0 \0; !!S \ x. \ x \ S; I S \; S \ S0 \ \ I (S - {x}) (f x \); !!\. I {} \ \ P \ \ \ P (it (\_. True) f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_P [of it S0 "\_ _. True" I \0 "\_. True" f P] by simp lemma set_iterator_no_cond_rule_insert_P: "\ set_iterator it S0; I {} \0; !!S \ x. \ x \ S0 - S; I S \; S \ S0 \ \ I (insert x S) (f x \); !!\. I S0 \ \ P \ \ \ P (it (\_. True) f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_insert_P [of it S0 "\_ _. True" I \0 "\_. True" f P] by simp lemma set_iterator_rule_P: "\ set_iterator it S0; I S0 \0; !!S \ x. \ c \; x \ S; I S \; S \ S0 \ \ I (S - {x}) (f x \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_P [of it S0 "\_ _. True" I \0 c f P] by simp lemma set_iterator_rule_insert_P: "\ set_iterator it S0; I {} \0; !!S \ x. \ c \; x \ S0 - S; I S \; S \ S0 \ \ I (insert x S) (f x \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_def using set_iterator_genord.iteratei_rule_insert_P [of it S0 "\_ _. True" I \0 c f P] by simp text\The following rules is adapted for maps. Instead of a set of key-value pairs the invariant now only sees the keys.\ lemma map_iterator_genord_rule_P: assumes "map_iterator_genord it m R" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \; \k' v'. k' \ it-{k} \ m k' = Some v' \ R (k, v) (k', v'); \k' v'. k' \ it \ m k' = Some v' \ R (k', v') (k, v)\ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \; \k v k' v'. k \ it \ m k = Some v \ k' \ it \ m k' = Some v' \ R (k, v) (k', v') \ \ P \" shows "P (it c f \0)" proof (rule set_iterator_genord.iteratei_rule_P [of it "map_to_set m" R "\S \. I (fst ` S) \" \0 c f P]) show "map_iterator_genord it m R" by fact next show "I (fst ` map_to_set m) \0" using I0 by (simp add: map_to_set_dom[symmetric]) next fix \ assume "I (fst ` {}) \" with IF show "P \" by simp next fix \ S assume "S \ map_to_set m" "S \ {}" "\ c \" "I (fst ` S) \" and R_prop: "\x\S. \y\map_to_set m - S. R y x" let ?S' = "fst ` S" show "P \" proof (rule II [where it = ?S']) from \S \ map_to_set m\ show "?S' \ dom m" unfolding map_to_set_dom by auto next from \S \ {}\ show "?S' \ {}" by auto next show "\ (c \)" by fact next show "I (fst ` S) \" by fact next show "\k v k' v'. k \ fst ` S \ m k = Some v \ k' \ fst ` S \ m k' = Some v' \ R (k, v) (k', v')" proof (intro allI impI, elim conjE ) fix k v k' v' assume pre_k: "k \ fst ` S" "m k = Some v" assume pre_k': "k' \ fst ` S" "m k' = Some v'" from \S \ map_to_set m\ pre_k' have kv'_in: "(k', v') \ S" unfolding map_to_set_def by auto from \S \ map_to_set m\ pre_k have kv_in: "(k, v) \ map_to_set m - S" unfolding map_to_set_def by (auto simp add: image_iff) from R_prop kv_in kv'_in show "R (k, v) (k',v')" by simp qed qed next fix \ S kv assume "S \ map_to_set m" "kv \ S" "c \" and I_S': "I (fst ` S) \" and R_S: "\kv'\S - {kv}. R kv kv'" and R_not_S: "\kv'\map_to_set m - S. R kv' kv" let ?S' = "fst ` S" obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) have "I (fst ` S - {k}) (f (k, v) \)" proof (rule IP) show "c \" by fact next from \kv \ S\ show "k \ ?S'" by (auto simp add: image_iff Bex_def) next from \kv \ S\ \S \ map_to_set m\ have "kv \ map_to_set m" by auto thus m_k_eq: "m k = Some v" unfolding map_to_set_def by simp next from \S \ map_to_set m\ show S'_subset: "?S' \ dom m" unfolding map_to_set_dom by auto next show "I (fst ` S) \" by fact next from \S \ map_to_set m\ \kv \ S\ have S_simp: "{(k', v'). k' \ (fst ` S) - {k} \ m k' = Some v'} = S - {kv}" unfolding map_to_set_def subset_iff apply (auto simp add: image_iff Bex_def) apply (metis option.inject) done from R_S[unfolded S_simp[symmetric]] R_not_S show "\k' v'. k' \ fst ` S - {k} \ m k' = Some v' \ R (k, v) (k', v') " by simp next from \S \ map_to_set m\ R_not_S show "\k' v'. k' \ fst ` S \ m k' = Some v' \ R (k', v') (k, v)" apply (simp add: Ball_def map_to_set_def subset_iff image_iff) apply metis done qed moreover from \S \ map_to_set m\ \kv \ S\ have "fst ` (S - {kv}) = fst ` S - {k}" apply (simp add: set_eq_iff image_iff Bex_def map_to_set_def subset_iff) apply (metis option.inject) done ultimately show "I (fst ` (S - {kv})) (f kv \)" by simp qed lemma map_iterator_genord_rule_insert_P: assumes "map_iterator_genord it m R" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \; \k' v'. k' \ (dom m - it) - {k} \ m k' = Some v' \ R (k, v) (k', v'); \k' v'. k' \ it \ m k' = Some v' \ R (k', v') (k, v)\ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \; \k v k' v'. k \ it \ m k = Some v \ k' \ it \ m k' = Some v' \ R (k, v) (k', v') \ \ P \" shows "P (it c f \0)" proof (rule map_iterator_genord_rule_P [of it m R "\S \. I (dom m - S) \"]) show "map_iterator_genord it m R" by fact next show "I (dom m - dom m) \0" using I0 by simp next fix \ assume "I (dom m - {}) \" with IF show "P \" by simp next fix \ it assume assms: "it \ dom m" "it \ {}" "\ c \" "I (dom m - it) \" "\k v k' v'. k \ it \ m k = Some v \ k' \ it \ m k' = Some v' \ R (k, v) (k', v')" from assms have "dom m - it \ dom m" by auto with II[of "dom m - it" \] assms show "P \" apply (simp add: subset_iff dom_def) apply (metis option.simps(2)) done next fix k v it \ assume assms: "c \" "k \ it" "m k = Some v" "it \ dom m" "I (dom m - it) \" "\k' v'. k' \ it - {k} \ m k' = Some v' \ R (k, v) (k', v')" "\k' v'. k' \ it \ m k' = Some v' \ R (k', v') (k, v)" hence "insert k (dom m - it) = (dom m - (it - {k}))" "dom m - (dom m - it) = it" by auto with assms IP[of \ k "dom m - it" v] show "I (dom m - (it - {k})) (f (k, v) \)" by (simp_all add: subset_iff) qed lemma map_iterator_rule_P: assumes "map_iterator it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \ \ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \ \ \ P \" shows "P (it c f \0)" using assms map_iterator_genord_rule_P[of it m "\_ _. True" I \0 c f P] unfolding set_iterator_def by simp lemma map_iterator_rule_insert_P: assumes "map_iterator it m" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \ \ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \ \ \ P \" shows "P (it c f \0)" using assms map_iterator_genord_rule_insert_P[of it m "\_ _. True" I \0 c f P] unfolding set_iterator_def by simp lemma map_iterator_no_cond_rule_P: assumes "map_iterator it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ k \ it; m k = Some v; it \ dom m; I it \ \ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" shows "P (it (\_. True) f \0)" using assms map_iterator_genord_rule_P[of it m "\_ _. True" I \0 "\_. True" f P] unfolding set_iterator_def by simp lemma map_iterator_no_cond_rule_insert_P: assumes "map_iterator it m" and I0: "I {} \0" and IP: "!!k v it \. \ k \ dom m - it; m k = Some v; it \ dom m; I it \ \ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" shows "P (it (\_. True) f \0)" using assms map_iterator_genord_rule_insert_P[of it m "\_ _. True" I \0 "\_. True" f P] unfolding set_iterator_def by simp subsection \Ordered Iterators\ text \Selecting according to a linear order is another case that is interesting. Ordered iterators over maps, i.\,e.\ iterators over key-value pairs, use an order on the keys.\ context linorder begin definition "set_iterator_linord it S0 \ set_iterator_genord it S0 (\)" definition "set_iterator_rev_linord it S0 \ set_iterator_genord it S0 (\)" definition "set_iterator_map_linord it S0 \ set_iterator_genord it S0 (\(k,_) (k',_). k\k')" definition "set_iterator_map_rev_linord it S0 \ set_iterator_genord it S0 (\(k,_) (k',_). k\k')" abbreviation "map_iterator_linord it m \ set_iterator_map_linord it (map_to_set m)" abbreviation "map_iterator_rev_linord it m \ set_iterator_map_rev_linord it (map_to_set m)" lemma set_iterator_linord_rule_P: "\ set_iterator_linord it S0; I S0 \0; !!S \ x. \ c \; x \ S; I S \; S \ S0; \x'. x' \ S0-S \ x' \ x; \x'. x' \ S \ x \ x'\ \ I (S - {x}) (f x \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\x x'. \x \ S; x' \ S0-S\ \ x' \ x) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_linord_rule_insert_P: "\ set_iterator_linord it S0; I {} \0; !!S \ x. \ c \; x \ S0 - S; I S \; S \ S0; \x'. x' \ S \ x' \ x; \x'. x' \ S0 - S \ x \ x'\ \ I (insert x S) (f x \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\x x'. \x \ S0-S; x' \ S\ \ x' \ x) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_rev_linord_rule_P: "\ set_iterator_rev_linord it S0; I S0 \0; !!S \ x. \ c \; x \ S; I S \; S \ S0; \x'. x' \ S0-S \ x \ x'; \x'. x' \ S \ x' \ x\ \ I (S - {x}) (f x \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\x x'. \x \ S; x' \ S0-S\ \ x \ x') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_rev_linord_rule_insert_P: "\ set_iterator_rev_linord it S0; I {} \0; !!S \ x. \ c \; x \ S0 - S; I S \; S \ S0; \x'. x' \ S \ x \ x'; \x'. x' \ S0 - S \ x' \ x\ \ I (insert x S) (f x \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\x x'. \x \ S0-S; x' \ S\ \ x \ x') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\)" I \0 c f P]) apply (simp_all add: Ball_def) apply (metis order_refl) done lemma set_iterator_map_linord_rule_P: "\ set_iterator_map_linord it S0; I S0 \0; !!S \ k v. \ c \; (k, v) \ S; I S \; S \ S0; \k' v'. (k', v') \ S0-S \ k' \ k; \k' v'. (k', v') \ S \ k \ k'\ \ I (S - {(k,v)}) (f (k,v) \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\k v k' v'. \(k, v) \ S0-S; (k', v') \ S\ \ k \ k') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma set_iterator_map_linord_rule_insert_P: "\ set_iterator_map_linord it S0; I {} \0; !!S \ k v. \ c \; (k, v) \ S0 - S; I S \; S \ S0; \k' v'. (k', v') \ S \ k' \ k; \k' v'. (k',v') \ S0 - S \ k \ k'\ \ I (insert (k,v) S) (f (k,v) \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\k v k' v'. \(k, v) \ S; (k', v') \ S0-S\ \ k \ k') \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma set_iterator_map_rev_linord_rule_P: "\ set_iterator_map_rev_linord it S0; I S0 \0; !!S \ k v. \ c \; (k, v) \ S; I S \; S \ S0; \k' v'. (k', v') \ S0-S \ k \ k'; \k' v'. (k', v') \ S \ k' \ k\ \ I (S - {(k,v)}) (f (k,v) \); !!\. I {} \ \ P \; !!\ S. S \ S0 \ S \ {} \ (\k v k' v'. \(k, v) \ S0-S; (k', v') \ S\ \ k' \ k) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma set_iterator_map_rev_linord_rule_insert_P: "\ set_iterator_map_rev_linord it S0; I {} \0; !!S \ k v. \ c \; (k, v) \ S0 - S; I S \; S \ S0; \k' v'. (k', v') \ S \ k \ k'; \k' v'. (k',v') \ S0 - S \ k' \ k\ \ I (insert (k,v) S) (f (k,v) \); !!\. I S0 \ \ P \; !!\ S. S \ S0 \ S \ S0 \ (\k v k' v'. \(k, v) \ S; (k', v') \ S0-S\ \ k' \ k) \ \ c \ \ I S \ \ P \ \ \ P (it c f \0)" unfolding set_iterator_map_rev_linord_def apply (rule set_iterator_genord.iteratei_rule_insert_P [of it S0 "(\(k,_) (k',_). k \ k')" I \0 c f P]) apply simp_all apply (auto simp add: Ball_def) apply (metis order_refl) apply metis done lemma map_iterator_linord_rule_P: assumes "map_iterator_linord it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \; \k'. k' \ it \ k \ k'; \k'. k' \ (dom m)-it \ k' \ k\ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \; \k k'. \k \ (dom m)-it; k' \ it\ \ k \ k'\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_linord_def by (rule map_iterator_genord_rule_P) auto lemma map_iterator_linord_rule_insert_P: assumes "map_iterator_linord it m" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \; \k'. k' \ (dom m - it) \ k \ k'; \k'. k' \ it \ k' \ k \ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \; \k k'. \k \ it; k' \ (dom m)-it\ \ k \ k'\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_linord_def by (rule map_iterator_genord_rule_insert_P) auto lemma map_iterator_rev_linord_rule_P: assumes "map_iterator_rev_linord it m" and I0: "I (dom m) \0" and IP: "!!k v it \. \ c \; k \ it; m k = Some v; it \ dom m; I it \; \k'. k' \ it \ k' \ k; \k'. k' \ (dom m)-it \ k \ k'\ \ I (it - {k}) (f (k, v) \)" and IF: "!!\. I {} \ \ P \" and II: "!!\ it. \ it \ dom m; it \ {}; \ c \; I it \; \k k'. \k \ (dom m)-it; k' \ it\ \ k' \ k\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_rev_linord_def by (rule map_iterator_genord_rule_P) auto lemma map_iterator_rev_linord_rule_insert_P: assumes "map_iterator_rev_linord it m" and I0: "I {} \0" and IP: "!!k v it \. \ c \; k \ dom m - it; m k = Some v; it \ dom m; I it \; \k'. k' \ (dom m - it) \ k' \ k; \k'. k' \ it \ k \ k'\ \ I (insert k it) (f (k, v) \)" and IF: "!!\. I (dom m) \ \ P \" and II: "!!\ it. \ it \ dom m; it \ dom m; \ c \; I it \; \k k'. \k \ it; k' \ (dom m)-it\ \ k' \ k\ \ P \" shows "P (it c f \0)" using assms unfolding set_iterator_map_rev_linord_def by (rule map_iterator_genord_rule_insert_P) auto end subsection \Conversions to foldli\ lemma set_iterator_genord_foldli_conv : "set_iterator_genord iti S R \ (\l0. distinct l0 \ S = set l0 \ sorted_wrt R l0 \ iti = foldli l0)" unfolding set_iterator_genord_def by simp lemma set_iterator_genord_I [intro] : "\distinct l0; S = set l0; sorted_wrt R l0; iti = foldli l0\ \ set_iterator_genord iti S R" unfolding set_iterator_genord_foldli_conv by blast lemma set_iterator_foldli_conv : "set_iterator iti S \ (\l0. distinct l0 \ S = set l0 \ iti = foldli l0)" unfolding set_iterator_def set_iterator_genord_def by simp lemma set_iterator_I [intro] : "\distinct l0; S = set l0; iti = foldli l0\ \ set_iterator iti S" unfolding set_iterator_foldli_conv by blast context linorder begin lemma set_iterator_linord_foldli_conv : "set_iterator_linord iti S \ (\l0. distinct l0 \ S = set l0 \ sorted l0 \ iti = foldli l0)" - unfolding set_iterator_linord_def set_iterator_genord_def by (simp add: sorted_sorted_wrt) + unfolding set_iterator_linord_def set_iterator_genord_def by simp lemma set_iterator_linord_I [intro] : "\distinct l0; S = set l0; sorted l0; iti = foldli l0\ \ set_iterator_linord iti S" unfolding set_iterator_linord_foldli_conv by blast lemma set_iterator_rev_linord_foldli_conv : "set_iterator_rev_linord iti S \ (\l0. distinct l0 \ S = set l0 \ sorted (rev l0) \ iti = foldli l0)" unfolding set_iterator_rev_linord_def set_iterator_genord_def by simp lemma set_iterator_rev_linord_I [intro] : "\distinct l0; S = set l0; sorted (rev l0); iti = foldli l0\ \ set_iterator_rev_linord iti S" unfolding set_iterator_rev_linord_foldli_conv by blast end lemma map_iterator_genord_foldli_conv : "map_iterator_genord iti m R \ (\(l0::('k \ 'v) list). distinct (map fst l0) \ m = map_of l0 \ sorted_wrt R l0 \ iti = foldli l0)" proof - { fix l0 :: "('k \ 'v) list" assume dist: "distinct l0" have "(map_to_set m = set l0) \ (distinct (map fst l0) \ m = map_of l0)" proof (cases "distinct (map fst l0)") case True thus ?thesis by (metis map_of_map_to_set) next case False note not_dist_fst = this with dist have "~(inj_on fst (set l0))" by (simp add: distinct_map) hence "set l0 \ map_to_set m" by (rule_tac notI) (simp add: map_to_set_def inj_on_def) with not_dist_fst show ?thesis by simp qed } thus ?thesis unfolding set_iterator_genord_def distinct_map by metis qed lemma map_iterator_genord_I [intro] : "\distinct (map fst l0); m = map_of l0; sorted_wrt R l0; iti = foldli l0\ \ map_iterator_genord iti m R" unfolding map_iterator_genord_foldli_conv by blast lemma map_iterator_foldli_conv : "map_iterator iti m \ (\l0. distinct (map fst l0) \ m = map_of l0 \ iti = foldli l0)" unfolding set_iterator_def map_iterator_genord_foldli_conv by simp lemma map_iterator_I [intro] : "\distinct (map fst l0); m = map_of l0; iti = foldli l0\ \ map_iterator iti m" unfolding map_iterator_foldli_conv by blast context linorder begin lemma sorted_wrt_keys_map_fst: "sorted_wrt (\(k,_) (k',_). R k k') l = sorted_wrt R (map fst l)" by (induct l) auto lemma map_iterator_linord_foldli_conv : "map_iterator_linord iti m \ (\l0. distinct (map fst l0) \ m = map_of l0 \ sorted (map fst l0) \ iti = foldli l0)" unfolding set_iterator_map_linord_def map_iterator_genord_foldli_conv - by (simp add: sorted_wrt_keys_map_fst sorted_sorted_wrt) + by (simp add: sorted_wrt_keys_map_fst) lemma map_iterator_linord_I [intro] : "\distinct (map fst l0); m = map_of l0; sorted (map fst l0); iti = foldli l0\ \ map_iterator_linord iti m" unfolding map_iterator_linord_foldli_conv by blast lemma map_iterator_rev_linord_foldli_conv : "map_iterator_rev_linord iti m \ (\l0. distinct (map fst l0) \ m = map_of l0 \ sorted (rev (map fst l0)) \ iti = foldli l0)" unfolding set_iterator_map_rev_linord_def map_iterator_genord_foldli_conv by (simp add: sorted_wrt_keys_map_fst) lemma map_iterator_rev_linord_I [intro] : "\distinct (map fst l0); m = map_of l0; sorted (rev (map fst l0)); iti = foldli l0\ \ map_iterator_rev_linord iti m" unfolding map_iterator_rev_linord_foldli_conv by blast end end diff --git a/thys/Containers/RBT_ext.thy b/thys/Containers/RBT_ext.thy --- a/thys/Containers/RBT_ext.thy +++ b/thys/Containers/RBT_ext.thy @@ -1,358 +1,358 @@ (* Title: Containers/RBT_ext.thy Author: Andreas Lochbihler, KIT *) theory RBT_ext imports "HOL-Library.RBT_Impl" Containers_Auxiliary List_Fusion begin section \More on red-black trees\ subsection \More lemmas\ context linorder begin lemma is_rbt_fold_rbt_insert_impl: "is_rbt t \ is_rbt (RBT_Impl.fold rbt_insert t' t)" by(simp add: rbt_insert_def is_rbt_fold_rbt_insertwk) lemma rbt_sorted_fold_insert: "rbt_sorted t \ rbt_sorted (RBT_Impl.fold rbt_insert t' t)" by(induct t' arbitrary: t)(simp_all add: rbt_insert_rbt_sorted) lemma rbt_lookup_rbt_insert': "rbt_sorted t \ rbt_lookup (rbt_insert k v t) = rbt_lookup t(k \ v)" by(simp add: rbt_insert_def rbt_lookup_rbt_insertwk fun_eq_iff split: option.split) lemma rbt_lookup_fold_rbt_insert_impl: "rbt_sorted t2 \ rbt_lookup (RBT_Impl.fold rbt_insert t1 t2) = rbt_lookup t2 ++ map_of (rev (RBT_Impl.entries t1))" proof(induction t1 arbitrary: t2) case Empty thus ?case by simp next case (Branch c l x k r) show ?case using Branch.prems by(simp add: map_add_def Branch.IH rbt_insert_rbt_sorted rbt_sorted_fold_insert rbt_lookup_rbt_insert' fun_eq_iff split: option.split) qed end subsection \Build the cross product of two RBTs\ context fixes f :: "'a \ 'b \ 'c \ 'd \ 'e" begin definition alist_product :: "('a \ 'b) list \ ('c \ 'd) list \ (('a \ 'c) \ 'e) list" where "alist_product xs ys = concat (map (\(a, b). map (\(c, d). ((a, c), f a b c d)) ys) xs)" lemma alist_product_simps [simp]: "alist_product [] ys = []" "alist_product xs [] = []" "alist_product ((a, b) # xs) ys = map (\(c, d). ((a, c), f a b c d)) ys @ alist_product xs ys" by(simp_all add: alist_product_def) lemma append_alist_product_conv_fold: "zs @ alist_product xs ys = rev (fold (\(a, b). fold (\(c, d) rest. ((a, c), f a b c d) # rest) ys) xs (rev zs))" proof(induction xs arbitrary: zs) case Nil thus ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by(cases x) have "\zs. fold (\(c, d). (#) ((a, c), f a b c d)) ys zs = rev (map (\(c, d). ((a, c), f a b c d)) ys) @ zs" by(induct ys) auto with Cons.IH[of "zs @ map (\(c, d). ((a, c), f a b c d)) ys"] x show ?case by simp qed lemma alist_product_code [code]: "alist_product xs ys = rev (fold (\(a, b). fold (\(c, d) rest. ((a, c), f a b c d) # rest) ys) xs [])" using append_alist_product_conv_fold[of "[]" xs ys] by simp lemma set_alist_product: "set (alist_product xs ys) = (\((a, b), (c, d)). ((a, c), f a b c d)) ` (set xs \ set ys)" by(auto 4 3 simp add: alist_product_def intro: rev_image_eqI rev_bexI) lemma distinct_alist_product: "\ distinct (map fst xs); distinct (map fst ys) \ \ distinct (map fst (alist_product xs ys))" proof(induct xs) case Nil thus ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by(cases x) have "distinct (map (fst \ (\(c, d). ((a, c), f a b c d))) ys)" using \distinct (map fst ys)\ by(induct ys)(auto intro: rev_image_eqI) thus ?case using Cons x by(auto simp add: set_alist_product intro: rev_image_eqI) qed lemma map_of_alist_product: "map_of (alist_product xs ys) (a, c) = (case map_of xs a of None \ None | Some b \ map_option (f a b c) (map_of ys c))" proof(induction xs) case Nil thus ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by (cases x) let ?map = "map (\(c, d). ((a, c), f a b c d)) ys" have "map_of ?map (a, c) = map_option (f a b c) (map_of ys c)" by(induct ys) auto moreover { fix a' assume "a' \ a" hence "map_of ?map (a', c) = None" by(induct ys) auto } ultimately show ?case using x Cons.IH by(auto simp add: map_add_def split: option.split) qed definition rbt_product :: "('a, 'b) rbt \ ('c, 'd) rbt \ ('a \ 'c, 'e) rbt" where "rbt_product rbt1 rbt2 = rbtreeify (alist_product (RBT_Impl.entries rbt1) (RBT_Impl.entries rbt2))" lemma rbt_product_code [code]: "rbt_product rbt1 rbt2 = rbtreeify (rev (RBT_Impl.fold (\a b. RBT_Impl.fold (\c d rest. ((a, c), f a b c d) # rest) rbt2) rbt1 []))" unfolding rbt_product_def alist_product_code RBT_Impl.fold_def .. end context fixes leq_a :: "'a \ 'a \ bool" (infix "\\<^sub>a" 50) and less_a :: "'a \ 'a \ bool" (infix "\\<^sub>a" 50) and leq_b :: "'b \ 'b \ bool" (infix "\\<^sub>b" 50) and less_b :: "'b \ 'b \ bool" (infix "\\<^sub>b" 50) assumes lin_a: "class.linorder leq_a less_a" and lin_b: "class.linorder leq_b less_b" begin abbreviation (input) less_eq_prod' :: "('a \ 'b) \ ('a \ 'b) \ bool" (infix "\" 50) where "less_eq_prod' \ less_eq_prod leq_a less_a leq_b" abbreviation (input) less_prod' :: "('a \ 'b) \ ('a \ 'b) \ bool" (infix "\" 50) where "less_prod' \ less_prod leq_a less_a less_b" lemmas linorder_prod = linorder_prod[OF lin_a lin_b] lemma sorted_alist_product: assumes xs: "linorder.sorted leq_a (map fst xs)" "distinct (map fst xs)" and ys: "linorder.sorted (\\<^sub>b) (map fst ys)" shows "linorder.sorted (\) (map fst (alist_product f xs ys))" proof - interpret a: linorder "(\\<^sub>a)" "(\\<^sub>a)" by(fact lin_a) note [simp] = - linorder.sorted.simps(1)[OF linorder_prod] linorder.sorted.simps(2)[OF linorder_prod] + linorder.sorted0[OF linorder_prod] linorder.sorted1[OF linorder_prod] linorder.sorted_append[OF linorder_prod] - linorder.sorted.simps(2)[OF lin_b] + linorder.sorted1[OF lin_b] show ?thesis using xs proof(induction xs) case Nil show ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by(cases x) have "linorder.sorted (\) (map fst (map (\(c, d). ((a, c), f a b c d)) ys))" using ys by(induct ys) auto thus ?case using x Cons by(fastforce simp add: set_alist_product a.not_less dest: bspec a.order_antisym intro: rev_image_eqI) qed qed lemma is_rbt_rbt_product: "\ ord.is_rbt (\\<^sub>a) rbt1; ord.is_rbt (\\<^sub>b) rbt2 \ \ ord.is_rbt (\) (rbt_product f rbt1 rbt2)" unfolding rbt_product_def by(blast intro: linorder.is_rbt_rbtreeify[OF linorder_prod] sorted_alist_product linorder.rbt_sorted_entries[OF lin_a] ord.is_rbt_rbt_sorted linorder.distinct_entries[OF lin_a] linorder.rbt_sorted_entries[OF lin_b] distinct_alist_product linorder.distinct_entries[OF lin_b]) lemma rbt_lookup_rbt_product: "\ ord.is_rbt (\\<^sub>a) rbt1; ord.is_rbt (\\<^sub>b) rbt2 \ \ ord.rbt_lookup (\) (rbt_product f rbt1 rbt2) (a, c) = (case ord.rbt_lookup (\\<^sub>a) rbt1 a of None \ None | Some b \ map_option (f a b c) (ord.rbt_lookup (\\<^sub>b) rbt2 c))" by(simp add: rbt_product_def linorder.rbt_lookup_rbtreeify[OF linorder_prod] linorder.is_rbt_rbtreeify[OF linorder_prod] sorted_alist_product linorder.rbt_sorted_entries[OF lin_a] ord.is_rbt_rbt_sorted linorder.distinct_entries[OF lin_a] linorder.rbt_sorted_entries[OF lin_b] distinct_alist_product linorder.distinct_entries[OF lin_b] map_of_alist_product linorder.map_of_entries[OF lin_a] linorder.map_of_entries[OF lin_b] cong: option.case_cong) end hide_const less_eq_prod' less_prod' subsection \Build an RBT where keys are paired with themselves\ primrec RBT_Impl_diag :: "('a, 'b) rbt \ ('a \ 'a, 'b) rbt" where "RBT_Impl_diag rbt.Empty = rbt.Empty" | "RBT_Impl_diag (rbt.Branch c l k v r) = rbt.Branch c (RBT_Impl_diag l) (k, k) v (RBT_Impl_diag r)" lemma entries_RBT_Impl_diag: "RBT_Impl.entries (RBT_Impl_diag t) = map (\(k, v). ((k, k), v)) (RBT_Impl.entries t)" by(induct t) simp_all lemma keys_RBT_Impl_diag: "RBT_Impl.keys (RBT_Impl_diag t) = map (\k. (k, k)) (RBT_Impl.keys t)" by(simp add: RBT_Impl.keys_def entries_RBT_Impl_diag split_beta) lemma rbt_sorted_RBT_Impl_diag: "ord.rbt_sorted lt t \ ord.rbt_sorted (less_prod leq lt lt) (RBT_Impl_diag t)" by(induct t)(auto simp add: ord.rbt_sorted.simps ord.rbt_less_prop ord.rbt_greater_prop keys_RBT_Impl_diag) lemma bheight_RBT_Impl_diag: "bheight (RBT_Impl_diag t) = bheight t" by(induct t) simp_all lemma inv_RBT_Impl_diag: assumes "inv1 t" "inv2 t" shows "inv1 (RBT_Impl_diag t)" "inv2 (RBT_Impl_diag t)" and "color_of t = color.B \ color_of (RBT_Impl_diag t) = color.B" using assms by(induct t)(auto simp add: bheight_RBT_Impl_diag) lemma is_rbt_RBT_Impl_diag: "ord.is_rbt lt t \ ord.is_rbt (less_prod leq lt lt) (RBT_Impl_diag t)" by(simp add: ord.is_rbt_def rbt_sorted_RBT_Impl_diag inv_RBT_Impl_diag) lemma (in linorder) rbt_lookup_RBT_Impl_diag: "ord.rbt_lookup (less_prod (\) (<) (<)) (RBT_Impl_diag t) = (\(k, k'). if k = k' then ord.rbt_lookup (<) t k else None)" by(induct t)(auto simp add: ord.rbt_lookup.simps fun_eq_iff) subsection \Folding and quantifiers over RBTs\ definition RBT_Impl_fold1 :: "('a \ 'a \ 'a) \ ('a, unit) RBT_Impl.rbt \ 'a" where "RBT_Impl_fold1 f rbt = fold f (tl (RBT_Impl.keys rbt)) (hd (RBT_Impl.keys rbt))" lemma RBT_Impl_fold1_simps [simp, code]: "RBT_Impl_fold1 f rbt.Empty = undefined" "RBT_Impl_fold1 f (Branch c rbt.Empty k v r) = RBT_Impl.fold (\k v. f k) r k" "RBT_Impl_fold1 f (Branch c (Branch c' l' k' v' r') k v r) = RBT_Impl.fold (\k v. f k) r (f k (RBT_Impl_fold1 f (Branch c' l' k' v' r')))" by(simp_all add: RBT_Impl_fold1_def RBT_Impl.keys_def fold_map RBT_Impl.fold_def split_def o_def tl_append hd_def split: list.split) definition RBT_Impl_rbt_all :: "('a \ 'b \ bool) \ ('a, 'b) rbt \ bool" where [code del]: "RBT_Impl_rbt_all P rbt = (\(k, v) \ set (RBT_Impl.entries rbt). P k v)" lemma RBT_Impl_rbt_all_simps [simp, code]: "RBT_Impl_rbt_all P rbt.Empty \ True" "RBT_Impl_rbt_all P (Branch c l k v r) \ P k v \ RBT_Impl_rbt_all P l \ RBT_Impl_rbt_all P r" by(auto simp add: RBT_Impl_rbt_all_def) definition RBT_Impl_rbt_ex :: "('a \ 'b \ bool) \ ('a, 'b) rbt \ bool" where [code del]: "RBT_Impl_rbt_ex P rbt = (\(k, v) \ set (RBT_Impl.entries rbt). P k v)" lemma RBT_Impl_rbt_ex_simps [simp, code]: "RBT_Impl_rbt_ex P rbt.Empty \ False" "RBT_Impl_rbt_ex P (Branch c l k v r) \ P k v \ RBT_Impl_rbt_ex P l \ RBT_Impl_rbt_ex P r" by(auto simp add: RBT_Impl_rbt_ex_def) subsection \List fusion for RBTs\ type_synonym ('a, 'b, 'c) rbt_generator_state = "('c \ ('a, 'b) RBT_Impl.rbt) list \ ('a, 'b) RBT_Impl.rbt" fun rbt_has_next :: "('a, 'b, 'c) rbt_generator_state \ bool" where "rbt_has_next ([], rbt.Empty) = False" | "rbt_has_next _ = True" fun rbt_keys_next :: "('a, 'b, 'a) rbt_generator_state \ 'a \ ('a, 'b, 'a) rbt_generator_state" where "rbt_keys_next ((k, t) # kts, rbt.Empty) = (k, kts, t)" | "rbt_keys_next (kts, rbt.Branch c l k v r) = rbt_keys_next ((k, r) # kts, l)" lemma rbt_generator_induct [case_names empty split shuffle]: assumes "P ([], rbt.Empty)" and "\k t kts. P (kts, t) \ P ((k, t) # kts, rbt.Empty)" and "\kts c l k v r. P ((f k v, r) # kts, l) \ P (kts, Branch c l k v r)" shows "P ktst" using assms apply(induction_schema) apply pat_completeness apply(relation "measure (\(kts, t). size_list (\(k, t). size_rbt (\_. 1) (\_. 1) t) kts + size_rbt (\_. 1) (\_. 1) t)") apply simp_all done lemma terminates_rbt_keys_generator: "terminates (rbt_has_next, rbt_keys_next)" proof fix ktst :: "('a \ ('a, 'b) rbt) list \ ('a, 'b) rbt" show "ktst \ terminates_on (rbt_has_next, rbt_keys_next)" by(induct ktst taking: "\k _. k" rule: rbt_generator_induct)(auto 4 3 intro: terminates_on.intros elim: terminates_on.cases) qed lift_definition rbt_keys_generator :: "('a, ('a, 'b, 'a) rbt_generator_state) generator" is "(rbt_has_next, rbt_keys_next)" by(rule terminates_rbt_keys_generator) definition rbt_init :: "('a, 'b) rbt \ ('a, 'b, 'c) rbt_generator_state" where "rbt_init = Pair []" lemma has_next_rbt_keys_generator [simp]: "list.has_next rbt_keys_generator = rbt_has_next" by transfer simp lemma next_rbt_keys_generator [simp]: "list.next rbt_keys_generator = rbt_keys_next" by transfer simp lemma unfoldr_rbt_keys_generator_aux: "list.unfoldr rbt_keys_generator (kts, t) = RBT_Impl.keys t @ concat (map (\(k, t). k # RBT_Impl.keys t) kts)" proof(induct "(kts, t)" arbitrary: kts t taking: "\k _. k" rule: rbt_generator_induct) case empty thus ?case by(simp add: list.unfoldr.simps) next case split thus ?case by(subst list.unfoldr.simps) simp next case shuffle thus ?case by(subst list.unfoldr.simps)(subst (asm) list.unfoldr.simps, simp) qed corollary unfoldr_rbt_keys_generator: "list.unfoldr rbt_keys_generator (rbt_init t) = RBT_Impl.keys t" by(simp add: unfoldr_rbt_keys_generator_aux rbt_init_def) fun rbt_entries_next :: "('a, 'b, 'a \ 'b) rbt_generator_state \ ('a \ 'b) \ ('a, 'b, 'a \ 'b) rbt_generator_state" where "rbt_entries_next ((kv, t) # kts, rbt.Empty) = (kv, kts, t)" | "rbt_entries_next (kts, rbt.Branch c l k v r) = rbt_entries_next (((k, v), r) # kts, l)" lemma terminates_rbt_entries_generator: "terminates (rbt_has_next, rbt_entries_next)" proof(rule terminatesI) fix ktst :: "('a, 'b, 'a \ 'b) rbt_generator_state" show "ktst \ terminates_on (rbt_has_next, rbt_entries_next)" by(induct ktst taking: Pair rule: rbt_generator_induct)(auto 4 3 intro: terminates_on.intros elim: terminates_on.cases) qed lift_definition rbt_entries_generator :: "('a \ 'b, ('a, 'b, 'a \ 'b) rbt_generator_state) generator" is "(rbt_has_next, rbt_entries_next)" by(rule terminates_rbt_entries_generator) lemma has_next_rbt_entries_generator [simp]: "list.has_next rbt_entries_generator = rbt_has_next" by transfer simp lemma next_rbt_entries_generator [simp]: "list.next rbt_entries_generator = rbt_entries_next" by transfer simp lemma unfoldr_rbt_entries_generator_aux: "list.unfoldr rbt_entries_generator (kts, t) = RBT_Impl.entries t @ concat (map (\(k, t). k # RBT_Impl.entries t) kts)" proof(induct "(kts, t)" arbitrary: kts t taking: Pair rule: rbt_generator_induct) case empty thus ?case by(simp add: list.unfoldr.simps) next case split thus ?case by(subst list.unfoldr.simps) simp next case shuffle thus ?case by(subst list.unfoldr.simps)(subst (asm) list.unfoldr.simps, simp) qed corollary unfoldr_rbt_entries_generator: "list.unfoldr rbt_entries_generator (rbt_init t) = RBT_Impl.entries t" by(simp add: unfoldr_rbt_entries_generator_aux rbt_init_def) end diff --git a/thys/Containers/Set_Linorder.thy b/thys/Containers/Set_Linorder.thy --- a/thys/Containers/Set_Linorder.thy +++ b/thys/Containers/Set_Linorder.thy @@ -1,2481 +1,2481 @@ (* Title: Containers/Set_Linorder.thy Author: Andreas Lochbihler, KIT *) theory Set_Linorder imports Containers_Auxiliary Lexicographic_Order Extend_Partial_Order "HOL-Library.Cardinality" begin section \An executable linear order on sets\ subsection \Definition of the linear order\ subsubsection \Extending finite and cofinite sets\ text \ Partition sets into finite and cofinite sets and distribute the rest arbitrarily such that complement switches between the two. \ consts infinite_complement_partition :: "'a set set" specification (infinite_complement_partition) finite_complement_partition: "finite (A :: 'a set) \ A \ infinite_complement_partition" complement_partition: "\ finite (UNIV :: 'a set) \ (A :: 'a set) \ infinite_complement_partition \ - A \ infinite_complement_partition" proof(cases "finite (UNIV :: 'a set)") case False define Field_r where "Field_r = {\

:: 'a set set. (\C \ \

. - C \ \

) \ (\A. finite A \ A \ \

)}" define r where "r = {(\

1, \

2). \

1 \ \

2 \ \

1 \ Field_r \ \

2 \ Field_r}" have in_Field_r [simp]: "\\

. \

\ Field_r \ (\C \ \

. - C \ \

) \ (\A. finite A \ A \ \

)" unfolding Field_r_def by simp have in_r [simp]: "\\

1 \

2. (\

1, \

2) \ r \ \

1 \ \

2 \ \

1 \ Field_r \ \

2 \ Field_r" unfolding r_def by simp have Field_r [simp]: "Field r = Field_r" by(auto simp add: Field_def Field_r_def) have "Partial_order r" by(auto simp add: Field_def r_def partial_order_on_def preorder_on_def intro!: refl_onI transI antisymI) moreover have "\\ \ Field r. \\ \ \. (\, \) \ r" if \: "\ \ Chains r" for \ proof - let ?\ = "\\ \ {A. finite A}" have *: "?\ \ Field r" using False \ by clarsimp(safe, drule (2) ChainsD, auto 4 4 dest: Chains_Field) hence "\\. \ \ \ \ (\, ?\) \ r" using \ by(auto simp del: in_Field_r dest: Chains_Field) with * show "\\ \ Field r. \\ \ \. (\, \) \ r" by blast qed ultimately have "\\

\ Field r. \\ \ Field r. (\

, \) \ r \ \ = \

" by(rule Zorns_po_lemma) then obtain \

where \

: "\

\ Field r" and max: "\\. \ \ \ Field r; (\

, \) \ r \ \ \ = \

" by blast have "\A. finite A \ A \ \

" using \

by simp moreover { fix C have "C \ \

\ - C \ \

" proof(rule ccontr) assume "\ ?thesis" hence C: "C \ \

" "- C \ \

" by simp_all let ?\ = "insert C \

" have *: "?\ \ Field r" using C \

by auto hence "(\

, ?\) \ r" using \

by auto with * have "?\ = \

" by(rule max) thus False using C by auto qed hence "C \ \

\ - C \ \

" using \

by auto } ultimately have "\\

:: 'a set set. (\A. finite A \ A \ \

) \ (\C. C \ \

\ - C \ \

)" by blast thus ?thesis by metis qed auto lemma not_in_complement_partition: "\ finite (UNIV :: 'a set) \ (A :: 'a set) \ infinite_complement_partition \ - A \ infinite_complement_partition" by(metis complement_partition) lemma not_in_complement_partition_False: "\ (A :: 'a set) \ infinite_complement_partition; \ finite (UNIV :: 'a set) \ \ - A \ infinite_complement_partition = False" by(simp add: not_in_complement_partition) lemma infinite_complement_partition_finite [simp]: "finite (UNIV :: 'a set) \ infinite_complement_partition = (UNIV :: 'a set set)" by(auto intro: finite_subset finite_complement_partition) lemma Compl_eq_empty_iff: "- A = {} \ A = UNIV" by auto subsubsection \A lexicographic-style order on finite subsets\ context ord begin definition set_less_aux :: "'a set \ 'a set \ bool" (infix "\''" 50) where "A \' B \ finite A \ finite B \ (\y \ B - A. \z \ (A - B) \ (B - A). y \ z \ (z \ y \ y = z))" definition set_less_eq_aux :: "'a set \ 'a set \ bool" (infix "\''" 50) where "A \' B \ A \ infinite_complement_partition \ A = B \ A \' B" lemma set_less_aux_irrefl [iff]: "\ A \' A" by(auto simp add: set_less_aux_def) lemma set_less_eq_aux_refl [iff]: "A \' A \ A \ infinite_complement_partition" by(simp add: set_less_eq_aux_def) lemma set_less_aux_empty [simp]: "\ A \' {}" by(auto simp add: set_less_aux_def intro: finite_subset finite_complement_partition) lemma set_less_eq_aux_empty [simp]: "A \' {} \ A = {}" by(auto simp add: set_less_eq_aux_def finite_complement_partition) lemma set_less_aux_antisym: "\ A \' B; B \' A \ \ False" by(auto simp add: set_less_aux_def split: if_split_asm) lemma set_less_aux_conv_set_less_eq_aux: "A \' B \ A \' B \ \ B \' A" by(auto simp add: set_less_eq_aux_def dest: set_less_aux_antisym) lemma set_less_eq_aux_antisym: "\ A \' B; B \' A \ \ A = B" by(auto simp add: set_less_eq_aux_def dest: set_less_aux_antisym) lemma set_less_aux_finiteD: "A \' B \ finite A \ B \ infinite_complement_partition" by(auto simp add: set_less_aux_def finite_complement_partition) lemma set_less_eq_aux_infinite_complement_partitionD: "A \' B \ A \ infinite_complement_partition \ B \ infinite_complement_partition" by(auto simp add: set_less_eq_aux_def dest: set_less_aux_finiteD intro: finite_complement_partition) lemma Compl_set_less_aux_Compl: "finite (UNIV :: 'a set) \ - A \' - B \ B \' A" by(auto simp add: set_less_aux_def finite_subset[OF subset_UNIV]) lemma Compl_set_less_eq_aux_Compl: "finite (UNIV :: 'a set) \ - A \' - B \ B \' A" by(auto simp add: set_less_eq_aux_def Compl_set_less_aux_Compl) lemma set_less_aux_insert_same: "x \ A \ x \ B \ insert x A \' insert x B \ A \' B" by(auto simp add: set_less_aux_def) lemma set_less_eq_aux_insert_same: "\ A \ infinite_complement_partition; insert x B \ infinite_complement_partition; x \ A \ x \ B \ \ insert x A \' insert x B \ A \' B" by(auto simp add: set_less_eq_aux_def set_less_aux_insert_same) end context order begin lemma set_less_aux_singleton_iff: "A \' {x} \ finite A \ (\a\A. x < a)" by(auto simp add: set_less_aux_def less_le Bex_def) end context linorder begin lemma wlog_le [case_names sym le]: assumes "\a b. P a b \ P b a" and "\a b. a \ b \ P a b" shows "P b a" by (metis assms linear) lemma empty_set_less_aux [simp]: "{} \' A \ A \ {} \ finite A" by(auto 4 3 simp add: set_less_aux_def intro!: Min_eqI intro: bexI[where x="Min A"] order_trans[where y="Min A"] Min_in) lemma empty_set_less_eq_aux [simp]: "{} \' A \ finite A" by(auto simp add: set_less_eq_aux_def finite_complement_partition) lemma set_less_aux_trans: assumes AB: "A \' B" and BC: "B \' C" shows "A \' C" proof - from AB BC have A: "finite A" and B: "finite B" and C: "finite C" by(auto simp add: set_less_aux_def) from AB A B obtain ab where ab: "ab \ B - A" and minAB: "\x. \ x \ A; x \ B \ \ ab \ x \ (x \ ab \ ab = x)" and minBA: "\x. \ x \ B; x \ A \ \ ab \ x \ (x \ ab \ ab = x)" unfolding set_less_aux_def by auto from BC B C obtain bc where bc: "bc \ C - B" and minBC: "\x. \ x \ B; x \ C \ \ bc \ x \ (x \ bc \ bc = x)" and minCB: "\x. \ x \ C; x \ B \ \ bc \ x \ (x \ bc \ bc = x)" unfolding set_less_aux_def by auto show ?thesis proof(cases "ab \ bc") case True hence "ab \ C - A" "ab \ A - C" using ab bc by(auto dest: minBC antisym) moreover { fix x assume x: "x \ (C - A) \ (A - C)" hence "ab \ x" by(cases "x \ B")(auto dest: minAB minBA minBC minCB intro: order_trans[OF True]) moreover hence "ab \ x \ \ x \ ab" using ab bc x by(cases "x \ B")(auto dest: antisym) moreover note calculation } ultimately show ?thesis using A C unfolding set_less_aux_def by auto next case False with linear[of ab bc] have "bc \ ab" by simp with ab bc have "bc \ C - A" "bc \ A - C" by(auto dest: minAB antisym) moreover { fix x assume x: "x \ (C - A) \ (A - C)" hence "bc \ x" by(cases "x \ B")(auto dest: minAB minBA minBC minCB intro: order_trans[OF \bc \ ab\]) moreover hence "bc \ x \ \ x \ bc" using ab bc x by(cases "x \ B")(auto dest: antisym) moreover note calculation } ultimately show ?thesis using A C unfolding set_less_aux_def by auto qed qed lemma set_less_eq_aux_trans [trans]: "\ A \' B; B \' C \ \ A \' C" by(auto simp add: set_less_eq_aux_def dest: set_less_aux_trans) lemma set_less_trans_set_less_eq [trans]: "\ A \' B; B \' C \ \ A \' C" by(auto simp add: set_less_eq_aux_def dest: set_less_aux_trans) lemma set_less_eq_aux_porder: "partial_order_on infinite_complement_partition {(A, B). A \' B}" by(auto simp add: preorder_on_def partial_order_on_def intro!: refl_onI transI antisymI dest: set_less_eq_aux_infinite_complement_partitionD intro: set_less_eq_aux_antisym set_less_eq_aux_trans del: equalityI) lemma psubset_finite_imp_set_less_aux: assumes AsB: "A \ B" and B: "finite B" shows "A \' B" proof - from AsB B have A: "finite A" by(auto intro: finite_subset) moreover from AsB B have "Min (B - A) \ B - A" by - (rule Min_in, auto) ultimately show ?thesis using B AsB by(auto simp add: set_less_aux_def intro!: rev_bexI[where x="Min (B - A)"] Min_eqI dest: Min_ge_iff[THEN iffD1, rotated 2]) qed lemma subset_finite_imp_set_less_eq_aux: "\ A \ B; finite B \ \ A \' B" by(cases "A = B")(auto simp add: set_less_eq_aux_def finite_complement_partition intro: psubset_finite_imp_set_less_aux) lemma empty_set_less_aux_finite_iff: "finite A \ {} \' A \ A \ {}" by(auto intro: psubset_finite_imp_set_less_aux) lemma set_less_aux_finite_total: assumes A: "finite A" and B: "finite B" shows "A \' B \ A = B \ B \' A" proof(cases "A \ B \ B \ A") case True thus ?thesis using A B psubset_finite_imp_set_less_aux[of A B] psubset_finite_imp_set_less_aux[of B A] by auto next case False hence A': "\ A \ B" and B': "\ B \ A" and AnB: "A \ B" by auto thus ?thesis using A B proof(induct "Min (B - A)" "Min (A - B)" arbitrary: A B rule: wlog_le) case (sym m n) from sym.hyps[OF refl refl] sym.prems show ?case by blast next case (le A B) note A = \finite A\ and B = \finite B\ and A' = \\ A \ B\ and B' = \\ B \ A\ { fix z assume z: "z \ (A - B) \ (B - A)" hence "Min (B - A) \ z \ (z \ Min (B - A) \ Min (B - A) = z)" proof assume "z \ B - A" hence "Min (B - A) \ z" by(simp add: B) thus ?thesis by auto next assume "z \ A - B" hence "Min (A - B) \ z" by(simp add: A) with le.hyps show ?thesis by(auto) qed } moreover have "Min (B - A) \ B - A" by(rule Min_in)(simp_all add: B B') ultimately have "A \' B" using A B by(auto simp add: set_less_aux_def) thus ?case .. qed qed lemma set_less_eq_aux_finite_total: "\ finite A; finite B \ \ A \' B \ A = B \ B \' A" by(drule (1) set_less_aux_finite_total)(auto simp add: set_less_eq_aux_def) lemma set_less_eq_aux_finite_total2: "\ finite A; finite B \ \ A \' B \ B \' A" by(drule (1) set_less_eq_aux_finite_total)(auto simp add: finite_complement_partition) lemma set_less_aux_rec: assumes A: "finite A" and B: "finite B" and A': "A \ {}" and B': "B \ {}" shows "A \' B \ Min B < Min A \ Min A = Min B \ A - {Min A} \' B - {Min A}" proof(cases "Min A = Min B") case True from A A' B B' have "insert (Min A) A = A" "insert (Min B) B = B" by(auto simp add: ex_in_conv[symmetric] exI) with True show ?thesis by(subst (2) set_less_aux_insert_same[symmetric, where x="Min A"]) simp_all next case False have "A \' B \ Min B < Min A" proof assume AB: "A \' B" with B A obtain ab where ab: "ab \ B - A" and AB: "\x. \ x \ A; x \ B \ \ ab \ x" by(auto simp add: set_less_aux_def) { fix a assume "a \ A" hence "Min B \ a" using A A' B B' ab by(cases "a \ B")(auto intro: order_trans[where y=ab] dest: AB) } hence "Min B \ Min A" using A A' by simp with False show "Min B < Min A" using False by auto next assume "Min B < Min A" hence "\z\A - B \ (B - A). Min B \ z \ (z \ Min B \ Min B = z)" using A B A' B' by(auto 4 4 intro: Min_in Min_eqI dest: bspec bspec[where x="Min B"]) moreover have "Min B \ A" using \Min B < Min A\ by (metis A Min_le not_less) ultimately show "A \' B" using A B A' B' by(simp add: set_less_aux_def bexI[where x="Min B"]) qed thus ?thesis using False by simp qed lemma set_less_eq_aux_rec: assumes "finite A" "finite B" "A \ {}" "B \ {}" shows "A \' B \ Min B < Min A \ Min A = Min B \ A - {Min A} \' B - {Min A}" proof(cases "A = B") case True thus ?thesis using assms by(simp add: finite_complement_partition) next case False moreover hence "Min A = Min B \ A - {Min A} \ B - {Min B}" by (metis (lifting) assms Min_in insert_Diff) ultimately show ?thesis using set_less_aux_rec[OF assms] by(simp add: set_less_eq_aux_def cong: conj_cong) qed lemma set_less_aux_Min_antimono: "\ Min A < Min B; finite A; finite B; A \ {} \ \ B \' A" using set_less_aux_rec[of B A] by(cases "B = {}")(simp_all add: empty_set_less_aux_finite_iff) lemma sorted_Cons_Min: "sorted (x # xs) \ Min (insert x (set xs)) = x" by(auto simp add: intro: Min_eqI) lemma set_less_aux_code: "\ sorted xs; distinct xs; sorted ys; distinct ys \ \ set xs \' set ys \ ord.lexordp (>) xs ys" apply(induct xs ys rule: list_induct2') apply(simp_all add: empty_set_less_aux_finite_iff sorted_Cons_Min set_less_aux_rec neq_Nil_conv) apply(auto cong: conj_cong) done lemma set_less_eq_aux_code: assumes "sorted xs" "distinct xs" "sorted ys" "distinct ys" shows "set xs \' set ys \ ord.lexordp_eq (>) xs ys" proof - have dual: "class.linorder (\) (>)" by(rule linorder.dual_linorder) unfold_locales from assms show ?thesis by(auto simp add: set_less_eq_aux_def finite_complement_partition linorder.lexordp_eq_conv_lexord[OF dual] set_less_aux_code intro: sorted_distinct_set_unique) qed end subsubsection \Extending @{term set_less_eq_aux} to have @{term "{}"} as least element\ context ord begin definition set_less_eq_aux' :: "'a set \ 'a set \ bool" (infix "\''''" 50) where "A \'' B \ A \' B \ A = {} \ B \ infinite_complement_partition" lemma set_less_eq_aux'_refl: "A \'' A \ A \ infinite_complement_partition" by(auto simp add: set_less_eq_aux'_def) lemma set_less_eq_aux'_antisym: "\ A \'' B; B \'' A \ \ A = B" by(auto simp add: set_less_eq_aux'_def intro: set_less_eq_aux_antisym del: equalityI) lemma set_less_eq_aux'_infinite_complement_partitionD: "A \'' B \ A \ infinite_complement_partition \ B \ infinite_complement_partition" by(auto simp add: set_less_eq_aux'_def intro: finite_complement_partition dest: set_less_eq_aux_infinite_complement_partitionD) lemma empty_set_less_eq_def [simp]: "{} \'' B \ B \ infinite_complement_partition" by(auto simp add: set_less_eq_aux'_def dest: set_less_eq_aux_infinite_complement_partitionD) end context linorder begin lemma set_less_eq_aux'_trans: "\ A \'' B; B \'' C \ \ A \'' C" by(auto simp add: set_less_eq_aux'_def del: equalityI intro: set_less_eq_aux_trans dest: set_less_eq_aux_infinite_complement_partitionD) lemma set_less_eq_aux'_porder: "partial_order_on infinite_complement_partition {(A, B). A \'' B}" by(auto simp add: partial_order_on_def preorder_on_def intro!: refl_onI transI antisymI dest: set_less_eq_aux'_antisym set_less_eq_aux'_infinite_complement_partitionD simp add: set_less_eq_aux'_refl intro: set_less_eq_aux'_trans) end subsubsection \Extend @{term set_less_eq_aux'} to a total order on @{term infinite_complement_partition}\ context ord begin definition set_less_eq_aux'' :: "'a set \ 'a set \ bool" (infix "\''''''" 50) where "set_less_eq_aux'' = (SOME sleq. (linear_order_on UNIV {(a, b). a \ b} \ linear_order_on infinite_complement_partition {(A, B). sleq A B}) \ order_consistent {(A, B). A \'' B} {(A, B). sleq A B})" lemma set_less_eq_aux''_spec: shows "linear_order {(a, b). a \ b} \ linear_order_on infinite_complement_partition {(A, B). A \''' B}" (is "PROP ?thesis1") and "order_consistent {(A, B). A \'' B} {(A, B). A \''' B}" (is ?thesis2) proof - let ?P = "\sleq. (linear_order {(a, b). a \ b} \ linear_order_on infinite_complement_partition {(A, B). sleq A B}) \ order_consistent {(A, B). A \'' B} {(A, B). sleq A B}" have "Ex ?P" proof(cases "linear_order {(a, b). a \ b}") case False have "antisym {(a, b). a \'' b}" by (rule antisymI) (simp add: set_less_eq_aux'_antisym) then show ?thesis using False by (auto intro: antisym_order_consistent_self) next case True hence "partial_order_on infinite_complement_partition {(A, B). A \'' B}" by(rule linorder.set_less_eq_aux'_porder[OF linear_order_imp_linorder]) then obtain s where "linear_order_on infinite_complement_partition s" and "order_consistent {(A, B). A \'' B} s" by(rule porder_extend_to_linorder) thus ?thesis by(auto intro: exI[where x="\A B. (A, B) \ s"]) qed hence "?P (Eps ?P)" by(rule someI_ex) thus "PROP ?thesis1" ?thesis2 by(simp_all add: set_less_eq_aux''_def) qed end context linorder begin lemma set_less_eq_aux''_linear_order: "linear_order_on infinite_complement_partition {(A, B). A \''' B}" by(rule set_less_eq_aux''_spec)(rule linear_order) lemma set_less_eq_aux''_refl [iff]: "A \''' A \ A \ infinite_complement_partition" using set_less_eq_aux''_linear_order by(auto simp add: linear_order_on_def partial_order_on_def preorder_on_def dest: refl_onD refl_onD1) lemma set_less_eq_aux'_into_set_less_eq_aux'': assumes "A \'' B" shows "A \''' B" proof(rule ccontr) assume nleq: "\ ?thesis" moreover from assms have A: "A \ infinite_complement_partition" and B: "B \ infinite_complement_partition" by(auto dest: set_less_eq_aux'_infinite_complement_partitionD) with set_less_eq_aux''_linear_order have "A \''' B \ A = B \ B \''' A" by(auto simp add: linear_order_on_def dest: total_onD) ultimately have "B \''' A" using B by auto with assms have "A = B" using set_less_eq_aux''_spec(2) by(simp add: order_consistent_def) with A nleq show False by simp qed lemma finite_set_less_eq_aux''_finite: assumes "finite A" and "finite B" shows "A \''' B \ A \'' B" proof assume "A \''' B" from assms have "A \' B \ B \' A" by(rule set_less_eq_aux_finite_total2) hence "A \'' B \ B \'' A" by(auto simp add: set_less_eq_aux'_def) thus "A \'' B" proof assume "B \'' A" hence "B \''' A" by(rule set_less_eq_aux'_into_set_less_eq_aux'') with \A \''' B\ set_less_eq_aux''_linear_order have "A = B" by(auto simp add: linear_order_on_def partial_order_on_def dest: antisymD) thus ?thesis using assms by(simp add: finite_complement_partition set_less_eq_aux'_def) qed qed(rule set_less_eq_aux'_into_set_less_eq_aux'') lemma set_less_eq_aux''_finite: "finite (UNIV :: 'a set) \ set_less_eq_aux'' = set_less_eq_aux" by(auto simp add: fun_eq_iff finite_set_less_eq_aux''_finite set_less_eq_aux'_def finite_subset[OF subset_UNIV]) lemma set_less_eq_aux''_antisym: "\ A \''' B; B \''' A; A \ infinite_complement_partition; B \ infinite_complement_partition \ \ A = B" using set_less_eq_aux''_linear_order by(auto simp add: linear_order_on_def partial_order_on_def dest: antisymD del: equalityI) lemma set_less_eq_aux''_trans: "\ A \''' B; B \''' C \ \ A \''' C" using set_less_eq_aux''_linear_order by(auto simp add: linear_order_on_def partial_order_on_def preorder_on_def dest: transD) lemma set_less_eq_aux''_total: "\ A \ infinite_complement_partition; B \ infinite_complement_partition \ \ A \''' B \ B \''' A" using set_less_eq_aux''_linear_order by(auto simp add: linear_order_on_def dest: total_onD) end subsubsection \Extend @{term set_less_eq_aux''} to cofinite sets\ context ord begin definition set_less_eq :: "'a set \ 'a set \ bool" (infix "\" 50) where "A \ B \ (if A \ infinite_complement_partition then A \''' B \ B \ infinite_complement_partition else B \ infinite_complement_partition \ - B \''' - A)" definition set_less :: "'a set \ 'a set \ bool" (infix "\" 50) where "A \ B \ A \ B \ \ B \ A" lemma set_less_eq_def2: "A \ B \ (if finite (UNIV :: 'a set) then A \''' B else if A \ infinite_complement_partition then A \''' B \ B \ infinite_complement_partition else B \ infinite_complement_partition \ - B \''' - A)" by(simp add: set_less_eq_def) end context linorder begin lemma set_less_eq_refl [iff]: "A \ A" by(auto simp add: set_less_eq_def2 not_in_complement_partition) lemma set_less_eq_antisym: "\ A \ B; B \ A \ \ A = B" by(auto simp add: set_less_eq_def2 set_less_eq_aux''_finite not_in_complement_partition not_in_complement_partition_False del: equalityI split: if_split_asm dest: set_less_eq_aux_antisym set_less_eq_aux''_antisym) lemma set_less_eq_trans: "\ A \ B; B \ C \ \ A \ C" by(auto simp add: set_less_eq_def split: if_split_asm intro: set_less_eq_aux''_trans) lemma set_less_eq_total: "A \ B \ B \ A" by(auto simp add: set_less_eq_def2 set_less_eq_aux''_finite not_in_complement_partition not_in_complement_partition_False intro: set_less_eq_aux_finite_total2 finite_subset[OF subset_UNIV] del: disjCI dest: set_less_eq_aux''_total) lemma set_less_eq_linorder: "class.linorder (\) (\)" by(unfold_locales)(auto simp add: set_less_def set_less_eq_antisym set_less_eq_total intro: set_less_eq_trans) lemma set_less_eq_conv_set_less: "set_less_eq A B \ A = B \ set_less A B" by(auto simp add: set_less_def del: equalityI dest: set_less_eq_antisym) lemma Compl_set_less_eq_Compl: "- A \ - B \ B \ A" by(auto simp add: set_less_eq_def2 not_in_complement_partition_False not_in_complement_partition set_less_eq_aux''_finite Compl_set_less_eq_aux_Compl) lemma Compl_set_less_Compl: "- A \ - B \ B \ A" by(simp add: set_less_def Compl_set_less_eq_Compl) lemma set_less_eq_finite_iff: "\ finite A; finite B \ \ A \ B \ A \' B" by(auto simp add: set_less_eq_def finite_complement_partition set_less_eq_aux'_def finite_set_less_eq_aux''_finite) lemma set_less_finite_iff: "\ finite A; finite B \ \ A \ B \ A \' B" by(simp add: set_less_def set_less_aux_conv_set_less_eq_aux set_less_eq_finite_iff) lemma infinite_set_less_eq_Complement: "\ finite A; finite B; \ finite (UNIV :: 'a set) \ \ A \ - B" by(simp add: set_less_eq_def finite_complement_partition not_in_complement_partition) lemma infinite_set_less_Complement: "\ finite A; finite B; \ finite (UNIV :: 'a set) \ \ A \ - B" by(auto simp add: set_less_def dest: set_less_eq_antisym intro: infinite_set_less_eq_Complement) lemma infinite_Complement_set_less_eq: "\ finite A; finite B; \ finite (UNIV :: 'a set) \ \ \ - A \ B" using infinite_set_less_eq_Complement[of A B] Compl_set_less_eq_Compl[of A "- B"] by(auto dest: set_less_eq_antisym) lemma infinite_Complement_set_less: "\ finite A; finite B; \ finite (UNIV :: 'a set) \ \ \ - A \ B" using infinite_Complement_set_less_eq[of A B] by(simp add: set_less_def) lemma empty_set_less_eq [iff]: "{} \ A" by(auto simp add: set_less_eq_def finite_complement_partition intro: set_less_eq_aux'_into_set_less_eq_aux'') lemma set_less_eq_empty [iff]: "A \ {} \ A = {}" by (metis empty_set_less_eq set_less_eq_antisym) lemma empty_set_less_iff [iff]: "{} \ A \ A \ {}" by(simp add: set_less_def) lemma not_set_less_empty [simp]: "\ A \ {}" by(simp add: set_less_def) lemma set_less_eq_UNIV [iff]: "A \ UNIV" using Compl_set_less_eq_Compl[of "- A" "{}"] by simp lemma UNIV_set_less_eq [iff]: "UNIV \ A \ A = UNIV" using Compl_set_less_eq_Compl[of "{}" "- A"] by(simp add: Compl_eq_empty_iff) lemma set_less_UNIV_iff [iff]: "A \ UNIV \ A \ UNIV" by(simp add: set_less_def) lemma not_UNIV_set_less [simp]: "\ UNIV \ A" by(simp add: set_less_def) end subsection \Implementation based on sorted lists\ type_synonym 'a proper_interval = "'a option \ 'a option \ bool" class proper_intrvl = ord + fixes proper_interval :: "'a proper_interval" class proper_interval = proper_intrvl + assumes proper_interval_simps: "proper_interval None None = True" "proper_interval None (Some y) = (\z. z < y)" "proper_interval (Some x) None = (\z. x < z)" "proper_interval (Some x) (Some y) = (\z. x < z \ z < y)" context proper_intrvl begin function set_less_eq_aux_Compl :: "'a option \ 'a list \ 'a list \ bool" where "set_less_eq_aux_Compl ao [] ys \ True" | "set_less_eq_aux_Compl ao xs [] \ True" | "set_less_eq_aux_Compl ao (x # xs) (y # ys) \ (if x < y then proper_interval ao (Some x) \ set_less_eq_aux_Compl (Some x) xs (y # ys) else if y < x then proper_interval ao (Some y) \ set_less_eq_aux_Compl (Some y) (x # xs) ys else proper_interval ao (Some y))" by(pat_completeness) simp_all termination by(lexicographic_order) fun Compl_set_less_eq_aux :: "'a option \ 'a list \ 'a list \ bool" where "Compl_set_less_eq_aux ao [] [] \ \ proper_interval ao None" | "Compl_set_less_eq_aux ao [] (y # ys) \ \ proper_interval ao (Some y) \ Compl_set_less_eq_aux (Some y) [] ys" | "Compl_set_less_eq_aux ao (x # xs) [] \ \ proper_interval ao (Some x) \ Compl_set_less_eq_aux (Some x) xs []" | "Compl_set_less_eq_aux ao (x # xs) (y # ys) \ (if x < y then \ proper_interval ao (Some x) \ Compl_set_less_eq_aux (Some x) xs (y # ys) else if y < x then \ proper_interval ao (Some y) \ Compl_set_less_eq_aux (Some y) (x # xs) ys else \ proper_interval ao (Some y))" fun set_less_aux_Compl :: "'a option \ 'a list \ 'a list \ bool" where "set_less_aux_Compl ao [] [] \ proper_interval ao None" | "set_less_aux_Compl ao [] (y # ys) \ proper_interval ao (Some y) \ set_less_aux_Compl (Some y) [] ys" | "set_less_aux_Compl ao (x # xs) [] \ proper_interval ao (Some x) \ set_less_aux_Compl (Some x) xs []" | "set_less_aux_Compl ao (x # xs) (y # ys) \ (if x < y then proper_interval ao (Some x) \ set_less_aux_Compl (Some x) xs (y # ys) else if y < x then proper_interval ao (Some y) \ set_less_aux_Compl (Some y) (x # xs) ys else proper_interval ao (Some y))" function Compl_set_less_aux :: "'a option \ 'a list \ 'a list \ bool" where "Compl_set_less_aux ao [] ys \ False" | "Compl_set_less_aux ao xs [] \ False" | "Compl_set_less_aux ao (x # xs) (y # ys) \ (if x < y then \ proper_interval ao (Some x) \ Compl_set_less_aux (Some x) xs (y # ys) else if y < x then \ proper_interval ao (Some y) \ Compl_set_less_aux (Some y) (x # xs) ys else \ proper_interval ao (Some y))" by pat_completeness simp_all termination by lexicographic_order end lemmas [code] = proper_intrvl.set_less_eq_aux_Compl.simps proper_intrvl.set_less_aux_Compl.simps proper_intrvl.Compl_set_less_eq_aux.simps proper_intrvl.Compl_set_less_aux.simps class linorder_proper_interval = linorder + proper_interval begin theorem assumes fin: "finite (UNIV :: 'a set)" and xs: "sorted xs" "distinct xs" and ys: "sorted ys" "distinct ys" shows set_less_eq_aux_Compl2_conv_set_less_eq_aux_Compl: "set xs \' - set ys \ set_less_eq_aux_Compl None xs ys" (is ?Compl2) and Compl1_set_less_eq_aux_conv_Compl_set_less_eq_aux: "- set xs \' set ys \ Compl_set_less_eq_aux None xs ys" (is ?Compl1) proof - note fin' [simp] = finite_subset[OF subset_UNIV fin] define above where "above = case_option UNIV (Collect \ less)" have above_simps [simp]: "above None = UNIV" "\x. above (Some x) = {y. x < y}" and above_upclosed: "\x y ao. \ x \ above ao; x < y \ \ y \ above ao" and proper_interval_Some2: "\x ao. proper_interval ao (Some x) \ (\z\above ao. z < x)" and proper_interval_None2: "\ao. proper_interval ao None \ above ao \ {}" by(simp_all add: proper_interval_simps above_def split: option.splits) { fix ao x A B assume ex: "proper_interval ao (Some x)" and A: "\a \ A. x \ a" and B: "\b \ B. x \ b" from ex obtain z where z_ao: "z \ above ao" and z: "z < x" by(auto simp add: proper_interval_Some2) with A have A_eq: "A \ above ao = A" by(auto simp add: above_upclosed) from z_ao z B have B_eq: "B \ above ao = B" by(auto simp add: above_upclosed) define w where "w = Min (above ao)" with z_ao have "w \ z" "\z \ above ao. w \ z" "w \ above ao" by(auto simp add: Min_le_iff intro: Min_in) hence "A \ above ao \' (- B) \ above ao" (is "?lhs \' ?rhs") using A B z by(auto simp add: set_less_aux_def intro!: bexI[where x="w"]) hence "A \' ?rhs" unfolding A_eq by(simp add: set_less_eq_aux_def) moreover from z_ao z A B have "z \ - A \ above ao" "z \ B" by auto hence neq: "- A \ above ao \ B \ above ao" by auto have "\ - A \ above ao \' B \ above ao" (is "\ ?lhs' \' ?rhs'") using A B z z_ao by(force simp add: set_less_aux_def not_less dest: bspec[where x=z]) with neq have "\ ?lhs' \' B" unfolding B_eq by(auto simp add: set_less_eq_aux_def) moreover note calculation } note proper_interval_set_less_eqI = this(1) and proper_interval_not_set_less_eq_auxI = this(2) { fix ao assume "set xs \ set ys \ above ao" with xs ys have "set_less_eq_aux_Compl ao xs ys \ set xs \' (- set ys) \ above ao" proof(induction ao xs ys rule: set_less_eq_aux_Compl.induct) case 1 thus ?case by simp next case 2 thus ?case by(auto intro: subset_finite_imp_set_less_eq_aux) next case (3 ao x xs y ys) note ao = \set (x # xs) \ set (y # ys) \ above ao\ hence x_ao: "x \ above ao" and y_ao: "y \ above ao" by simp_all note yys = \sorted (y # ys)\ \distinct (y # ys)\ hence ys: "sorted ys" "distinct ys" and y_Min: "\y' \ set ys. y < y'" by(auto simp add: less_le) note xxs = \sorted (x # xs)\ \distinct (x # xs)\ hence xs: "sorted xs" "distinct xs" and x_Min: "\x' \ set xs. x < x'" by(auto simp add: less_le) let ?lhs = "set (x # xs)" and ?rhs = "- set (y # ys) \ above ao" show ?case proof(cases "x < y") case True show ?thesis proof(cases "proper_interval ao (Some x)") case True hence "?lhs \' ?rhs" using x_Min y_Min \x < y\ by(auto intro!: proper_interval_set_less_eqI) with True show ?thesis using \x < y\ by simp next case False have "set xs \ set (y # ys) \ above (Some x)" using True x_Min y_Min by auto with True xs yys have IH: "set_less_eq_aux_Compl (Some x) xs (y # ys) = (set xs \' - set (y # ys) \ above (Some x))" by(rule "3.IH") from True y_Min x_ao have "x \ - set (y # ys) \ above ao" by auto hence "?rhs \ {}" by blast moreover have "Min ?lhs = x" using x_Min x_ao by(auto intro!: Min_eqI) moreover have "Min ?rhs = x" using \x < y\ y_Min x_ao False by(auto intro!: Min_eqI simp add: proper_interval_Some2) moreover have "set xs = set xs - {x}" using ao x_Min by auto moreover have "- set (y # ys) \ above (Some x) = - set (y # ys) \ above ao - {x}" using False x_ao by(auto simp add: proper_interval_Some2 intro: above_upclosed) ultimately show ?thesis using True False IH by(simp del: set_simps)(subst (2) set_less_eq_aux_rec, simp_all add: x_ao) qed next case False show ?thesis proof(cases "y < x") case True show ?thesis proof(cases "proper_interval ao (Some y)") case True hence "?lhs \' ?rhs" using x_Min y_Min \\ x < y\ by(auto intro!: proper_interval_set_less_eqI) with True show ?thesis using \\ x < y\ by simp next case False have "set (x # xs) \ set ys \ above (Some y)" using \y < x\ x_Min y_Min by auto with \\ x < y\ \y < x\ xxs ys have IH: "set_less_eq_aux_Compl (Some y) (x # xs) ys = (set (x # xs) \' - set ys \ above (Some y))" by(rule "3.IH") moreover have "- set ys \ above (Some y) = ?rhs" using y_ao False by(auto intro: above_upclosed simp add: proper_interval_Some2) ultimately show ?thesis using \\ x < y\ True False by simp qed next case False with \\ x < y\ have "x = y" by auto { assume "proper_interval ao (Some y)" hence "?lhs \' ?rhs" using x_Min y_Min \x = y\ by(auto intro!: proper_interval_set_less_eqI) } moreover { assume "?lhs \' ?rhs" moreover have "?lhs \ ?rhs" proof assume eq: "?lhs = ?rhs" have "x \ ?lhs" using x_ao by simp also note eq also note \x = y\ finally show False by simp qed ultimately obtain z where "z \ above ao" "z < y" using \x = y\ y_ao by(fastforce simp add: set_less_eq_aux_def set_less_aux_def not_le dest!: bspec[where x=y]) hence "proper_interval ao (Some y)" by(auto simp add: proper_interval_Some2) } ultimately show ?thesis using \x = y\ \\ x < y\ \\ y < x\ by auto qed qed qed } from this[of None] show ?Compl2 by simp { fix ao assume "set xs \ set ys \ above ao" with xs ys have "Compl_set_less_eq_aux ao xs ys \ (- set xs) \ above ao \' set ys" proof(induction ao xs ys rule: Compl_set_less_eq_aux.induct) case 1 thus ?case by(simp add: proper_interval_None2) next case (2 ao y ys) from \sorted (y # ys)\ \distinct (y # ys)\ have ys: "sorted ys" "distinct ys" and y_Min: "\y' \ set ys. y < y'" by(auto simp add: less_le) show ?case proof(cases "proper_interval ao (Some y)") case True hence "\ - set [] \ above ao \' set (y # ys)" using y_Min by -(erule proper_interval_not_set_less_eq_auxI, auto) thus ?thesis using True by simp next case False note ao = \set [] \ set (y # ys) \ above ao\ hence y_ao: "y \ above ao" by simp from ao y_Min have "set [] \ set ys \ above (Some y)" by auto with \sorted []\ \distinct []\ ys have "Compl_set_less_eq_aux (Some y) [] ys \ - set [] \ above (Some y) \' set ys" by(rule "2.IH") moreover have "above ao \ {}" using y_ao by auto moreover have "Min (above ao) = y" and "Min (set (y # ys)) = y" using y_ao False ao by(auto intro!: Min_eqI simp add: proper_interval_Some2 not_less) moreover have "above ao - {y} = above (Some y)" using False y_ao by(auto simp add: proper_interval_Some2 intro: above_upclosed) moreover have "set ys - {y} = set ys" using y_Min y_ao by(auto) ultimately show ?thesis using False y_ao by(simp)(subst (2) set_less_eq_aux_rec, simp_all) qed next case (3 ao x xs) from \sorted (x # xs)\ \distinct (x # xs)\ have xs: "sorted xs" "distinct xs" and x_Min: "\x'\set xs. x < x'" by(auto simp add: less_le) show ?case proof(cases "proper_interval ao (Some x)") case True then obtain z where "z \ above ao" "z < x" by(auto simp add: proper_interval_Some2) hence "z \ - set (x # xs) \ above ao" using x_Min by auto thus ?thesis using True by auto next case False note ao = \set (x # xs) \ set [] \ above ao\ hence x_ao: "x \ above ao" by simp from ao have "set xs \ set [] \ above (Some x)" using x_Min by auto with xs \sorted []\ \distinct []\ have "Compl_set_less_eq_aux (Some x) xs [] \ - set xs \ above (Some x) \' set []" by(rule "3.IH") moreover have "- set (x # xs) \ above ao = - set xs \ above (Some x)" using False x_ao by(auto simp add: proper_interval_Some2 intro: above_upclosed) ultimately show ?thesis using False by simp qed next case (4 ao x xs y ys) note ao = \set (x # xs) \ set (y # ys) \ above ao\ hence x_ao: "x \ above ao" and y_ao: "y \ above ao" by simp_all note xxs = \sorted (x # xs)\ \distinct (x # xs)\ hence xs: "sorted xs" "distinct xs" and x_Min: "\x'\set xs. x < x'" by(auto simp add: less_le) note yys = \sorted (y # ys)\ \distinct (y # ys)\ hence ys: "sorted ys" "distinct ys" and y_Min: "\y'\set ys. y < y'" by(auto simp add: less_le) let ?lhs = "- set (x # xs) \ above ao" and ?rhs = "set (y # ys)" show ?case proof(cases "x < y") case True show ?thesis proof(cases "proper_interval ao (Some x)") case True hence "\ ?lhs \' ?rhs" using x_Min y_Min \x < y\ by -(erule proper_interval_not_set_less_eq_auxI, auto) thus ?thesis using True \x < y\ by simp next case False have "set xs \ set (y # ys) \ above (Some x)" using ao x_Min y_Min True by auto with True xs yys have "Compl_set_less_eq_aux (Some x) xs (y # ys) \ - set xs \ above (Some x) \' set (y # ys)" by(rule "4.IH") moreover have "- set xs \ above (Some x) = ?lhs" using x_ao False by(auto intro: above_upclosed simp add: proper_interval_Some2) ultimately show ?thesis using False True by simp qed next case False show ?thesis proof(cases "y < x") case True show ?thesis proof(cases "proper_interval ao (Some y)") case True hence "\ ?lhs \' ?rhs" using x_Min y_Min \y < x\ by -(erule proper_interval_not_set_less_eq_auxI, auto) thus ?thesis using True \y < x\ \\ x < y\ by simp next case False from ao True x_Min y_Min have "set (x # xs) \ set ys \ above (Some y)" by auto with \\ x < y\ True xxs ys have "Compl_set_less_eq_aux (Some y) (x # xs) ys \ - set (x # xs) \ above (Some y) \' set ys" by(rule "4.IH") moreover have "y \ ?lhs" using True x_Min y_ao by auto hence "?lhs \ {}" by auto moreover have "Min ?lhs = y" using True False x_Min y_ao by(auto intro!: Min_eqI simp add: not_le not_less proper_interval_Some2) moreover have "Min ?rhs = y" using y_Min y_ao by(auto intro!: Min_eqI) moreover have "- set (x # xs) \ above (Some y) = ?lhs - {y}" using y_ao False by(auto intro: above_upclosed simp add: proper_interval_Some2) moreover have "set ys = set ys - {y}" using y_ao y_Min by(auto intro: above_upclosed) ultimately show ?thesis using True False \\ x < y\ y_ao by(simp)(subst (2) set_less_eq_aux_rec, simp_all) qed next case False with \\ x < y\ have "x = y" by auto { assume "proper_interval ao (Some y)" hence "\ ?lhs \' ?rhs" using x_Min y_Min \x = y\ by -(erule proper_interval_not_set_less_eq_auxI, auto) } moreover { assume "\ ?lhs \' ?rhs" also have "?rhs = set (y # ys) \ above ao" using ao by auto finally obtain z where "z \ above ao" "z < y" using \x = y\ x_ao x_Min[unfolded Ball_def] by(fastforce simp add: set_less_eq_aux_def set_less_aux_def simp add: less_le not_le dest!: bspec[where x=y]) hence "proper_interval ao (Some y)" by(auto simp add: proper_interval_Some2) } ultimately show ?thesis using \x = y\ by auto qed qed qed } from this[of None] show ?Compl1 by simp qed lemma set_less_aux_Compl_iff: "set_less_aux_Compl ao xs ys \ set_less_eq_aux_Compl ao xs ys \ \ Compl_set_less_eq_aux ao ys xs" by(induct ao xs ys rule: set_less_aux_Compl.induct)(auto simp add: not_less_iff_gr_or_eq) lemma Compl_set_less_aux_Compl_iff: "Compl_set_less_aux ao xs ys \ Compl_set_less_eq_aux ao xs ys \ \ set_less_eq_aux_Compl ao ys xs" by(induct ao xs ys rule: Compl_set_less_aux.induct)(auto simp add: not_less_iff_gr_or_eq) theorem assumes fin: "finite (UNIV :: 'a set)" and xs: "sorted xs" "distinct xs" and ys: "sorted ys" "distinct ys" shows set_less_aux_Compl2_conv_set_less_aux_Compl: "set xs \' - set ys \ set_less_aux_Compl None xs ys" (is ?Compl2) and Compl1_set_less_aux_conv_Compl_set_less_aux: "- set xs \' set ys \ Compl_set_less_aux None xs ys" (is ?Compl1) using assms by(simp_all only: set_less_aux_conv_set_less_eq_aux set_less_aux_Compl_iff Compl_set_less_aux_Compl_iff set_less_eq_aux_Compl2_conv_set_less_eq_aux_Compl Compl1_set_less_eq_aux_conv_Compl_set_less_eq_aux) end subsection \Implementation of proper intervals for sets\ definition length_last :: "'a list \ nat \ 'a" where "length_last xs = (length xs, last xs)" lemma length_last_Nil [code]: "length_last [] = (0, undefined)" by(simp add: length_last_def last_def) lemma length_last_Cons_code [symmetric, code]: "fold (\x (n, _) . (n + 1, x)) xs (1, x) = length_last (x # xs)" by(induct xs rule: rev_induct)(simp_all add: length_last_def) context proper_intrvl begin fun exhaustive_above :: "'a \ 'a list \ bool" where "exhaustive_above x [] \ \ proper_interval (Some x) None" | "exhaustive_above x (y # ys) \ \ proper_interval (Some x) (Some y) \ exhaustive_above y ys" fun exhaustive :: "'a list \ bool" where "exhaustive [] = False" | "exhaustive (x # xs) \ \ proper_interval None (Some x) \ exhaustive_above x xs" fun proper_interval_set_aux :: "'a list \ 'a list \ bool" where "proper_interval_set_aux xs [] \ False" | "proper_interval_set_aux [] (y # ys) \ ys \ [] \ proper_interval (Some y) None" | "proper_interval_set_aux (x # xs) (y # ys) \ (if x < y then False else if y < x then proper_interval (Some y) (Some x) \ ys \ [] \ \ exhaustive_above x xs else proper_interval_set_aux xs ys)" fun proper_interval_set_Compl_aux :: "'a option \ nat \ 'a list \ 'a list \ bool" where "proper_interval_set_Compl_aux ao n [] [] \ CARD('a) > n + 1" | "proper_interval_set_Compl_aux ao n [] (y # ys) \ (let m = CARD('a) - n; (len_y, y') = length_last (y # ys) in m \ len_y \ (m = len_y + 1 \ \ proper_interval (Some y') None))" | "proper_interval_set_Compl_aux ao n (x # xs) [] \ (let m = CARD('a) - n; (len_x, x') = length_last (x # xs) in m \ len_x \ (m = len_x + 1 \ \ proper_interval (Some x') None))" | "proper_interval_set_Compl_aux ao n (x # xs) (y # ys) \ (if x < y then proper_interval ao (Some x) \ proper_interval_set_Compl_aux (Some x) (n + 1) xs (y # ys) else if y < x then proper_interval ao (Some y) \ proper_interval_set_Compl_aux (Some y) (n + 1) (x # xs) ys else proper_interval ao (Some x) \ (let m = card (UNIV :: 'a set) - n in m - length ys \ 2 \ m - length xs \ 2))" fun proper_interval_Compl_set_aux :: "'a option \ 'a list \ 'a list \ bool" where "proper_interval_Compl_set_aux ao (x # xs) (y # ys) \ (if x < y then \ proper_interval ao (Some x) \ proper_interval_Compl_set_aux (Some x) xs (y # ys) else if y < x then \ proper_interval ao (Some y) \ proper_interval_Compl_set_aux (Some y) (x # xs) ys else \ proper_interval ao (Some x) \ (ys = [] \ xs \ []))" | "proper_interval_Compl_set_aux ao _ _ \ False" end lemmas [code] = proper_intrvl.exhaustive_above.simps proper_intrvl.exhaustive.simps proper_intrvl.proper_interval_set_aux.simps proper_intrvl.proper_interval_set_Compl_aux.simps proper_intrvl.proper_interval_Compl_set_aux.simps context linorder_proper_interval begin lemma exhaustive_above_iff: "\ sorted xs; distinct xs; \x'\set xs. x < x' \ \ exhaustive_above x xs \ set xs = {z. z > x}" proof(induction x xs rule: exhaustive_above.induct) case 1 thus ?case by(simp add: proper_interval_simps) next case (2 x y ys) from \sorted (y # ys)\ \distinct (y # ys)\ have ys: "sorted ys" "distinct ys" and y: "\y'\set ys. y < y'" by(auto simp add: less_le) hence "exhaustive_above y ys = (set ys = {z. y < z})" by(rule "2.IH") moreover from \\y'\set (y # ys). x < y'\ have "x < y" by simp ultimately show ?case using y by(fastforce simp add: proper_interval_simps) qed lemma exhaustive_correct: assumes "sorted xs" "distinct xs" shows "exhaustive xs \ set xs = UNIV" proof(cases xs) case Nil thus ?thesis by auto next case Cons show ?thesis using assms unfolding Cons exhaustive.simps apply(subst exhaustive_above_iff) apply(auto simp add: less_le proper_interval_simps not_less intro: order_antisym) done qed theorem proper_interval_set_aux: assumes fin: "finite (UNIV :: 'a set)" and xs: "sorted xs" "distinct xs" and ys: "sorted ys" "distinct ys" shows "proper_interval_set_aux xs ys \ (\A. set xs \' A \ A \' set ys)" proof - note [simp] = finite_subset[OF subset_UNIV fin] show ?thesis using xs ys proof(induction xs ys rule: proper_interval_set_aux.induct) case 1 thus ?case by simp next case (2 y ys) hence "\y'\set ys. y < y'" by(auto simp add: less_le) thus ?case by(cases ys)(auto simp add: proper_interval_simps set_less_aux_singleton_iff intro: psubset_finite_imp_set_less_aux) next case (3 x xs y ys) from \sorted (x # xs)\ \distinct (x # xs)\ have xs: "sorted xs" "distinct xs" and x: "\x'\set xs. x < x'" by(auto simp add: less_le) from \sorted (y # ys)\ \distinct (y # ys)\ have ys: "sorted ys" "distinct ys" and y: "\y'\set ys. y < y'" by(auto simp add: less_le) have Minxxs: "Min (set (x # xs)) = x" and xnxs: "x \ set xs" using x by(auto intro!: Min_eqI) have Minyys: "Min (set (y # ys)) = y" and ynys: "y \ set ys" using y by(auto intro!: Min_eqI) show ?case proof(cases "x < y") case True hence "set (y # ys) \' set (x # xs)" using Minxxs Minyys by -(rule set_less_aux_Min_antimono, simp_all) thus ?thesis using True by(auto dest: set_less_aux_trans set_less_aux_antisym) next case False show ?thesis proof(cases "y < x") case True { assume "proper_interval (Some y) (Some x)" then obtain z where z: "y < z" "z < x" by(auto simp add: proper_interval_simps) hence "set (x # xs) \' {z}" using x by -(rule set_less_aux_Min_antimono, auto) moreover have "{z} \' set (y # ys)" using z y Minyys by -(rule set_less_aux_Min_antimono, auto) ultimately have "\A. set (x # xs) \' A \ A \' set (y # ys)" by blast } moreover { assume "ys \ []" hence "{y} \' set (y # ys)" using y by -(rule psubset_finite_imp_set_less_aux, auto simp add: neq_Nil_conv) moreover have "set (x # xs) \' {y}" using False True x by -(rule set_less_aux_Min_antimono, auto) ultimately have "\A. set (x # xs) \' A \ A \' set (y # ys)" by blast } moreover { assume "\ exhaustive_above x xs" then obtain z where z: "z > x" "z \ set xs" using x by(auto simp add: exhaustive_above_iff[OF xs x]) let ?A = "insert z (set (x # xs))" have "set (x # xs) \' ?A" using z by -(rule psubset_finite_imp_set_less_aux, auto) moreover have "?A \' set (y # ys)" using Minyys False True z x by -(rule set_less_aux_Min_antimono, auto) ultimately have "\A. set (x # xs) \' A \ A \' set (y # ys)" by blast } moreover { fix A assume A: "set (x # xs) \' A" and A': "A \' {y}" and pi: "\ proper_interval (Some y) (Some x)" from A have nempty: "A \ {}" by auto have "y \ A" proof assume "y \ A" moreover with A' have "A \ {y}" by auto ultimately have "{y} \' A" by -(rule psubset_finite_imp_set_less_aux, auto) with A' show False by(rule set_less_aux_antisym) qed have "y < Min A" unfolding not_le[symmetric] proof assume "Min A \ y" moreover have "Min A \ y" using \y \ A\ nempty by clarsimp ultimately have "Min A < Min {y}" by simp hence "{y} \' A" by(rule set_less_aux_Min_antimono)(simp_all add: nempty) with A' show False by(rule set_less_aux_antisym) qed with pi nempty have "x \ Min A" by(auto simp add: proper_interval_simps) moreover from A obtain z where z: "z \ A" "z \ set (x # xs)" by(auto simp add: set_less_aux_def) with \x \ Min A\ nempty have "x < z" by auto with z have "\ exhaustive_above x xs" by(auto simp add: exhaustive_above_iff[OF xs x]) } ultimately show ?thesis using True False by fastforce next case False with \\ x < y\ have "x = y" by auto from \\ x < y\ False have "proper_interval_set_aux xs ys = (\A. set xs \' A \ A \' set ys)" using xs ys by(rule "3.IH") also have "\ = (\A. set (x # xs) \' A \ A \' set (y # ys))" (is "?lhs = ?rhs") proof assume ?lhs then obtain A where A: "set xs \' A" and A': "A \' set ys" by blast hence nempty: "A \ {}" "ys \ []" by auto let ?A = "insert y A" { assume "Min A \ y" also from y nempty have "y < Min (set ys)" by auto finally have "set ys \' A" by(rule set_less_aux_Min_antimono)(simp_all add: nempty) with A' have False by(rule set_less_aux_antisym) } hence MinA: "y < Min A" by(metis not_le) with nempty have "y \ A" by auto moreover with MinA nempty have MinyA: "Min ?A = y" by -(rule Min_eqI, auto) ultimately have A1: "set (x # xs) \' ?A" using \x = y\ A Minxxs xnxs by(subst set_less_aux_rec) simp_all moreover have "?A \' set (y # ys)" using \x = y\ MinyA \y \ A\ A' Minyys ynys by(subst set_less_aux_rec) simp_all ultimately show ?rhs by blast next assume "?rhs" then obtain A where A: "set (x # xs) \' A" and A': "A \' set (y # ys)" by blast let ?A = "A - {x}" from A have nempty: "A \ {}" by auto { assume "x < Min A" hence "Min (set (x # xs)) < Min A" unfolding Minxxs . hence "A \' set (x # xs)" by(rule set_less_aux_Min_antimono) simp_all with A have False by(rule set_less_aux_antisym) } moreover { assume "Min A < x" hence "Min A < Min (set (y # ys))" unfolding \x = y\ Minyys . hence "set (y # ys) \' A" by(rule set_less_aux_Min_antimono)(simp_all add: nempty) with A' have False by(rule set_less_aux_antisym) } ultimately have MinA: "Min A = x" by(metis less_linear) hence "x \ A" using nempty by(metis Min_in \finite A\) from A nempty Minxxs xnxs have "set xs \' ?A" by(subst (asm) set_less_aux_rec)(auto simp add: MinA) moreover from A' \x = y\ nempty Minyys MinA ynys have "?A \' set ys" by(subst (asm) set_less_aux_rec) simp_all ultimately show ?lhs by blast qed finally show ?thesis using \x = y\ by simp qed qed qed qed lemma proper_interval_set_Compl_aux: assumes fin: "finite (UNIV :: 'a set)" and xs: "sorted xs" "distinct xs" and ys: "sorted ys" "distinct ys" shows "proper_interval_set_Compl_aux None 0 xs ys \ (\A. set xs \' A \ A \' - set ys)" proof - note [simp] = finite_subset[OF subset_UNIV fin] define above where "above = case_option UNIV (Collect \ less)" have above_simps [simp]: "above None = UNIV" "\x. above (Some x) = {y. x < y}" and above_upclosed: "\x y ao. \ x \ above ao; x < y \ \ y \ above ao" and proper_interval_Some2: "\x ao. proper_interval ao (Some x) \ (\z\above ao. z < x)" by(simp_all add: proper_interval_simps above_def split: option.splits) { fix ao n assume "set xs \ above ao" "set ys \ above ao" from xs \set xs \ above ao\ ys \set ys \ above ao\ have "proper_interval_set_Compl_aux ao (card (UNIV - above ao)) xs ys \ (\A \ above ao. set xs \' A \ A \' - set ys \ above ao)" proof(induct ao n\"card (UNIV - above ao)" xs ys rule: proper_interval_set_Compl_aux.induct) case (1 ao) have "card (UNIV - above ao) + 1 < CARD('a) \ (\A \ above ao. A \ {} \ A \' above ao)" (is "?lhs \ ?rhs") proof assume ?lhs hence "card (UNIV - (UNIV - above ao)) > 1" by(simp add: card_Diff_subset) from card_gt_1D[OF this] obtain x y where above: "x \ above ao" "y \ above ao" and neq: "x \ y" by blast hence "{x} \' {x, y} \ above ao" by(simp_all add: psubsetI psubset_finite_imp_set_less_aux) also have "\ \' above ao" by(auto intro: subset_finite_imp_set_less_eq_aux) finally show ?rhs using above by blast next assume ?rhs then obtain A where nempty: "A \ above ao \ {}" and subset: "A \ above ao" and less: "A \' above ao" by blast from nempty obtain x where x: "x \ A" "x \ above ao" by blast show ?lhs proof(rule ccontr) assume "\ ?lhs" moreover have "CARD('a) \ card (UNIV - above ao)" by(rule card_mono) simp_all moreover from card_Un_disjoint[of "UNIV - above ao" "above ao"] have "CARD('a) = card (UNIV - above ao) + card (above ao)" by auto ultimately have "card (above ao) = 1" using x by(cases "card (above ao)")(auto simp add: not_less_eq less_Suc_eq_le) with x have "above ao = {x}" unfolding card_eq_1_iff by auto moreover with x subset have A: "A = {x}" by auto ultimately show False using less by simp qed qed thus ?case by simp next case (2 ao y ys) note ys = \sorted (y # ys)\ \distinct (y # ys)\ \set (y # ys) \ above ao\ have len_ys: "length ys = card (set ys)" using ys by(auto simp add: List.card_set intro: sym) define m where "m = CARD('a) - card (UNIV - above ao)" have "CARD('a) = card (above ao) + card (UNIV - above ao)" using card_Un_disjoint[of "above ao" "UNIV - above ao"] by auto hence m_eq: "m = card (above ao)" unfolding m_def by simp have "m \ length ys + 1 \ (m = length ys + 2 \ \ proper_interval (Some (last (y # ys))) None) \ (\A \ above ao. A \ {} \ A \' - set (y # ys) \ above ao)" (is "?lhs \ ?rhs") proof assume ?lhs hence m: "m \ length ys + 1" and pi: "m = length ys + 2 \ \ proper_interval (Some (last (y # ys))) None" by simp_all have "length ys + 1 = card (set (y # ys))" using ys len_ys by simp also have "\ \ m" unfolding m_eq by(rule card_mono)(simp, rule ys) finally have "length ys + 2 \ m" using m by simp show ?rhs proof(cases "m = length ys + 2") case True hence "card (UNIV - (UNIV - above ao) - set (y # ys)) = 1" using ys len_ys by(subst card_Diff_subset)(auto simp add: m_def card_Diff_subset) then obtain z where z: "z \ above ao" "z \ y" "z \ set ys" unfolding card_eq_1_iff by auto from True have "\ proper_interval (Some (last (y # ys))) None" by(rule pi) hence "z \ last (y # ys)" by(simp add: proper_interval_simps not_less del: last.simps) moreover have ly: "last (y # ys) \ set (y # ys)" by(rule last_in_set) simp with z have "z \ last (y # ys)" by auto ultimately have "z < last (y # ys)" by simp hence "{last (y # ys)} \' {z}" using z ly ys by(auto 4 3 simp add: set_less_aux_def) also have "\ \' - set (y # ys) \ above ao" using z by(auto intro: subset_finite_imp_set_less_eq_aux) also have "{last (y # ys)} \ {}" using ly ys by blast moreover have "{last (y # ys)} \ above ao" using ys by auto ultimately show ?thesis by blast next case False with \length ys + 2 \ m\ ys len_ys have "card (UNIV - (UNIV - above ao) - set (y # ys)) > 1" by(subst card_Diff_subset)(auto simp add: card_Diff_subset m_def) from card_gt_1D[OF this] obtain x x' where x: "x \ above ao" "x \ y" "x \ set ys" and x': "x' \ above ao" "x' \ y" "x' \ set ys" and neq: "x \ x'" by auto hence "{x} \' {x, x'} \ above ao" by(simp_all add: psubsetI psubset_finite_imp_set_less_aux) also have "\ \' - set (y # ys) \ above ao" using x x' ys by(auto intro: subset_finite_imp_set_less_eq_aux) also have "{x} \ above ao \ {}" using x by simp ultimately show ?rhs by blast qed next assume ?rhs then obtain A where nempty: "A \ {}" and less: "A \' - set (y # ys) \ above ao" and subset: "A \ above ao" by blast have "card (set (y # ys)) \ card (above ao)" using ys(3) by(simp add: card_mono) hence "length ys + 1 \ m" unfolding m_eq using ys by(simp add: len_ys) have "m \ length ys + 1" proof assume "m = length ys + 1" hence "card (above ao) \ card (set (y # ys))" unfolding m_eq using ys len_ys by auto from card_seteq[OF _ _ this] ys have "set (y # ys) = above ao" by simp with nempty less show False by(auto simp add: set_less_aux_def) qed moreover { assume "m = length ys + 2" hence "card (above ao - set (y # ys)) = 1" using ys len_ys m_eq by(auto simp add: card_Diff_subset) then obtain z where z: "above ao - set (y # ys) = {z}" unfolding card_eq_1_iff .. hence eq_z: "- set (y # ys) \ above ao = {z}" by auto with less have "A \' {z}" by simp have "\ proper_interval (Some (last (y # ys))) None" proof assume "proper_interval (Some (last (y # ys))) None" then obtain z' where z': "last (y # ys) < z'" by(clarsimp simp add: proper_interval_simps) have "last (y # ys) \ set (y # ys)" by(rule last_in_set) simp with ys z' have "z' \ above ao" "z' \ set (y # ys)" - by(auto simp del: last.simps sorted.simps(2) intro: above_upclosed dest: sorted_last) + using above_upclosed local.not_less local.sorted_last ys(1) z' by blast+ with eq_z have "z = z'" by fastforce from z' have "\x. x \ set (y # ys) \ x < z'" using ys - by(auto dest: sorted_last simp del: sorted.simps(2)) + by(auto dest: sorted_last simp del: sorted_wrt.simps(2)) with eq_z \z = z'\ have "\x. x \ above ao \ x \ z'" by(fastforce) with \A \' {z}\ nempty \z = z'\ subset show False by(auto simp add: set_less_aux_def) qed } ultimately show ?lhs by simp qed thus ?case by(simp add: length_last_def m_def Let_def) next case (3 ao x xs) note xs = \sorted (x # xs)\ \distinct (x # xs)\ \set (x # xs) \ above ao\ have len_xs: "length xs = card (set xs)" using xs by(auto simp add: List.card_set intro: sym) define m where "m = CARD('a) - card (UNIV - above ao)" have "CARD('a) = card (above ao) + card (UNIV - above ao)" using card_Un_disjoint[of "above ao" "UNIV - above ao"] by auto hence m_eq: "m = card (above ao)" unfolding m_def by simp have "m \ length xs + 1 \ (m = length xs + 2 \ \ proper_interval (Some (last (x # xs))) None) \ (\A \ above ao. set (x # xs) \' A \ A \' above ao)" (is "?lhs \ ?rhs") proof assume ?lhs hence m: "m \ length xs + 1" and pi: "m = length xs + 2 \ \ proper_interval (Some (last (x # xs))) None" by simp_all have "length xs + 1 = card (set (x # xs))" using xs len_xs by simp also have "\ \ m" unfolding m_eq by(rule card_mono)(simp, rule xs) finally have "length xs + 2 \ m" using m by simp show ?rhs proof(cases "m = length xs + 2") case True hence "card (UNIV - (UNIV - above ao) - set (x # xs)) = 1" using xs len_xs by(subst card_Diff_subset)(auto simp add: m_def card_Diff_subset) then obtain z where z: "z \ above ao" "z \ x" "z \ set xs" unfolding card_eq_1_iff by auto define A where "A = insert z {y. y \ set (x # xs) \ y < z}" from True have "\ proper_interval (Some (last (x # xs))) None" by(rule pi) hence "z \ last (x # xs)" by(simp add: proper_interval_simps not_less del: last.simps) moreover have lx: "last (x # xs) \ set (x # xs)" by(rule last_in_set) simp with z have "z \ last (x # xs)" by auto ultimately have "z < last (x # xs)" by simp hence "set (x # xs) \' A" using z xs by(auto simp add: A_def set_less_aux_def intro: rev_bexI[where x=z]) moreover have "last (x # xs) \ A" using xs \z < last (x # xs)\ by(auto simp add: A_def simp del: last.simps) hence "A \ insert (last (x # xs)) A" by blast hence less': "A \' insert (last (x # xs)) A" by(rule psubset_finite_imp_set_less_aux) simp have "\ \ above ao" using xs lx z by(auto simp del: last.simps simp add: A_def) hence "insert (last (x # xs)) A \' above ao" by(auto intro: subset_finite_imp_set_less_eq_aux) with less' have "A \' above ao" by(rule set_less_trans_set_less_eq) moreover have "A \ above ao" using xs z by(auto simp add: A_def) ultimately show ?thesis by blast next case False with \length xs + 2 \ m\ xs len_xs have "card (UNIV - (UNIV - above ao) - set (x # xs)) > 1" by(subst card_Diff_subset)(auto simp add: card_Diff_subset m_def) from card_gt_1D[OF this] obtain y y' where y: "y \ above ao" "y \ x" "y \ set xs" and y': "y' \ above ao" "y' \ x" "y' \ set xs" and neq: "y \ y'" by auto define A where "A = insert y (set (x # xs) \ above ao)" hence "set (x # xs) \ A" using y xs by auto hence "set (x # xs) \' \" by(fastforce intro: psubset_finite_imp_set_less_aux) moreover have *: "\ \ above ao" using y y' neq by(auto simp add: A_def) moreover from * have "A \' above ao" by(auto intro: psubset_finite_imp_set_less_aux) ultimately show ?thesis by blast qed next assume ?rhs then obtain A where lessA: "set (x # xs) \' A" and Aless: "A \' above ao" and subset: "A \ above ao" by blast have "card (set (x # xs)) \ card (above ao)" using xs(3) by(simp add: card_mono) hence "length xs + 1 \ m" unfolding m_eq using xs by(simp add: len_xs) have "m \ length xs + 1" proof assume "m = length xs + 1" hence "card (above ao) \ card (set (x # xs))" unfolding m_eq using xs len_xs by auto from card_seteq[OF _ _ this] xs have "set (x # xs) = above ao" by simp with lessA Aless show False by(auto dest: set_less_aux_antisym) qed moreover { assume "m = length xs + 2" hence "card (above ao - set (x # xs)) = 1" using xs len_xs m_eq by(auto simp add: card_Diff_subset) then obtain z where z: "above ao - set (x # xs) = {z}" unfolding card_eq_1_iff .. have "\ proper_interval (Some (last (x # xs))) None" proof assume "proper_interval (Some (last (x # xs))) None" then obtain z' where z': "last (x # xs) < z'" by(clarsimp simp add: proper_interval_simps) have "last (x # xs) \ set (x # xs)" by(rule last_in_set) simp with xs z' have "z' \ above ao" "z' \ set (x # xs)" - by(auto simp del: last.simps sorted.simps(2) intro: above_upclosed dest: sorted_last) + by(auto simp del: last.simps sorted_wrt.simps(2) intro: above_upclosed dest: sorted_last) with z have "z = z'" by fastforce from z' have y_less: "\y. y \ set (x # xs) \ y < z'" using xs - by(auto simp del: sorted.simps(2) dest: sorted_last) + by(auto simp del: sorted_wrt.simps(2) dest: sorted_last) with z \z = z'\ have "\y. y \ above ao \ y \ z'" by(fastforce) from lessA subset obtain y where y: "y \ A" "y \ above ao" "y \ set (x # xs)" and min: "\y'. \ y' \ set (x # xs); y' \ above ao; y' \ A \ \ y \ y'" by(auto simp add: set_less_aux_def) with z \z = z'\ have "y = z'" by auto have "set (x # xs) \ A" proof fix y' assume y': "y' \ set (x # xs)" show "y' \ A" proof(rule ccontr) assume "y' \ A" from y' xs have "y' \ above ao" by auto with y' have "y \ y'" using \y' \ A\ by(rule min) moreover from y' have "y' < z'" by(rule y_less) ultimately show False using \y = z'\ by simp qed qed moreover from z xs have "above ao = insert z (set (x # xs))" by auto ultimately have "A = above ao" using y \y = z'\ \z = z'\ subset by auto with Aless show False by simp qed } ultimately show ?lhs by simp qed thus ?case by(simp add: length_last_def m_def Let_def del: last.simps) next case (4 ao x xs y ys) note xxs = \sorted (x # xs)\ \distinct (x # xs)\ and yys = \sorted (y # ys)\ \distinct (y # ys)\ and xxs_above = \set (x # xs) \ above ao\ and yys_above = \set (y # ys) \ above ao\ from xxs have xs: "sorted xs" "distinct xs" and x_Min: "\x'\set xs. x < x'" by(auto simp add: less_le) from yys have ys: "sorted ys" "distinct ys" and y_Min: "\y'\set ys. y < y'" by(auto simp add: less_le) have len_xs: "length xs = card (set xs)" using xs by(auto simp add: List.card_set intro: sym) have len_ys: "length ys = card (set ys)" using ys by(auto simp add: List.card_set intro: sym) show ?case proof(cases "x < y") case True have "proper_interval ao (Some x) \ proper_interval_set_Compl_aux (Some x) (card (UNIV - above ao) + 1) xs (y # ys) \ (\A \ above ao. set (x # xs) \' A \ A \' - set (y # ys) \ above ao)" (is "?lhs \ ?rhs") proof(cases "proper_interval ao (Some x)") case True then obtain z where z: "z \ above ao" "z < x" by(clarsimp simp add: proper_interval_Some2) moreover with xxs have "\x'\set xs. z < x'" by(auto) ultimately have "set (x # xs) \' {z}" by(auto simp add: set_less_aux_def intro!: bexI[where x=z]) moreover { from z yys \x < y\ have "z < y" "\y'\set ys. z < y'" by(auto) hence subset: "{z} \ - set (y # ys) \ above ao" using ys \x < y\ z by auto moreover have "x \ \" using yys xxs \x < y\ xxs_above by(auto) ultimately have "{z} \ \" using \z < x\ by fastforce hence "{z} \' \" by(fastforce intro: psubset_finite_imp_set_less_aux) } moreover have "{z} \ above ao" using z by simp ultimately have ?rhs by blast thus ?thesis using True by simp next case False hence above_eq: "above ao = insert x (above (Some x))" using xxs_above by(auto simp add: proper_interval_Some2 intro: above_upclosed) moreover have "card (above (Some x)) < CARD('a)" by(rule psubset_card_mono)(auto) ultimately have card_eq: "card (UNIV - above ao) + 1 = card (UNIV - above (Some x))" by(simp add: card_Diff_subset) from xxs_above x_Min have xs_above: "set xs \ above (Some x)" by(auto) from \x < y\ y_Min have "set (y # ys) \ above (Some x)" by(auto) with \x < y\ card_eq xs xs_above yys have "proper_interval_set_Compl_aux (Some x) (card (UNIV - above ao) + 1) xs (y # ys) \ (\A \ above (Some x). set xs \' A \ A \' - set (y # ys) \ above (Some x))" by(subst card_eq)(rule 4) also have "\ \ ?rhs" (is "?lhs' \ _") proof assume ?lhs' then obtain A where less_A: "set xs \' A" and A_less: "A \' - set (y # ys) \ above (Some x)" and subset: "A \ above (Some x)" by blast let ?A = "insert x A" have Min_A': "Min ?A = x" using xxs_above False subset by(auto intro!: Min_eqI simp add: proper_interval_Some2) moreover have "Min (set (x # xs)) = x" using x_Min by(auto intro!: Min_eqI) moreover have Amx: "A - {x} = A" using False subset by(auto simp add: proper_interval_Some2 intro: above_upclosed) moreover have "set xs - {x} = set xs" using x_Min by auto ultimately have less_A': "set (x # xs) \' ?A" using less_A xxs_above x_Min by(subst set_less_aux_rec) simp_all have "x \ - insert y (set ys) \ above ao" using \x < y\ xxs_above y_Min by auto hence "- insert y (set ys) \ above ao \ {}" by auto moreover have "Min (- insert y (set ys) \ above ao) = x" using yys y_Min xxs_above \x < y\ False by(auto intro!: Min_eqI simp add: proper_interval_Some2) moreover have "- set (y # ys) \ above ao - {x} = - set (y # ys) \ above (Some x)" using yys_above False xxs_above by(auto simp add: proper_interval_Some2 intro: above_upclosed) ultimately have A'_less: "?A \' - set (y # ys) \ above ao" using Min_A' A_less Amx xxs_above by(subst set_less_aux_rec) simp_all moreover have "?A \ above ao" using subset xxs_above by(auto intro: above_upclosed) ultimately show ?rhs using less_A' by blast next assume ?rhs then obtain A where less_A: "set (x # xs) \' A" and A_less: "A \' - set (y # ys) \ above ao" and subset: "A \ above ao" by blast let ?A = "A - {x}" from less_A subset xxs_above have "set (x # xs) \ above ao \' A \ above ao" by(simp add: Int_absorb2) with False xxs_above subset have "x \ A" by(auto simp add: set_less_aux_def proper_interval_Some2) hence "\ \ {}" by auto moreover from \x \ A\ False subset have Min_A: "Min A = x" by(auto intro!: Min_eqI simp add: proper_interval_Some2 not_less) moreover have "Min (set (x # xs)) = x" using x_Min by(auto intro!: Min_eqI) moreover have eq_A: "?A \ above (Some x)" using xxs_above False subset by(fastforce simp add: proper_interval_Some2 not_less intro: above_upclosed) moreover have "set xs - {x} = set xs" using x_Min by(auto) ultimately have less_A': "set xs \' ?A" using xxs_above less_A by(subst (asm) set_less_aux_rec)(simp_all cong: conj_cong) have "x \ - insert y (set ys) \ above ao" using \x < y\ xxs_above y_Min by auto hence "- insert y (set ys) \ above ao \ {}" by auto moreover have "Min (- set (y # ys) \ above ao) = x" using yys y_Min xxs_above \x < y\ False by(auto intro!: Min_eqI simp add: proper_interval_Some2) moreover have "- set (y # ys) \ above (Some x) = - set (y # ys) \ above ao - {x}" by(auto simp add: above_eq) ultimately have "?A \' - set (y # ys) \ above (Some x)" using A_less \A \ {}\ eq_A Min_A by(subst (asm) set_less_aux_rec) simp_all with less_A' eq_A show ?lhs' by blast qed finally show ?thesis using False by simp qed thus ?thesis using True by simp next case False show ?thesis proof(cases "y < x") case True have "proper_interval ao (Some y) \ proper_interval_set_Compl_aux (Some y) (card (UNIV - above ao) + 1) (x # xs) ys \ (\A \ above ao. set (x # xs) \' A \ A \' - set (y # ys) \ above ao)" (is "?lhs \ ?rhs") proof(cases "proper_interval ao (Some y)") case True then obtain z where z: "z \ above ao" "z < y" by(clarsimp simp add: proper_interval_Some2) from xxs \y < x\ have "\x'\set (x # xs). y < x'" by(auto) hence less_A: "set (x # xs) \' {y}" by(auto simp add: set_less_aux_def intro!: bexI[where x=y]) have "{y} \' {z}" using z y_Min by(auto simp add: set_less_aux_def intro: bexI[where x=z]) also have "\ \ - set (y # ys) \ above ao" using z y_Min by auto hence "{z} \' \" by(auto intro: subset_finite_imp_set_less_eq_aux) finally have "{y} \' \" . moreover have "{y} \ above ao" using yys_above by auto ultimately have ?rhs using less_A by blast thus ?thesis using True by simp next case False hence above_eq: "above ao = insert y (above (Some y))" using yys_above by(auto simp add: proper_interval_Some2 intro: above_upclosed) moreover have "card (above (Some y)) < CARD('a)" by(rule psubset_card_mono)(auto) ultimately have card_eq: "card (UNIV - above ao) + 1 = card (UNIV - above (Some y))" by(simp add: card_Diff_subset) from yys_above y_Min have ys_above: "set ys \ above (Some y)" by(auto) have eq_ys: "- set ys \ above (Some y) = - set (y # ys) \ above ao" by(auto simp add: above_eq) from \y < x\ x_Min have "set (x # xs) \ above (Some y)" by(auto) with \\ x < y\ \y < x\ card_eq xxs ys ys_above have "proper_interval_set_Compl_aux (Some y) (card (UNIV - above ao) + 1) (x # xs) ys \ (\A \ above (Some y). set (x # xs) \' A \ A \' - set ys \ above (Some y))" by(subst card_eq)(rule 4) also have "\ \ ?rhs" (is "?lhs' \ _") proof assume ?lhs' then obtain A where "set (x # xs) \' A" and subset: "A \ above (Some y)" and "A \' - set ys \ above (Some y)" by blast moreover from subset have "A - {y} = A" by auto ultimately have "set (x # xs) \' A - {y}" and "A - {y} \' - set (y # ys) \ above ao" using eq_ys by simp_all moreover from subset have "A - {y} \ above ao" using yys_above by(auto intro: above_upclosed) ultimately show ?rhs by blast next assume ?rhs then obtain A where "set (x # xs) \' A" and A_less: "A \' - set (y # ys) \ above ao" and subset: "A \ above ao" by blast moreover from A_less False yys_above have "y \ A" by(auto simp add: set_less_aux_def proper_interval_Some2 not_less) ultimately have "set (x # xs) \' A" and "A \' - set ys \ above (Some y)" using eq_ys by simp_all moreover from \y \ A\ subset above_eq have "A \ above (Some y)" by auto ultimately show ?lhs' by blast qed finally show ?thesis using False by simp qed with False True show ?thesis by simp next case False with \\ x < y\ have "x = y" by simp have "proper_interval ao (Some x) \ (CARD('a) - (card (UNIV - above ao) + length ys) \ 2 \ CARD('a) - (card (UNIV - above ao) + length xs) \ 2) \ (\A \ above ao. set (x # xs) \' A \ A \' - set (y # ys) \ above ao)" (is "?below \ ?card \ ?rhs") proof(cases "?below") case False hence "- set (y # ys) \ above ao \' set (x # xs)" using \x = y\ yys_above xxs_above y_Min by(auto simp add: not_less set_less_aux_def proper_interval_Some2 intro!: bexI[where x=y]) with False show ?thesis by(auto dest: set_less_aux_trans) next case True then obtain z where z: "z \ above ao" "z < x" by(clarsimp simp add: proper_interval_Some2) have "?card \ ?rhs" proof assume ?rhs then obtain A where less_A: "set (x # xs) \' A" and A_less: "A \' - set (y # ys) \ above ao" and subset: "A \ above ao" by blast { assume c_ys: "CARD('a) - (card (UNIV - above ao) + length ys) = 2" and c_xs: "CARD('a) - (card (UNIV - above ao) + length xs) = 2" from c_ys yys_above len_ys y_Min have "card (UNIV - (UNIV - above ao) - set (y # ys)) = 1" by(subst card_Diff_subset)(auto simp add: card_Diff_subset) then obtain z' where eq_y: "- set (y # ys) \ above ao = {z'}" unfolding card_eq_1_iff by auto moreover from z have "z \ set (y # ys)" using \x = y\ y_Min by auto ultimately have "z' = z" using z by fastforce from c_xs xxs_above len_xs x_Min have "card (UNIV - (UNIV - above ao) - set (x # xs)) = 1" by(subst card_Diff_subset)(auto simp add: card_Diff_subset) then obtain z'' where eq_x: "- set (x # xs) \ above ao = {z''}" unfolding card_eq_1_iff by auto moreover from z have "z \ set (x # xs)" using x_Min by auto ultimately have "z'' = z" using z by fastforce from less_A subset obtain q where "q \ A" "q \ above ao" "q \ set (x # xs)" by(auto simp add: set_less_aux_def) hence "q \ {z''}" unfolding eq_x[symmetric] by simp hence "q = z''" by simp with \q \ A\ \z' = z\ \z'' = z\ z have "- set (y # ys) \ above ao \ A" unfolding eq_y by simp hence "- set (y # ys) \ above ao \' A" by(auto intro: subset_finite_imp_set_less_eq_aux) with A_less have False by(auto dest: set_less_trans_set_less_eq) } thus ?card by auto next assume ?card (is "?card_ys \ ?card_xs") thus ?rhs proof assume ?card_ys let ?YS = "UNIV - (UNIV - above ao) - set (y # ys)" from \?card_ys\ yys_above len_ys y_Min have "card ?YS \ 1" by(subst card_Diff_subset)(auto simp add: card_Diff_subset) moreover have "?YS \ {}" using True y_Min yys_above \x = y\ by(fastforce simp add: proper_interval_Some2) hence "card ?YS \ 0" by simp ultimately have "card ?YS > 1" by(cases "card ?YS") simp_all from card_gt_1D[OF this] obtain x' y' where x': "x' \ above ao" "x' \ set (y # ys)" and y': "y' \ above ao" "y' \ set (y # ys)" and neq: "x' \ y'" by auto let ?A = "{z}" have "set (x # xs) \' ?A" using z x_Min by(auto simp add: set_less_aux_def intro!: rev_bexI) moreover { have "?A \ - set (y # ys) \ above ao" using z \x = y\ y_Min by(auto) moreover have "x' \ ?A \ y' \ ?A" using neq by auto with x' y' have "?A \ - set (y # ys) \ above ao" by auto ultimately have "?A \ - set (y # ys) \ above ao" by(rule psubsetI) hence "?A \' \" by(fastforce intro: psubset_finite_imp_set_less_aux) } ultimately show ?thesis using z by blast next assume ?card_xs let ?XS = "UNIV - (UNIV - above ao) - set (x # xs)" from \?card_xs\ xxs_above len_xs x_Min have "card ?XS \ 1" by(subst card_Diff_subset)(auto simp add: card_Diff_subset) moreover have "?XS \ {}" using True x_Min xxs_above by(fastforce simp add: proper_interval_Some2) hence "card ?XS \ 0" by simp ultimately have "card ?XS > 1" by(cases "card ?XS") simp_all from card_gt_1D[OF this] obtain x' y' where x': "x' \ above ao" "x' \ set (x # xs)" and y': "y' \ above ao" "y' \ set (x # xs)" and neq: "x' \ y'" by auto define A where "A = (if x' = Min (above ao) then insert y' (set (x # xs)) else insert x' (set (x # xs)))" hence "set (x # xs) \ A" by auto moreover have "set (x # xs) \ \" using neq x' y' by(auto simp add: A_def) ultimately have "set (x # xs) \ A" .. hence "set (x # xs) \' \" by(fastforce intro: psubset_finite_imp_set_less_aux) moreover { have nempty: "above ao \ {}" using z by auto have "A \' {Min (above ao)}" using z x' y' neq \x = y\ x_Min xxs_above by(auto 6 4 simp add: set_less_aux_def A_def nempty intro!: rev_bexI Min_eqI) also have "Min (above ao) \ z" using z by(simp) hence "Min (above ao) < x" using \z < x\ by(rule le_less_trans) with \x = y\ y_Min have "Min (above ao) \ set (y # ys)" by auto hence "{Min (above ao)} \ - set (y # ys) \ above ao" by(auto simp add: nempty) hence "{Min (above ao)} \' \" by(auto intro: subset_finite_imp_set_less_eq_aux) finally have "A \' \" . } moreover have "A \ above ao" using xxs_above yys_above x' y' by(auto simp add: A_def) ultimately show ?rhs by blast qed qed thus ?thesis using True by simp qed thus ?thesis using \x = y\ by simp qed qed qed } from this[of None] show ?thesis by(simp) qed lemma proper_interval_Compl_set_aux: assumes fin: "finite (UNIV :: 'a set)" and xs: "sorted xs" "distinct xs" and ys: "sorted ys" "distinct ys" shows "proper_interval_Compl_set_aux None xs ys \ (\A. - set xs \' A \ A \' set ys)" proof - note [simp] = finite_subset[OF subset_UNIV fin] define above where "above = case_option UNIV (Collect \ less)" have above_simps [simp]: "above None = UNIV" "\x. above (Some x) = {y. x < y}" and above_upclosed: "\x y ao. \ x \ above ao; x < y \ \ y \ above ao" and proper_interval_Some2: "\x ao. proper_interval ao (Some x) \ (\z\above ao. z < x)" by(simp_all add: proper_interval_simps above_def split: option.splits) { fix ao n assume "set xs \ above ao" "set ys \ above ao" from xs \set xs \ above ao\ ys \set ys \ above ao\ have "proper_interval_Compl_set_aux ao xs ys \ (\A. - set xs \ above ao \' A \ above ao \ A \ above ao \' set ys \ above ao)" proof(induction ao xs ys rule: proper_interval_Compl_set_aux.induct) case ("2_1" ao ys) { fix A assume "above ao \' A \ above ao" also have "\ \ above ao" by simp hence "A \ above ao \' above ao" by(auto intro: subset_finite_imp_set_less_eq_aux) finally have False by simp } thus ?case by auto next case ("2_2" ao xs) thus ?case by simp next case (1 ao x xs y ys) note xxs = \sorted (x # xs)\ \distinct (x # xs)\ hence xs: "sorted xs" "distinct xs" and x_Min: "\x' \ set xs. x < x'" by(auto simp add: less_le) note yys = \sorted (y # ys)\ \distinct (y # ys)\ hence ys: "sorted ys" "distinct ys" and y_Min: "\y'\set ys. y < y'" by(auto simp add: less_le) note xxs_above = \set (x # xs) \ above ao\ note yys_above = \set (y # ys) \ above ao\ show ?case proof(cases "x < y") case True have "\ proper_interval ao (Some x) \ proper_interval_Compl_set_aux (Some x) xs (y # ys) \ (\A. - set (x # xs) \ above ao \' A \ above ao \ A \ above ao \' set (y # ys) \ above ao)" (is "?lhs \ ?rhs") proof(cases "proper_interval ao (Some x)") case True then obtain z where z: "z < x" "z \ above ao" by(auto simp add: proper_interval_Some2) hence nempty: "above ao \ {}" by auto with z have "Min (above ao) \ z" by auto hence "Min (above ao) < x" using \z < x\ by(rule le_less_trans) hence "set (y # ys) \ above ao \' - set (x # xs) \ above ao" using y_Min x_Min z \x < y\ by(fastforce simp add: set_less_aux_def nempty intro!: Min_eqI bexI[where x="Min (above ao)"]) thus ?thesis using True by(auto dest: set_less_aux_trans set_less_aux_antisym) next case False hence above_eq: "above ao = insert x (above (Some x))" using xxs_above by(auto simp add: proper_interval_Some2 intro: above_upclosed) from x_Min have xs_above: "set xs \ above (Some x)" by auto from \x < y\ y_Min have ys_above: "set (y # ys) \ above (Some x)" by auto have eq_xs: "- set xs \ above (Some x) = - set (x # xs) \ above ao" using above_eq by auto have eq_ys: "set (y # ys) \ above (Some x) = set (y # ys) \ above ao" using y_Min \x < y\ xxs_above by(auto intro: above_upclosed) from \x < y\ xs xs_above yys ys_above have "proper_interval_Compl_set_aux (Some x) xs (y # ys) \ (\A. - set xs \ above (Some x) \' A \ above (Some x) \ A \ above (Some x) \' set (y # ys) \ above (Some x))" by(rule "1.IH") also have "\ \ ?rhs" (is "?lhs \ _") proof assume ?lhs then obtain A where "- set xs \ above (Some x) \' A \ above (Some x)" and "A \ above (Some x) \' set (y # ys) \ above (Some x)" by blast moreover have "A \ above (Some x) = (A - {x}) \ above ao" using above_eq by auto ultimately have "- set (x # xs) \ above ao \' (A - {x}) \ above ao" and "(A - {x}) \ above ao \' set (y # ys) \ above ao" using eq_xs eq_ys by simp_all thus ?rhs by blast next assume ?rhs then obtain A where "- set (x # xs) \ above ao \' A \ above ao" and A_less: "A \ above ao \' set (y # ys) \ above ao" by blast moreover have "x \ A" proof assume "x \ A" hence "set (y # ys) \ above ao \' A \ above ao" using y_Min \x < y\ by(auto simp add: above_eq set_less_aux_def intro!: bexI[where x=x]) with A_less show False by(auto dest: set_less_aux_antisym) qed hence "A \ above ao = A \ above (Some x)" using above_eq by auto ultimately show ?lhs using eq_xs eq_ys by auto qed finally show ?thesis using False by simp qed thus ?thesis using True by simp next case False show ?thesis proof(cases "y < x") case True show ?thesis (is "?lhs \ ?rhs") proof(cases "proper_interval ao (Some y)") case True then obtain z where z: "z < y" "z \ above ao" by(auto simp add: proper_interval_Some2) hence nempty: "above ao \ {}" by auto with z have "Min (above ao) \ z" by auto hence "Min (above ao) < y" using \z < y\ by(rule le_less_trans) hence "set (y # ys) \ above ao \' - set (x # xs) \ above ao" using y_Min x_Min z \y < x\ by(fastforce simp add: set_less_aux_def nempty intro!: Min_eqI bexI[where x="Min (above ao)"]) thus ?thesis using True \y < x\ by(auto dest: set_less_aux_trans set_less_aux_antisym) next case False hence above_eq: "above ao = insert y (above (Some y))" using yys_above by(auto simp add: proper_interval_Some2 intro: above_upclosed) from y_Min have ys_above: "set ys \ above (Some y)" by auto from \y < x\ x_Min have xs_above: "set (x # xs) \ above (Some y)" by auto have "y \ - set (x # xs) \ above ao" using \y < x\ x_Min yys_above by auto hence nempty: "- set (x # xs) \ above ao \ {}" by auto have Min_x: "Min (- set (x # xs) \ above ao) = y" using above_eq \y < x\ x_Min by(auto intro!: Min_eqI) have Min_y: "Min (set (y # ys) \ above ao) = y" using y_Min above_eq by(auto intro!: Min_eqI) have eq_xs: "- set (x # xs) \ above ao - {y} = - set (x # xs) \ above (Some y)" by(auto simp add: above_eq) have eq_ys: "set ys \ above ao - {y} = set ys \ above (Some y)" using y_Min above_eq by auto from \\ x < y\ \y < x\ xxs xs_above ys ys_above have "proper_interval_Compl_set_aux (Some y) (x # xs) ys \ (\A. - set (x # xs) \ above (Some y) \' A \ above (Some y) \ A \ above (Some y) \' set ys \ above (Some y))" by(rule "1.IH") also have "\ \ ?rhs" (is "?lhs' \ _") proof assume ?lhs' then obtain A where less_A: "- set (x # xs) \ above (Some y) \' A \ above (Some y)" and A_less: "A \ above (Some y) \' set ys \ above (Some y)" by blast let ?A = "insert y A" have Min_A: "Min (?A \ above ao) = y" using above_eq by(auto intro!: Min_eqI) moreover have A_eq: "A \ above ao - {y} = A \ above (Some y)" using above_eq by auto ultimately have less_A': "- set (x # xs) \ above ao \' ?A \ above ao" using nempty yys_above less_A Min_x eq_xs by(subst set_less_aux_rec) simp_all have A'_less: "?A \ above ao \' set (y # ys) \ above ao" using yys_above nempty Min_A A_eq A_less Min_y eq_ys by(subst set_less_aux_rec) simp_all with less_A' show ?rhs by blast next assume ?rhs then obtain A where less_A: "- set (x # xs) \ above ao \' A \ above ao" and A_less: "A \ above ao \' set (y # ys) \ above ao" by blast from less_A have nempty': "A \ above ao \ {}" by auto moreover have A_eq: "A \ above ao - {y} = A \ above (Some y)" using above_eq by auto moreover have y_in_xxs: "y \ - set (x # xs) \ above ao" using \y < x\ x_Min yys_above by auto moreover have "y \ A" proof(rule ccontr) assume "y \ A" hence "A \ above ao \' - set (x # xs) \ above ao" using \y < x\ x_Min y_in_xxs by(auto simp add: set_less_aux_def above_eq intro: bexI[where x=y]) with less_A show False by(rule set_less_aux_antisym) qed hence Min_A: "Min (A \ above ao) = y" using above_eq y_Min by(auto intro!: Min_eqI) ultimately have less_A': "- set (x # xs) \ above (Some y) \' A \ above (Some y)" using nempty less_A Min_x eq_xs by(subst (asm) set_less_aux_rec)(auto dest: bspec[where x=y]) have A'_less: "A \ above (Some y) \' set ys \ above (Some y)" using A_less nempty' yys_above Min_A Min_y A_eq eq_ys by(subst (asm) set_less_aux_rec) simp_all with less_A' show ?lhs' by blast qed finally show ?thesis using \\ x < y\ \y < x\ False by simp qed next case False with \\ x < y\ have "x = y" by auto thus ?thesis (is "?lhs \ ?rhs") proof(cases "proper_interval ao (Some x)") case True then obtain z where z: "z < x" "z \ above ao" by(auto simp add: proper_interval_Some2) hence nempty: "above ao \ {}" by auto with z have "Min (above ao) \ z" by auto hence "Min (above ao) < x" using \z < x\ by(rule le_less_trans) hence "set (y # ys) \ above ao \' - set (x # xs) \ above ao" using y_Min x_Min z \x = y\ by(fastforce simp add: set_less_aux_def nempty intro!: Min_eqI bexI[where x="Min (above ao)"]) thus ?thesis using True \x = y\ by(auto dest: set_less_aux_trans set_less_aux_antisym) next case False hence above_eq: "above ao = insert x (above (Some x))" using xxs_above by(auto simp add: proper_interval_Some2 intro: above_upclosed) have "(ys = [] \ xs \ []) \ ?rhs" (is "?lhs' \ _") proof(intro iffI strip notI) assume ?lhs' show ?rhs proof(cases ys) case Nil with \?lhs'\ obtain x' xs' where xs_eq: "xs = x' # xs'" by(auto simp add: neq_Nil_conv) with xs have x'_Min: "\x'' \ set xs'. x' < x''" by(auto simp add: less_le) let ?A = "- set (x # xs')" have "- set (x # xs) \ above ao \ ?A \ above ao" using xs_eq by auto moreover have "x' \ - set (x # xs) \ above ao" "x' \ ?A \ above ao" using xs_eq xxs_above x'_Min x_Min by auto ultimately have "- set (x # xs) \ above ao \ ?A \ above ao" by blast hence "- set (x # xs) \ above ao \' \ " by(fastforce intro: psubset_finite_imp_set_less_aux) moreover have "\ \' set (y # ys) \ above ao" using Nil \x = y\ by(auto simp add: set_less_aux_def above_eq) ultimately show ?thesis by blast next case (Cons y' ys') let ?A = "{y}" have "- set (x # xs) \ above ao \' ?A \ above ao" using \x = y\ x_Min by(auto simp add: set_less_aux_def above_eq) moreover have "\ \ set (y # ys) \ above ao" using yys_above yys Cons by auto hence "?A \ above ao \' \" by(fastforce intro: psubset_finite_imp_set_less_aux) ultimately show ?thesis by blast qed next assume Nil: "ys = []" "xs = []" and ?rhs then obtain A where less_A: "- {x} \ above ao \' A \ above ao" and A_less: "A \ above ao \' {x}" using \x = y\ above_eq by auto have "x \ A" using A_less by(auto simp add: set_less_aux_def above_eq) hence "A \ above ao \ - {x} \ above ao" by auto hence "A \ above ao \' \" by(auto intro: subset_finite_imp_set_less_eq_aux) with less_A have "\ \' \" by(rule set_less_trans_set_less_eq) thus False by simp qed with \x = y\ False show ?thesis by simp qed qed qed qed } from this[of None] show ?thesis by simp qed end subsection \Proper intervals for HOL types\ instantiation unit :: proper_interval begin fun proper_interval_unit :: "unit proper_interval" where "proper_interval_unit None None = True" | "proper_interval_unit _ _ = False" instance by intro_classes auto end instantiation bool :: proper_interval begin fun proper_interval_bool :: "bool proper_interval" where "proper_interval_bool (Some x) (Some y) \ False" | "proper_interval_bool (Some x) None \ \ x" | "proper_interval_bool None (Some y) \ y" | "proper_interval_bool None None = True" instance by intro_classes auto end instantiation nat :: proper_interval begin fun proper_interval_nat :: "nat proper_interval" where "proper_interval_nat no None = True" | "proper_interval_nat None (Some x) \ x > 0" | "proper_interval_nat (Some x) (Some y) \ y - x > 1" instance by intro_classes auto end instantiation int :: proper_interval begin fun proper_interval_int :: "int proper_interval" where "proper_interval_int (Some x) (Some y) \ y - x > 1" | "proper_interval_int _ _ = True" instance by intro_classes (auto intro: less_add_one, metis less_add_one minus_less_iff) end instantiation integer :: proper_interval begin context includes integer.lifting begin lift_definition proper_interval_integer :: "integer proper_interval" is "proper_interval" . instance by(intro_classes)(transfer, simp only: proper_interval_simps)+ end end lemma proper_interval_integer_simps [code]: includes integer.lifting fixes x y :: integer and xo yo :: "integer option" shows "proper_interval (Some x) (Some y) = (1 < y - x)" "proper_interval None yo = True" "proper_interval xo None = True" by(transfer, simp)+ instantiation natural :: proper_interval begin context includes natural.lifting begin lift_definition proper_interval_natural :: "natural proper_interval" is "proper_interval" . instance by(intro_classes)(transfer, simp only: proper_interval_simps)+ end end lemma proper_interval_natural_simps [code]: includes natural.lifting fixes x y :: natural and xo :: "natural option" shows "proper_interval xo None = True" "proper_interval None (Some y) \ y > 0" "proper_interval (Some x) (Some y) \ y - x > 1" by(transfer, simp)+ lemma char_less_iff_nat_of_char: "x < y \ of_char x < (of_char y :: nat)" by (fact less_char_def) lemma nat_of_char_inject [simp]: "of_char x = (of_char y :: nat) \ x = y" by (fact of_char_eq_iff) lemma char_le_iff_nat_of_char: "x \ y \ of_char x \ (of_char y :: nat)" by (fact less_eq_char_def) instantiation char :: proper_interval begin fun proper_interval_char :: "char proper_interval" where "proper_interval_char None None \ True" | "proper_interval_char None (Some x) \ x \ CHR 0x00" | "proper_interval_char (Some x) None \ x \ CHR 0xFF" | "proper_interval_char (Some x) (Some y) \ of_char y - of_char x > (1 :: nat)" instance proof fix y :: char { assume "y \ CHR 0x00" have "CHR 0x00 < y" proof (rule ccontr) assume "\ CHR 0x00 < y" then have "of_char y = (of_char CHR 0x00 :: nat)" by (simp add: not_less char_le_iff_nat_of_char) then have "y = CHR 0x00" using nat_of_char_inject [of y "CHR 0x00"] by simp with \y \ CHR 0x00\ show False by simp qed } moreover { fix z :: char assume "z < CHR 0x00" hence False by (simp add: char_less_iff_nat_of_char of_char_eq_iff [symmetric]) } ultimately show "proper_interval None (Some y) = (\z. z < y)" by auto fix x :: char { assume "x \ CHR 0xFF" then have "x < CHR 0xFF" by (auto simp add: neq_iff char_less_iff_nat_of_char) (insert nat_of_char_less_256 [of x], simp) hence "\z. x < z" .. } moreover { fix z :: char assume "CHR 0xFF < z" hence "False" by (simp add: char_less_iff_nat_of_char) (insert nat_of_char_less_256 [of z], simp) } ultimately show "proper_interval (Some x) None = (\z. x < z)" by auto { assume gt: "of_char y - of_char x > (1 :: nat)" let ?z = "char_of (of_char x + (1 :: nat))" from gt nat_of_char_less_256 [of y] have 255: "of_char x < (255 :: nat)" by arith with gt have "x < ?z" "?z < y" by (simp_all add: char_less_iff_nat_of_char) hence "\z. x < z \ z < y" by blast } moreover { fix z assume "x < z" "z < y" hence "(1 :: nat) < of_char y - of_char x" by (simp add: char_less_iff_nat_of_char) } ultimately show "proper_interval (Some x) (Some y) = (\z>x. z < y)" by auto qed simp end instantiation Enum.finite_1 :: proper_interval begin definition proper_interval_finite_1 :: "Enum.finite_1 proper_interval" where "proper_interval_finite_1 x y \ x = None \ y = None" instance by intro_classes (simp_all add: proper_interval_finite_1_def less_finite_1_def) end instantiation Enum.finite_2 :: proper_interval begin fun proper_interval_finite_2 :: "Enum.finite_2 proper_interval" where "proper_interval_finite_2 None None \ True" | "proper_interval_finite_2 None (Some x) \ x = finite_2.a\<^sub>2" | "proper_interval_finite_2 (Some x) None \ x = finite_2.a\<^sub>1" | "proper_interval_finite_2 (Some x) (Some y) \ False" instance by intro_classes (auto simp add: less_finite_2_def) end instantiation Enum.finite_3 :: proper_interval begin fun proper_interval_finite_3 :: "Enum.finite_3 proper_interval" where "proper_interval_finite_3 None None \ True" | "proper_interval_finite_3 None (Some x) \ x \ finite_3.a\<^sub>1" | "proper_interval_finite_3 (Some x) None \ x \ finite_3.a\<^sub>3" | "proper_interval_finite_3 (Some x) (Some y) \ x = finite_3.a\<^sub>1 \ y = finite_3.a\<^sub>3" instance proof fix x y :: Enum.finite_3 show "proper_interval None (Some y) = (\z. z < y)" by(cases y)(auto simp add: less_finite_3_def split: finite_3.split) show "proper_interval (Some x) None = (\z. x < z)" by(cases x)(auto simp add: less_finite_3_def) show "proper_interval (Some x) (Some y) = (\z>x. z < y)" by(auto simp add: less_finite_3_def split: finite_3.split_asm) qed simp end subsection \List fusion for the order and proper intervals on @{typ "'a set"}\ definition length_last_fusion :: "('a, 's) generator \ 's \ nat \ 'a" where "length_last_fusion g s = length_last (list.unfoldr g s)" lemma length_last_fusion_code [code]: "length_last_fusion g s = (if list.has_next g s then let (x, s') = list.next g s in fold_fusion g (\x (n, _). (n + 1, x)) s' (1, x) else (0, undefined))" unfolding length_last_fusion_def by(subst list.unfoldr.simps)(simp add: length_last_Nil length_last_Cons_code fold_fusion_def split_beta) declare length_last_fusion_def [symmetric, code_unfold] context proper_intrvl begin definition set_less_eq_aux_Compl_fusion :: "('a, 's1) generator \ ('a, 's2) generator \ 'a option \ 's1 \ 's2 \ bool" where "set_less_eq_aux_Compl_fusion g1 g2 ao s1 s2 = set_less_eq_aux_Compl ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)" definition Compl_set_less_eq_aux_fusion :: "('a, 's1) generator \ ('a, 's2) generator \ 'a option \ 's1 \ 's2 \ bool" where "Compl_set_less_eq_aux_fusion g1 g2 ao s1 s2 = Compl_set_less_eq_aux ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)" definition set_less_aux_Compl_fusion :: "('a, 's1) generator \ ('a, 's2) generator \ 'a option \ 's1 \ 's2 \ bool" where "set_less_aux_Compl_fusion g1 g2 ao s1 s2 = set_less_aux_Compl ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)" definition Compl_set_less_aux_fusion :: "('a, 's1) generator \ ('a, 's2) generator \ 'a option \ 's1 \ 's2 \ bool" where "Compl_set_less_aux_fusion g1 g2 ao s1 s2 = Compl_set_less_aux ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)" definition exhaustive_above_fusion :: "('a, 's) generator \ 'a \ 's \ bool" where "exhaustive_above_fusion g a s = exhaustive_above a (list.unfoldr g s)" definition exhaustive_fusion :: "('a, 's) generator \ 's \ bool" where "exhaustive_fusion g s = exhaustive (list.unfoldr g s)" definition proper_interval_set_aux_fusion :: "('a, 's1) generator \ ('a, 's2) generator \ 's1 \ 's2 \ bool" where "proper_interval_set_aux_fusion g1 g2 s1 s2 = proper_interval_set_aux (list.unfoldr g1 s1) (list.unfoldr g2 s2)" definition proper_interval_set_Compl_aux_fusion :: "('a, 's1) generator \ ('a, 's2) generator \ 'a option \ nat \ 's1 \ 's2 \ bool" where "proper_interval_set_Compl_aux_fusion g1 g2 ao n s1 s2 = proper_interval_set_Compl_aux ao n (list.unfoldr g1 s1) (list.unfoldr g2 s2)" definition proper_interval_Compl_set_aux_fusion :: "('a, 's1) generator \ ('a, 's2) generator \ 'a option \ 's1 \ 's2 \ bool" where "proper_interval_Compl_set_aux_fusion g1 g2 ao s1 s2 = proper_interval_Compl_set_aux ao (list.unfoldr g1 s1) (list.unfoldr g2 s2)" lemma set_less_eq_aux_Compl_fusion_code: "set_less_eq_aux_Compl_fusion g1 g2 ao s1 s2 \ (list.has_next g1 s1 \ list.has_next g2 s2 \ (let (x, s1') = list.next g1 s1; (y, s2') = list.next g2 s2 in if x < y then proper_interval ao (Some x) \ set_less_eq_aux_Compl_fusion g1 g2 (Some x) s1' s2 else if y < x then proper_interval ao (Some y) \ set_less_eq_aux_Compl_fusion g1 g2 (Some y) s1 s2' else proper_interval ao (Some y)))" unfolding set_less_eq_aux_Compl_fusion_def by(subst (1 2 4 5) list.unfoldr.simps)(simp add: split_beta) lemma Compl_set_less_eq_aux_fusion_code: "Compl_set_less_eq_aux_fusion g1 g2 ao s1 s2 \ (if list.has_next g1 s1 then let (x, s1') = list.next g1 s1 in if list.has_next g2 s2 then let (y, s2') = list.next g2 s2 in if x < y then \ proper_interval ao (Some x) \ Compl_set_less_eq_aux_fusion g1 g2 (Some x) s1' s2 else if y < x then \ proper_interval ao (Some y) \ Compl_set_less_eq_aux_fusion g1 g2 (Some y) s1 s2' else \ proper_interval ao (Some y) else \ proper_interval ao (Some x) \ Compl_set_less_eq_aux_fusion g1 g2 (Some x) s1' s2 else if list.has_next g2 s2 then let (y, s2') = list.next g2 s2 in \ proper_interval ao (Some y) \ Compl_set_less_eq_aux_fusion g1 g2 (Some y) s1 s2' else \ proper_interval ao None)" unfolding Compl_set_less_eq_aux_fusion_def by(subst (1 2 4 5 8 9) list.unfoldr.simps)(simp add: split_beta) lemma set_less_aux_Compl_fusion_code: "set_less_aux_Compl_fusion g1 g2 ao s1 s2 \ (if list.has_next g1 s1 then let (x, s1') = list.next g1 s1 in if list.has_next g2 s2 then let (y, s2') = list.next g2 s2 in if x < y then proper_interval ao (Some x) \ set_less_aux_Compl_fusion g1 g2 (Some x) s1' s2 else if y < x then proper_interval ao (Some y) \ set_less_aux_Compl_fusion g1 g2 (Some y) s1 s2' else proper_interval ao (Some y) else proper_interval ao (Some x) \ set_less_aux_Compl_fusion g1 g2 (Some x) s1' s2 else if list.has_next g2 s2 then let (y, s2') = list.next g2 s2 in proper_interval ao (Some y) \ set_less_aux_Compl_fusion g1 g2 (Some y) s1 s2' else proper_interval ao None)" unfolding set_less_aux_Compl_fusion_def by(subst (1 2 4 5 8 9) list.unfoldr.simps)(simp add: split_beta) lemma Compl_set_less_aux_fusion_code: "Compl_set_less_aux_fusion g1 g2 ao s1 s2 \ list.has_next g1 s1 \ list.has_next g2 s2 \ (let (x, s1') = list.next g1 s1; (y, s2') = list.next g2 s2 in if x < y then \ proper_interval ao (Some x) \ Compl_set_less_aux_fusion g1 g2 (Some x) s1' s2 else if y < x then \ proper_interval ao (Some y) \ Compl_set_less_aux_fusion g1 g2 (Some y) s1 s2' else \ proper_interval ao (Some y))" unfolding Compl_set_less_aux_fusion_def by(subst (1 2 4 5) list.unfoldr.simps)(simp add: split_beta) lemma exhaustive_above_fusion_code: "exhaustive_above_fusion g y s \ (if list.has_next g s then let (x, s') = list.next g s in \ proper_interval (Some y) (Some x) \ exhaustive_above_fusion g x s' else \ proper_interval (Some y) None)" unfolding exhaustive_above_fusion_def by(subst list.unfoldr.simps)(simp add: split_beta) lemma exhaustive_fusion_code: "exhaustive_fusion g s = (list.has_next g s \ (let (x, s') = list.next g s in \ proper_interval None (Some x) \ exhaustive_above_fusion g x s'))" unfolding exhaustive_fusion_def exhaustive_above_fusion_def by(subst (1) list.unfoldr.simps)(simp add: split_beta) lemma proper_interval_set_aux_fusion_code: "proper_interval_set_aux_fusion g1 g2 s1 s2 \ list.has_next g2 s2 \ (let (y, s2') = list.next g2 s2 in if list.has_next g1 s1 then let (x, s1') = list.next g1 s1 in if x < y then False else if y < x then proper_interval (Some y) (Some x) \ list.has_next g2 s2' \ \ exhaustive_above_fusion g1 x s1' else proper_interval_set_aux_fusion g1 g2 s1' s2' else list.has_next g2 s2' \ proper_interval (Some y) None)" unfolding proper_interval_set_aux_fusion_def exhaustive_above_fusion_def by(subst (1 2) list.unfoldr.simps)(simp add: split_beta) lemma proper_interval_set_Compl_aux_fusion_code: "proper_interval_set_Compl_aux_fusion g1 g2 ao n s1 s2 \ (if list.has_next g1 s1 then let (x, s1') = list.next g1 s1 in if list.has_next g2 s2 then let (y, s2') = list.next g2 s2 in if x < y then proper_interval ao (Some x) \ proper_interval_set_Compl_aux_fusion g1 g2 (Some x) (n + 1) s1' s2 else if y < x then proper_interval ao (Some y) \ proper_interval_set_Compl_aux_fusion g1 g2 (Some y) (n + 1) s1 s2' else proper_interval ao (Some x) \ (let m = CARD('a) - n in m - length_fusion g2 s2' \ 2 \ m - length_fusion g1 s1' \ 2) else let m = CARD('a) - n; (len_x, x') = length_last_fusion g1 s1 in m \ len_x \ (m = len_x + 1 \ \ proper_interval (Some x') None) else if list.has_next g2 s2 then let (y, s2') = list.next g2 s2; m = CARD('a) - n; (len_y, y') = length_last_fusion g2 s2 in m \ len_y \ (m = len_y + 1 \ \ proper_interval (Some y') None) else CARD('a) > n + 1)" unfolding proper_interval_set_Compl_aux_fusion_def length_last_fusion_def length_fusion_def by(subst (1 2 4 5 9 10) list.unfoldr.simps)(simp add: split_beta) lemma proper_interval_Compl_set_aux_fusion_code: "proper_interval_Compl_set_aux_fusion g1 g2 ao s1 s2 \ list.has_next g1 s1 \ list.has_next g2 s2 \ (let (x, s1') = list.next g1 s1; (y, s2') = list.next g2 s2 in if x < y then \ proper_interval ao (Some x) \ proper_interval_Compl_set_aux_fusion g1 g2 (Some x) s1' s2 else if y < x then \ proper_interval ao (Some y) \ proper_interval_Compl_set_aux_fusion g1 g2 (Some y) s1 s2' else \ proper_interval ao (Some x) \ (list.has_next g2 s2' \ list.has_next g1 s1'))" unfolding proper_interval_Compl_set_aux_fusion_def by(subst (1 2 4 5) list.unfoldr.simps)(auto simp add: split_beta) end lemmas [code] = set_less_eq_aux_Compl_fusion_code proper_intrvl.set_less_eq_aux_Compl_fusion_code Compl_set_less_eq_aux_fusion_code proper_intrvl.Compl_set_less_eq_aux_fusion_code set_less_aux_Compl_fusion_code proper_intrvl.set_less_aux_Compl_fusion_code Compl_set_less_aux_fusion_code proper_intrvl.Compl_set_less_aux_fusion_code exhaustive_above_fusion_code proper_intrvl.exhaustive_above_fusion_code exhaustive_fusion_code proper_intrvl.exhaustive_fusion_code proper_interval_set_aux_fusion_code proper_intrvl.proper_interval_set_aux_fusion_code proper_interval_set_Compl_aux_fusion_code proper_intrvl.proper_interval_set_Compl_aux_fusion_code proper_interval_Compl_set_aux_fusion_code proper_intrvl.proper_interval_Compl_set_aux_fusion_code lemmas [symmetric, code_unfold] = set_less_eq_aux_Compl_fusion_def proper_intrvl.set_less_eq_aux_Compl_fusion_def Compl_set_less_eq_aux_fusion_def proper_intrvl.Compl_set_less_eq_aux_fusion_def set_less_aux_Compl_fusion_def proper_intrvl.set_less_aux_Compl_fusion_def Compl_set_less_aux_fusion_def proper_intrvl.Compl_set_less_aux_fusion_def exhaustive_above_fusion_def proper_intrvl.exhaustive_above_fusion_def exhaustive_fusion_def proper_intrvl.exhaustive_fusion_def proper_interval_set_aux_fusion_def proper_intrvl.proper_interval_set_aux_fusion_def proper_interval_set_Compl_aux_fusion_def proper_intrvl.proper_interval_set_Compl_aux_fusion_def proper_interval_Compl_set_aux_fusion_def proper_intrvl.proper_interval_Compl_set_aux_fusion_def subsection \Drop notation\ context ord begin no_notation set_less_aux (infix "\''" 50) and set_less_eq_aux (infix "\''" 50) and set_less_eq_aux' (infix "\''''" 50) and set_less_eq_aux'' (infix "\''''''" 50) and set_less_eq (infix "\" 50) and set_less (infix "\" 50) end end diff --git a/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy b/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy --- a/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy +++ b/thys/Extended_Finite_State_Machine_Inference/code-targets/Code_Target_FSet.thy @@ -1,164 +1,164 @@ section\Finite Sets\ text\Here we define the operations on the \texttt{fset} datatype in terms of lists rather than sets. This allows the Scala implementation to skip a case match each time, which makes for cleaner and slightly faster code.\ theory Code_Target_FSet imports "Extended_Finite_State_Machines.FSet_Utils" begin code_datatype fset_of_list declare FSet.fset_of_list.rep_eq [code] lemma fprod_code [code]: "fprod (fset_of_list xs) (fset_of_list ys) = fset_of_list (remdups [(x, y). x \ xs, y \ ys])" apply (simp add: fprod_def fset_of_list_def fset_both_sides Abs_fset_inverse) by auto lemma fminus_fset_filter [code]: "fset_of_list A - xs = fset_of_list (remdups (filter (\x. x |\| xs) A))" by auto lemma sup_fset_fold [code]: "(fset_of_list f1) |\| (fset_of_list f2) = fset_of_list (remdups (f1@f2))" by simp lemma bot_fset [code]: "{||} = fset_of_list []" by simp lemma finsert [code]: "finsert a (fset_of_list as) = fset_of_list (List.insert a as)" by (simp add: List.insert_def finsert_absorb fset_of_list_elem) lemma ffilter_filter [code]: "ffilter f (fset_of_list as) = fset_of_list (List.filter f (remdups as))" by simp lemma fimage_map [code]: "fimage f (fset_of_list as) = fset_of_list (List.map f (remdups as))" by simp lemma ffUnion_fold [code]: "ffUnion (fset_of_list as) = fold (|\|) as {||}" by (simp add: fold_union_ffUnion) lemma fmember [code]: "a |\| (fset_of_list as) = List.member as a" by (simp add: fset_of_list_elem member_def) lemma fthe_elem [code]: "fthe_elem (fset_of_list [x]) = x" by simp lemma size [code]: "size (fset_of_list as) = length (remdups as)" proof(induct as) case Nil then show ?case by simp next case (Cons a as) then show ?case by (simp add: fset_of_list.rep_eq insert_absorb) qed lemma fMax_fold [code]: "fMax (fset_of_list (a#as)) = fold max as a" by (metis Max.set_eq_fold fMax.F.rep_eq fset_of_list.rep_eq) lemma fMin_fold [code]: "fMin (fset_of_list (h#t)) = fold min t h" apply (simp add: fset_of_list_def) by (metis Min.set_eq_fold fMin_Min fset_of_list.abs_eq list.simps(15)) lemma fremove_code [code]: "fremove a (fset_of_list A) = fset_of_list (filter (\x. x \ a) A)" apply (simp add: fremove_def minus_fset_def ffilter_def fset_both_sides Abs_fset_inverse fset_of_list.rep_eq) by auto lemma fsubseteq [code]: "(fset_of_list l) |\| A = List.list_all (\x. x |\| A) l" by (induct l, auto) lemma fsum_fold [code]: "fSum (fset_of_list l) = fold (+) (remdups l) 0" proof(induct l) case Nil then show ?case by (simp add: fsum.F.rep_eq fSum_def) next case (Cons a l) then show ?case apply simp apply standard apply (simp add: finsert_absorb fset_of_list_elem) by (simp add: add.commute fold_plus_sum_list_rev fset_of_list.rep_eq fsum.F.rep_eq fSum_def) qed lemma code_fset_eq [code]: "HOL.equal X (fset_of_list Y) \ size X = length (remdups Y) \ (\x |\| X. List.member Y x)" apply (simp only: HOL.equal_class.equal_eq fset_eq_alt) apply (simp only: size) using fmember by fastforce lemma code_fsubset [code]: "s |\| s' = (s |\| s' \ size s < size s')" apply standard apply (simp only: size_fsubset) by auto lemma code_fset [code]: "fset (fset_of_list l) = fold insert l {}" using fset_of_list.rep_eq union_set_fold by fastforce lemma code_fBall [code]: "fBall (fset_of_list l) f = list_all f l" by (simp add: Ball_set fBall.rep_eq fset_of_list.rep_eq) lemma code_fBex [code]: "fBex (fset_of_list l) f = list_ex f l" by (meson Bex_set fBexE fset_of_list_elem rev_fBexI) definition "nativeSort = sort" code_printing constant nativeSort \ (Scala) "_.sortWith((Orderings.less))" lemma [code]: "sorted_list_of_fset (fset_of_list l) = nativeSort (remdups l)" by (simp add: nativeSort_def sorted_list_of_fset_sort) lemma [code]: "sorted_list_of_set (set l) = nativeSort (remdups l)" by (simp add: nativeSort_def sorted_list_of_set_sort_remdups) lemma [code]: "fMin (fset_of_list (h#t)) = hd (nativeSort (h#t))" by (metis fMin_Min hd_sort_Min list.distinct(1) nativeSort_def) lemma sorted_Max_Cons: "l \ [] \ sorted (a#l) \ Max (set (a#l)) = Max (set l)" using eq_iff by fastforce lemma sorted_Max: "l \ [] \ sorted l \ Max (set l) = hd (rev l)" proof(induct l) case Nil then show ?case by simp next case (Cons a l) then show ?case - by (metis sorted_Max_Cons Max_singleton hd_rev last.simps list.set(1) list.simps(15) sorted.simps(2)) + by (metis sorted_Max_Cons Max_singleton hd_rev last.simps list.set(1) list.simps(15) sorted_simps(2)) qed lemma [code]: "fMax (fset_of_list (h#t)) = last (nativeSort (h#t))" by (metis Max.set_eq_fold fMax_fold hd_rev list.simps(3) nativeSort_def set_empty2 set_sort sorted_Max sorted_sort) definition "list_max l = fold max l" code_printing constant list_max \ (Scala) "_.par.fold((_))(Orderings.max)" lemma [code]: "fMax (fset_of_list (h#t)) = list_max t h" by (metis fMax_fold list_max_def) definition "list_min l = fold min l" code_printing constant list_min \ (Scala) "_.par.fold((_))(Orderings.min)" lemma [code]: "fMin (fset_of_list (h#t)) = list_min t h" by (metis fMin_fold list_min_def) lemma fis_singleton_code [code]: "fis_singleton s = (size s = 1)" apply (simp add: fis_singleton_def is_singleton_def) by (simp add: card_Suc_eq) end diff --git a/thys/Extended_Finite_State_Machines/FSet_Utils.thy b/thys/Extended_Finite_State_Machines/FSet_Utils.thy --- a/thys/Extended_Finite_State_Machines/FSet_Utils.thy +++ b/thys/Extended_Finite_State_Machines/FSet_Utils.thy @@ -1,341 +1,341 @@ section\FSet Utilities\ text\This theory provides various additional lemmas, definitions, and syntax over the fset data type.\ theory FSet_Utils imports "HOL-Library.FSet" begin notation (latex output) "FSet.fempty" ("\") and "FSet.fmember" ("\") syntax (ASCII) "_fBall" :: "pttrn \ 'a fset \ bool \ bool" ("(3ALL (_/:_)./ _)" [0, 0, 10] 10) "_fBex" :: "pttrn \ 'a fset \ bool \ bool" ("(3EX (_/:_)./ _)" [0, 0, 10] 10) "_fBex1" :: "pttrn \ 'a fset \ bool \ bool" ("(3EX! (_/:_)./ _)" [0, 0, 10] 10) syntax (input) "_fBall" :: "pttrn \ 'a fset \ bool \ bool" ("(3! (_/:_)./ _)" [0, 0, 10] 10) "_fBex" :: "pttrn \ 'a fset \ bool \ bool" ("(3? (_/:_)./ _)" [0, 0, 10] 10) "_fBex1" :: "pttrn \ 'a fset \ bool \ bool" ("(3?! (_/:_)./ _)" [0, 0, 10] 10) syntax "_fBall" :: "pttrn \ 'a fset \ bool \ bool" ("(3\(_/|\|_)./ _)" [0, 0, 10] 10) "_fBex" :: "pttrn \ 'a fset \ bool \ bool" ("(3\(_/|\|_)./ _)" [0, 0, 10] 10) "_fBnex" :: "pttrn \ 'a fset \ bool \ bool" ("(3\(_/|\|_)./ _)" [0, 0, 10] 10) "_fBex1" :: "pttrn \ 'a fset \ bool \ bool" ("(3\!(_/|\|_)./ _)" [0, 0, 10] 10) translations "\x|\|A. P" \ "CONST fBall A (\x. P)" "\x|\|A. P" \ "CONST fBex A (\x. P)" "\x|\|A. P" \ "CONST fBall A (\x. \P)" "\!x|\|A. P" \ "\!x. x |\| A \ P" lemma fset_of_list_remdups [simp]: "fset_of_list (remdups l) = fset_of_list l" apply (induct l) apply simp by (simp add: finsert_absorb fset_of_list_elem) definition "fSum \ fsum (\x. x)" lemma fset_both_sides: "(Abs_fset s = f) = (fset (Abs_fset s) = fset f)" by (simp add: fset_inject) lemma Abs_ffilter: "(ffilter f s = s') = ({e \ (fset s). f e} = (fset s'))" by (simp add: ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def) lemma size_ffilter_card: "size (ffilter f s) = card ({e \ (fset s). f e})" by (simp add: ffilter_def fset_both_sides Abs_fset_inverse Set.filter_def) lemma ffilter_empty [simp]: "ffilter f {||} = {||}" by auto lemma ffilter_finsert: "ffilter f (finsert a s) = (if f a then finsert a (ffilter f s) else (ffilter f s))" apply simp apply standard apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse) apply auto[1] apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse) by auto lemma fset_equiv: "(f1 = f2) = (fset f1 = fset f2)" by (simp add: fset_inject) lemma finsert_equiv: "(finsert e f = f') = (insert e (fset f) = (fset f'))" by (simp add: finsert_def fset_both_sides Abs_fset_inverse) lemma filter_elements: "x |\| Abs_fset (Set.filter f (fset s)) = (x \ (Set.filter f (fset s)))" by (metis ffilter.rep_eq fset_inverse notin_fset) lemma sorted_list_of_fempty [simp]: "sorted_list_of_fset {||} = []" by (simp add: sorted_list_of_fset_def) lemma fmember_implies_member: "e |\| f \ e \ fset f" by (simp add: fmember_def) lemma fold_union_ffUnion: "fold (|\|) l {||} = ffUnion (fset_of_list l)" by(induct l rule: rev_induct, auto) lemma filter_filter: "ffilter P (ffilter Q xs) = ffilter (\x. Q x \ P x) xs" by auto lemma fsubset_strict: "x2 |\| x1 \ \e. e |\| x1 \ e |\| x2" by auto lemma fsubset: "x2 |\| x1 \ \e. e |\| x2 \ e |\| x1" by auto lemma size_fsubset_elem: assumes "\e. e |\| x1 \ e |\| x2" and "\e. e |\| x2 \ e |\| x1" shows "size x2 < size x1" using assms apply (simp add: fmember_def) by (metis card_seteq finite_fset linorder_not_le subsetI) lemma size_fsubset: "x2 |\| x1 \ size x2 < size x1" by (metis fsubset fsubset_strict size_fsubset_elem) definition fremove :: "'a \ 'a fset \ 'a fset" where [code_abbrev]: "fremove x A = A - {|x|}" lemma arg_cong_ffilter: "\e |\| f. p e = p' e \ ffilter p f = ffilter p' f" by auto lemma ffilter_singleton: "f e \ ffilter f {|e|} = {|e|}" apply (simp add: ffilter_def fset_both_sides Abs_fset_inverse) by auto lemma fset_eq_alt: "(x = y) = (x |\| y \ size x = size y)" by (metis exists_least_iff le_less size_fsubset) lemma ffold_empty [simp]: "ffold f b {||} = b" by (simp add: ffold_def) lemma sorted_list_of_fset_sort: "sorted_list_of_fset (fset_of_list l) = sort (remdups l)" by (simp add: fset_of_list.rep_eq sorted_list_of_fset.rep_eq sorted_list_of_set_sort_remdups) lemma fMin_Min: "fMin (fset_of_list l) = Min (set l)" by (simp add: fMin.F.rep_eq fset_of_list.rep_eq) lemma sorted_hd_Min: "sorted l \ l \ [] \ hd l = Min (set l)" - by (metis List.finite_set Min_eqI eq_iff hd_Cons_tl insertE list.set_sel(1) list.simps(15) sorted.simps(2)) + by (metis List.finite_set Min_eqI eq_iff hd_Cons_tl insertE list.set_sel(1) list.simps(15) sorted_simps(2)) lemma hd_sort_Min: "l \ [] \ hd (sort l) = Min (set l)" by (metis sorted_hd_Min set_empty set_sort sorted_sort) lemma hd_sort_remdups: "hd (sort (remdups l)) = hd (sort l)" by (metis hd_sort_Min remdups_eq_nil_iff set_remdups) lemma exists_fset_of_list: "\l. f = fset_of_list l" using exists_fset_of_list by fastforce lemma hd_sorted_list_of_fset: "s \ {||} \ hd (sorted_list_of_fset s) = (fMin s)" apply (insert exists_fset_of_list[of s]) apply (erule exE) apply simp apply (simp add: sorted_list_of_fset_sort fMin_Min hd_sort_remdups) by (metis fset_of_list_simps(1) hd_sort_Min) lemma fminus_filter_singleton: "fset_of_list l |-| {|x|} = fset_of_list (filter (\e. e \ x) l)" by auto lemma card_minus_fMin: "s \ {||} \ card (fset s - {fMin s}) < card (fset s)" by (metis Min_in bot_fset.rep_eq card_Diff1_less fMin.F.rep_eq finite_fset fset_equiv) (* Provides a deterministic way to fold fsets similar to List.fold that works with the code generator *) function ffold_ord :: "(('a::linorder) \ 'b \ 'b) \ 'a fset \ 'b \ 'b" where "ffold_ord f s b = ( if s = {||} then b else let h = fMin s; t = s - {|h|} in ffold_ord f t (f h b) )" by auto termination apply (relation "measures [\(a, s, ab). size s]") apply simp by (simp add: card_minus_fMin) lemma sorted_list_of_fset_Cons: "\h t. (sorted_list_of_fset (finsert s ss)) = h#t" apply (simp add: sorted_list_of_fset_def) by (cases "insort s (sorted_list_of_set (fset ss - {s}))", auto) lemma list_eq_hd_tl: "l \ [] \ hd l = h \ tl l = t \ l = (h#t)" by auto lemma fset_of_list_sort: "fset_of_list l = fset_of_list (sort l)" by (simp add: fset_of_list.abs_eq) lemma exists_sorted_distinct_fset_of_list: "\l. sorted l \ distinct l \ f = fset_of_list l" by (metis distinct_sorted_list_of_set sorted_list_of_fset.rep_eq sorted_list_of_fset_simps(2) sorted_sorted_list_of_set) lemma fset_of_list_empty [simp]: "(fset_of_list l = {||}) = (l = [])" by (metis fset_of_list.rep_eq fset_of_list_simps(1) set_empty) lemma ffold_ord_cons: assumes sorted: "sorted (h#t)" and distinct: "distinct (h#t)" shows "ffold_ord f (fset_of_list (h#t)) b = ffold_ord f (fset_of_list t) (f h b)" proof- have h_is_min: "h = fMin (fset_of_list (h#t))" by (metis sorted fMin_Min list.sel(1) list.simps(3) sorted_hd_Min) have remove_min: "fset_of_list t = (fset_of_list (h#t)) - {|h|}" using distinct fset_of_list_elem by force show ?thesis apply (simp only: ffold_ord.simps[of f "fset_of_list (h#t)"]) by (metis h_is_min remove_min fset_of_list_empty list.distinct(1)) qed lemma sorted_distinct_ffold_ord: assumes "sorted l" and "distinct l" shows "ffold_ord f (fset_of_list l) b = fold f l b" using assms apply (induct l arbitrary: b) apply simp - by (metis distinct.simps(2) ffold_ord_cons fold_simps(2) sorted.simps(2)) + by (metis distinct.simps(2) ffold_ord_cons fold_simps(2) sorted_simps(2)) lemma ffold_ord_fold_sorted: "ffold_ord f s b = fold f (sorted_list_of_fset s) b" by (metis exists_sorted_distinct_fset_of_list sorted_distinct_ffold_ord distinct_remdups_id sorted_list_of_fset_sort sorted_sort_id) context includes fset.lifting begin lift_definition fprod :: "'a fset \ 'b fset \ ('a \ 'b) fset " (infixr "|\|" 80) is "\a b. fset a \ fset b" by simp lift_definition fis_singleton :: "'a fset \ bool" is "\A. is_singleton (fset A)". end lemma fprod_empty_l: "{||} |\| a = {||}" using bot_fset_def fprod.abs_eq by force lemma fprod_empty_r: "a |\| {||} = {||}" by (simp add: fprod_def bot_fset_def Abs_fset_inverse) lemmas fprod_empty = fprod_empty_l fprod_empty_r lemma fprod_finsert: "(finsert a as) |\| (finsert b bs) = finsert (a, b) (fimage (\b. (a, b)) bs |\| fimage (\a. (a, b)) as |\| (as |\| bs))" apply (simp add: fprod_def fset_both_sides Abs_fset_inverse) by auto lemma fprod_member: "x |\| xs \ y |\| ys \ (x, y) |\| xs |\| ys" by (simp add: fmember_def fprod_def Abs_fset_inverse) lemma fprod_subseteq: "x |\| x' \ y |\| y' \ x |\| y |\| x' |\| y'" apply (simp add: fprod_def less_eq_fset_def Abs_fset_inverse) by auto lemma fimage_fprod: "(a, b) |\| A |\| B \ f a b |\| (\(x, y). f x y) |`| (A |\| B)" by force lemma fprod_singletons: "{|a|} |\| {|b|} = {|(a, b)|}" apply (simp add: fprod_def) by (metis fset_inverse fset_simps(1) fset_simps(2)) lemma fprod_equiv: "(fset (f |\| f') = s) = (((fset f) \ (fset f')) = s)" by (simp add: fprod_def Abs_fset_inverse) lemma fis_singleton_alt: "fis_singleton f = (\e. f = {|e|})" by (metis fis_singleton.rep_eq fset_inverse fset_simps(1) fset_simps(2) is_singleton_def) lemma singleton_singleton [simp]: "fis_singleton {|a|}" by (simp add: fis_singleton_def) lemma not_singleton_empty [simp]: "\ fis_singleton {||}" apply (simp add: fis_singleton_def) by (simp add: is_singleton_altdef) lemma fis_singleton_fthe_elem: "fis_singleton A \ A = {|fthe_elem A|}" by (metis fis_singleton_alt fthe_felem_eq) lemma fBall_ffilter: "\x |\| X. f x \ ffilter f X = X" by auto lemma fBall_ffilter2: "X = Y \ \x |\| X. f x \ ffilter f X = Y" by auto lemma size_fset_of_list: "size (fset_of_list l) = length (remdups l)" apply (induct l) apply simp by (simp add: fset_of_list.rep_eq insert_absorb) lemma size_fsingleton: "(size f = 1) = (\e. f = {|e|})" apply (insert exists_fset_of_list[of f]) apply clarify apply (simp only: size_fset_of_list) apply (simp add: fset_of_list_def fset_both_sides Abs_fset_inverse) by (metis List.card_set One_nat_def card.insert card_1_singletonE card.empty empty_iff finite.intros(1)) lemma ffilter_mono: "(ffilter X xs = f) \ \x |\| xs. X x = Y x \ (ffilter Y xs = f)" by auto lemma size_fimage: "size (fimage f s) \ size s" apply (induct s) apply simp by (simp add: card_insert_if) lemma size_ffilter: "size (ffilter P f) \ size f" apply (induct f) apply simp apply (simp only: ffilter_finsert) apply (case_tac "P x") apply (simp add: fmember.rep_eq) by (simp add: card_insert_if) lemma fimage_size_le: "\f s. size s \ n \ size (fimage f s) \ n" using le_trans size_fimage by blast lemma ffilter_size_le: "\f s. size s \ n \ size (ffilter f s) \ n" using dual_order.trans size_ffilter by blast lemma set_membership_eq: "A = B \ (\x. Set.member x A) = (\x. Set.member x B)" apply standard apply simp by (meson equalityI subsetI) lemmas ffilter_eq_iff = Abs_ffilter set_membership_eq fun_eq_iff lemma size_le_1: "size f \ 1 = (f = {||} \ (\e. f = {|e|}))" apply standard apply (metis bot.not_eq_extremum gr_implies_not0 le_neq_implies_less less_one size_fsingleton size_fsubset) by auto lemma size_gt_1: "1 < size f \ \e1 e2 f'. e1 \ e2 \ f = finsert e1 (finsert e2 f')" apply (induct f) apply simp apply (rule_tac x=x in exI) by (metis finsertCI leD not_le_imp_less size_le_1) end diff --git a/thys/Imperative_Insertion_Sort/Imperative_Insertion_Sort.thy b/thys/Imperative_Insertion_Sort/Imperative_Insertion_Sort.thy --- a/thys/Imperative_Insertion_Sort/Imperative_Insertion_Sort.thy +++ b/thys/Imperative_Insertion_Sort/Imperative_Insertion_Sort.thy @@ -1,473 +1,473 @@ (* Title: Imperative_Insertion_Sort Author: Christian Sternagel *) section \Insertion Sort\ theory Imperative_Insertion_Sort imports Imperative_Loops "HOL-Library.Multiset" begin subsection \The Algorithm\ abbreviation array_update :: "'a::heap array \ nat \ 'a \ 'a array Heap" ("(_.'(_') \/ _)" [1000, 0, 13] 14) where "a.(i) \ x \ Array.upd i x a" abbreviation array_nth :: "'a::heap array \ nat \ 'a Heap" ("_.'(_')" [1000, 0] 14) where "a.(i) \ Array.nth a i" text \ A definition of insertion sort as given by Cormen et al.\ in \emph{Introduction to Algorithms}. Compared to the informal textbook version the variant below is a bit unwieldy due to explicit dereferencing of variables on the heap. To avoid ambiguities with existing syntax we use OCaml's notation for accessing (@{term "a.(i)"}) and updating (@{term "a.(i) \ x"}) an array @{term a} at position @{term i}. \ definition "insertion_sort a = do { l \ Array.len a; for [1 ..< l] (\j. do { \ \Insert \a[j]\ into the sorted subarray \a[1 .. j - 1].\\ key \ a.(j); i \ ref j; while (do { i' \ ! i; if i' > 0 then do {x \ a.(i' - 1); return (x > key)} else return False}) (do { i' \ ! i; x \ a.(i' - 1); a.(i') \ x; i := i' - 1 }); i' \ ! i; a.(i') \ key }) }" text \ The following definitions decompose the nested loops of the algorithm into more manageable chunks. \ definition "shiftr_p a (key::'a::{heap, linorder}) i = (do {i' \ ! i; if i' > 0 then do {x \ a.(i' - 1); return (x > key)} else return False})" definition "shiftr_f a i = do { i' \ ! i; x \ a.(i' - 1); a.(i') \ x; i := i' - 1 }" definition "shiftr a key i = while (shiftr_p a key i) (shiftr_f a i)" definition "insert_elt a = (\j. do { key \ a.(j); i \ ref j; shiftr a key i; i' \ ! i; a.(i') \ key })" definition "sort_upto a = (\l. for [1 ..< l] (insert_elt a))" lemma insertion_sort_alt_def: "insertion_sort a = (Array.len a \ sort_upto a)" by (simp add: insertion_sort_def sort_upto_def shiftr_def shiftr_p_def shiftr_f_def insert_elt_def) subsection \Partial Correctness\ lemma effect_shiftr_f: assumes "effect (shiftr_f a i) h h' u" shows "Ref.get h' i = Ref.get h i - 1 \ Array.get h' a = list_update (Array.get h a) (Ref.get h i) (Array.get h a ! (Ref.get h i - 1))" using assms by (auto simp: shiftr_f_def elim!: effect_elims) lemma success_shiftr_p: "Ref.get h i < Array.length h a \ success (shiftr_p a key i) h" by (auto simp: success_def shiftr_p_def execute_simps) interpretation ro_shiftr_p: ro_cond "shiftr_p a key i" for a key i by (unfold_locales) (auto simp: shiftr_p_def success_def execute_simps execute_bind_case split: option.split, metis effectI effect_nthE) (*[0 .. j - 1]*) definition [simp]: "ini h a j = take j (Array.get h a)" (*[0 .. i - 1]*) definition [simp]: "left h a i = take (Ref.get h i) (Array.get h a)" (*[i+1 .. j]*) definition [simp]: "right h a j i = take (j - Ref.get h i) (drop (Ref.get h i + 1) (Array.get h a))" (*[0 .. i - 1, i + 1 .. j]*) definition [simp]: "both h a j i = left h a i @ right h a j i" lemma effect_shiftr: assumes "Ref.get h i = j" (is "?i h = _") and "j < Array.length h a" and "sorted (take j (Array.get h a))" and "effect (while (shiftr_p a key i) (shiftr_f a i)) h h' u" shows "Array.length h a = Array.length h' a \ ?i h' \ j \ mset (list_update (Array.get h a) j key) = mset (list_update (Array.get h' a) (?i h') key) \ ini h a j = both h' a j i \ sorted (both h' a j i) \ (\x \ set (right h' a j i). x > key)" using assms(4, 2) proof (induction rule: ro_shiftr_p.effect_while_induct) case base show ?case using assms by auto next case (step h' h'' u) from \success (shiftr_p a key i) h'\ and \cond (shiftr_p a key i) h'\ have "?i h' > 0" and key: "Array.get h' a ! (?i h' - 1) > key" by (auto dest!: ro_shiftr_p.success_cond_effect) (auto simp: shiftr_p_def elim!: effect_elims effect_ifE) from effect_shiftr_f [OF \effect (shiftr_f a i) h' h'' u\] have [simp]: "?i h'' = ?i h' - 1" "Array.get h'' a = list_update (Array.get h' a) (?i h') (Array.get h' a ! (?i h' - 1))" by auto from step have *: "?i h' < length (Array.get h' a)" and **: "?i h' - (Suc 0) \ ?i h'" "?i h' \ length (Array.get h' a)" and "?i h' \ j" and "?i h' < Suc j" and IH: "ini h a j = both h' a j i" by (auto simp add: Array.length_def) have "Array.length h a = Array.length h'' a" using step by (simp add: Array.length_def) moreover have "?i h'' \ j" using step by auto moreover have "mset (list_update (Array.get h a) j key) = mset (list_update (Array.get h'' a) (?i h'') key)" proof - have "?i h' < length (Array.get h' a)" and "?i h' - 1 < length (Array.get h' a)" using * by auto then show ?thesis using step by (simp add: mset_update ac_simps nth_list_update) qed moreover have "ini h a j = both h'' a j i" using \0 < ?i h'\ and \?i h' \ j\ and \?i h' < length (Array.get h' a)\ and ** and IH by (auto simp: upd_conv_take_nth_drop Suc_diff_le min_absorb1) (metis Suc_lessD Suc_pred take_Suc_conv_app_nth) moreover have "sorted (both h'' a j i)" using step and \0 < ?i h'\ and \?i h' \ j\ and \?i h' < length (Array.get h' a)\ and ** by (auto simp: IH upd_conv_take_nth_drop Suc_diff_le min_absorb1) (metis Suc_lessD Suc_pred append.simps append_assoc take_Suc_conv_app_nth) moreover have "\x \ set (right h'' a j i). x > key" using step and \0 < ?i h'\ and \?i h' < length (Array.get h' a)\ and key by (auto simp: upd_conv_take_nth_drop Suc_diff_le) ultimately show ?case by blast qed lemma sorted_take_nth: assumes "0 < i" and "i < length xs" and "xs ! (i - 1) \ y" and "sorted (take i xs)" shows "\x \ set (take i xs). x \ y" proof - have "take i xs = take (i - 1) xs @ [xs ! (i - 1)]" using \0 < i\ and \i < length xs\ by (metis Suc_diff_1 less_imp_diff_less take_Suc_conv_app_nth) then show ?thesis using \sorted (take i xs)\ and \xs ! (i - 1) \y\ by (auto simp: sorted_append) qed lemma effect_for_insert_elt: assumes "l \ Array.length h a" and "1 \ l" and "effect (for [1 ..< l] (insert_elt a)) h h' u" shows "Array.length h a = Array.length h' a \ sorted (take l (Array.get h' a)) \ mset (Array.get h a) = mset (Array.get h' a)" using assms(2-) proof (induction l h' rule: effect_for_induct) case base show ?case by (cases "Array.get h a") simp_all next case (step j h' h'' u) with assms(1) have "j < Array.length h' a" by auto from step have sorted: "sorted (take j (Array.get h' a))" by blast from step(3) [unfolded insert_elt_def] obtain key and h\<^sub>1 and i and h\<^sub>2 and i' where key: "key = Array.get h' a ! j" and "effect (ref j) h' h\<^sub>1 i" and ref\<^sub>1: "Ref.get h\<^sub>1 i = j" and shiftr': "effect (shiftr a key i) h\<^sub>1 h\<^sub>2 ()" and [simp]: "Ref.get h\<^sub>2 i = i'" and [simp]: "h'' = Array.update a i' key h\<^sub>2" and "i' < Array.length h\<^sub>2 a" by (elim effect_bindE effect_nthE effect_lookupE effect_updE) (auto intro: effect_intros, metis effect_refE) from \effect (ref j) h' h\<^sub>1 i\ have [simp]: "Array.get h\<^sub>1 a = Array.get h' a" by (metis array_get_alloc effectE execute_ref option.sel) have [simp]: "Array.length h\<^sub>1 a = Array.length h' a" by (simp add: Array.length_def) from step and assms(1) have "j < Array.length h\<^sub>1 a" "sorted (take j (Array.get h\<^sub>1 a))" by auto note shiftr = effect_shiftr [OF ref\<^sub>1 this shiftr' [unfolded shiftr_def], simplified] have "i' \ j" using shiftr by simp have "i' < length (Array.get h\<^sub>2 a)" by (metis \i' < Array.length h\<^sub>2 a\ length_def) have [simp]: "min (Suc j) i' = i'" using \i' \ j\ by simp have [simp]: "min (length (Array.get h\<^sub>2 a)) i' = i'" using \i' < length (Array.get h\<^sub>2 a)\ by (simp) have take_Suc_j: "take (Suc j) (list_update (Array.get h\<^sub>2 a) i' key) = take i' (Array.get h\<^sub>2 a) @ key # take (j - i') (drop (Suc i') (Array.get h\<^sub>2 a))" unfolding upd_conv_take_nth_drop [OF \i' < length (Array.get h\<^sub>2 a)\] by (auto) (metis Suc_diff_le \i' \ j\ take_Suc_Cons) have "Array.length h a = Array.length h'' a" using shiftr by (auto) (metis step.IH) moreover have "mset (Array.get h a) = mset (Array.get h'' a)" using shiftr and step by (simp add: key) moreover have "sorted (take (Suc j) (Array.get h'' a))" proof - from ro_shiftr_p.effect_while_post [OF shiftr' [unfolded shiftr_def]] have "i' = 0 \ (0 < i' \ key \ Array.get h\<^sub>2 a ! (i' - 1))" by (auto dest!: ro_shiftr_p.success_not_cond_effect) (auto elim!: effect_elims simp: shiftr_p_def) then show ?thesis proof assume [simp]: "i' = 0" have *: "take (Suc j) (list_update (Array.get h\<^sub>2 a) 0 key) = key # take j (drop 1 (Array.get h\<^sub>2 a))" by (simp) (metis \i' = 0\ append_Nil take_Suc_j diff_zero take_0) from sorted and shiftr have "sorted (take j (drop 1 (Array.get h\<^sub>2 a)))" and "\x \ set (take j (drop 1 (Array.get h\<^sub>2 a))). key < x" by simp_all then have "sorted (key # take j (drop 1 (Array.get h\<^sub>2 a)))" - by (metis less_imp_le sorted.simps(2)) + by (metis less_imp_le sorted_simps(2)) then show ?thesis by (simp add: *) next assume "0 < i' \ key \ Array.get h\<^sub>2 a ! (i' - 1)" moreover have "sorted (take i' (Array.get h\<^sub>2 a) @ take (j - i') (drop (Suc i') (Array.get h\<^sub>2 a)))" and "\x \ set (take (j - i') (drop (Suc i') (Array.get h\<^sub>2 a))). key < x" using shiftr by auto ultimately have "\x \ set (take i' (Array.get h\<^sub>2 a)). x \ key" using sorted_take_nth [OF _ \i' < length (Array.get h\<^sub>2 a)\, of key] by (simp add: sorted_append) then show ?thesis using shiftr by (auto simp: take_Suc_j sorted_append less_imp_le) qed qed ultimately show ?case by blast qed lemma effect_insertion_sort: assumes "effect (insertion_sort a) h h' u" shows "mset (Array.get h a) = mset (Array.get h' a) \ sorted (Array.get h' a)" using assms apply (cases "Array.length h a") apply (auto elim!: effect_elims simp: insertion_sort_def Array.length_def)[1] unfolding insertion_sort_def unfolding shiftr_p_def [symmetric] shiftr_f_def [symmetric] unfolding shiftr_def [symmetric] insert_elt_def [symmetric] apply (elim effect_elims) apply (simp only:) apply (subgoal_tac "Suc nat \ Array.length h a") apply (drule effect_for_insert_elt) apply (auto simp: Array.length_def) done subsection \Total Correctness\ lemma success_shiftr_f: assumes "Ref.get h i < Array.length h a" shows "success (shiftr_f a i) h" using assms by (auto simp: success_def shiftr_f_def execute_simps) lemma success_shiftr: assumes "Ref.get h i < Array.length h a" shows "success (while (shiftr_p a key i) (shiftr_f a i)) h" proof - have "wf (measure (\h. Ref.get h i))" by (metis wf_measure) then show ?thesis proof (induct taking: "\h. Ref.get h i < Array.length h a" rule: ro_shiftr_p.success_while_induct) case (success_cond h) then show ?case by (metis success_shiftr_p) next case (success_body h) then show ?case by (blast intro: success_shiftr_f) next case (step h h' r) then show ?case by (auto dest!: effect_shiftr_f ro_shiftr_p.success_cond_effect simp: length_def) (auto simp: shiftr_p_def elim!: effect_elims effect_ifE) qed fact qed lemma effect_shiftr_index: assumes "effect (shiftr a key i) h h' u" shows "Ref.get h' i \ Ref.get h i" using assms unfolding shiftr_def by (induct h' rule: ro_shiftr_p.effect_while_induct) (auto dest: effect_shiftr_f) lemma effect_shiftr_length: assumes "effect (shiftr a key i) h h' u" shows "Array.length h' a = Array.length h a" using assms unfolding shiftr_def by (induct h' rule: ro_shiftr_p.effect_while_induct) (auto simp: length_def dest: effect_shiftr_f) lemma success_insert_elt: assumes "k < Array.length h a" shows "success (insert_elt a k) h" proof - obtain key where "effect (a.(k)) h h key" using assms by (auto intro: effect_intros) moreover obtain i and h\<^sub>1 where "effect (ref k) h h\<^sub>1 i" and [simp]: "Ref.get h\<^sub>1 i = k" and [simp]: "Array.length h\<^sub>1 a = Array.length h a" by (auto simp: ref_def length_def) (metis Ref.get_alloc array_get_alloc effect_heapI) moreover obtain h\<^sub>2 where *: "effect (shiftr a key i) h\<^sub>1 h\<^sub>2 ()" using success_shiftr [of h\<^sub>1 i a key] and assms by (auto simp: success_def effect_def shiftr_def) moreover have "effect (! i) h\<^sub>2 h\<^sub>2 (Ref.get h\<^sub>2 i)" and "Ref.get h\<^sub>2 i \ Ref.get h\<^sub>1 i" and "Ref.get h\<^sub>2 i < Array.length h\<^sub>2 a" using effect_shiftr_index [OF *] and effect_shiftr_length [OF *] and assms by (auto intro!: effect_intros) moreover then obtain h\<^sub>3 and r where "effect (a.(Ref.get h\<^sub>2 i) \ key) h\<^sub>2 h\<^sub>3 r" using assms by (auto simp: effect_def execute_simps) ultimately have "effect (insert_elt a k) h h\<^sub>3 r" by (auto simp: insert_elt_def intro: effect_intros) then show ?thesis by (metis effectE) qed lemma for_insert_elt_correct: assumes "l \ Array.length h a" and "1 \ l" shows "\h'. effect (for [1 ..< l] (insert_elt a)) h h' () \ Array.length h a = Array.length h' a \ sorted (take l (Array.get h' a)) \ mset (Array.get h a) = mset (Array.get h' a)" using assms(2) proof (induction rule: for_induct) case (succeed k h) then show ?case using assms and success_insert_elt [of k h a] by auto next case base show ?case by (cases "Array.get h a") simp_all next case (step j h' h'' u) with assms(1) have "j < Array.length h' a" by auto from step have sorted: "sorted (take j (Array.get h' a))" by blast from step(4) [unfolded insert_elt_def] obtain key and h\<^sub>1 and i and h\<^sub>2 and i' where key: "key = Array.get h' a ! j" and "effect (ref j) h' h\<^sub>1 i" and ref\<^sub>1: "Ref.get h\<^sub>1 i = j" and shiftr': "effect (shiftr a key i) h\<^sub>1 h\<^sub>2 ()" and [simp]: "Ref.get h\<^sub>2 i = i'" and [simp]: "h'' = Array.update a i' key h\<^sub>2" and "i' < Array.length h\<^sub>2 a" by (elim effect_bindE effect_nthE effect_lookupE effect_updE) (auto intro: effect_intros, metis effect_refE) from \effect (ref j) h' h\<^sub>1 i\ have [simp]: "Array.get h\<^sub>1 a = Array.get h' a" by (metis array_get_alloc effectE execute_ref option.sel) have [simp]: "Array.length h\<^sub>1 a = Array.length h' a" by (simp add: Array.length_def) from step and assms(1) have "j < Array.length h\<^sub>1 a" "sorted (take j (Array.get h\<^sub>1 a))" by auto note shiftr = effect_shiftr [OF ref\<^sub>1 this shiftr' [unfolded shiftr_def], simplified] have "i' \ j" using shiftr by simp have "i' < length (Array.get h\<^sub>2 a)" by (metis \i' < Array.length h\<^sub>2 a\ length_def) have [simp]: "min (Suc j) i' = i'" using \i' \ j\ by simp have [simp]: "min (length (Array.get h\<^sub>2 a)) i' = i'" using \i' < length (Array.get h\<^sub>2 a)\ by (simp) have take_Suc_j: "take (Suc j) (list_update (Array.get h\<^sub>2 a) i' key) = take i' (Array.get h\<^sub>2 a) @ key # take (j - i') (drop (Suc i') (Array.get h\<^sub>2 a))" unfolding upd_conv_take_nth_drop [OF \i' < length (Array.get h\<^sub>2 a)\] by (auto) (metis Suc_diff_le \i' \ j\ take_Suc_Cons) have "Array.length h a = Array.length h'' a" using shiftr by (auto) (metis step.hyps(1)) moreover have "mset (Array.get h a) = mset (Array.get h'' a)" using shiftr and step by (simp add: key) moreover have "sorted (take (Suc j) (Array.get h'' a))" proof - from ro_shiftr_p.effect_while_post [OF shiftr' [unfolded shiftr_def]] have "i' = 0 \ (0 < i' \ key \ Array.get h\<^sub>2 a ! (i' - 1))" by (auto dest!: ro_shiftr_p.success_not_cond_effect) (auto elim!: effect_elims simp: shiftr_p_def) then show ?thesis proof assume [simp]: "i' = 0" have *: "take (Suc j) (list_update (Array.get h\<^sub>2 a) 0 key) = key # take j (drop 1 (Array.get h\<^sub>2 a))" by (simp) (metis \i' = 0\ append_Nil take_Suc_j diff_zero take_0) from sorted and shiftr have "sorted (take j (drop 1 (Array.get h\<^sub>2 a)))" and "\x \ set (take j (drop 1 (Array.get h\<^sub>2 a))). key < x" by simp_all then have "sorted (key # take j (drop 1 (Array.get h\<^sub>2 a)))" - by (metis less_imp_le sorted.simps(2)) + by (metis less_imp_le sorted_simps(2)) then show ?thesis by (simp add: *) next assume "0 < i' \ key \ Array.get h\<^sub>2 a ! (i' - 1)" moreover have "sorted (take i' (Array.get h\<^sub>2 a) @ take (j - i') (drop (Suc i') (Array.get h\<^sub>2 a)))" and "\x \ set (take (j - i') (drop (Suc i') (Array.get h\<^sub>2 a))). key < x" using shiftr by auto ultimately have "\x \ set (take i' (Array.get h\<^sub>2 a)). x \ key" using sorted_take_nth [OF _ \i' < length (Array.get h\<^sub>2 a)\, of key] by (simp add: sorted_append) then show ?thesis using shiftr by (auto simp: take_Suc_j sorted_append less_imp_le) qed qed ultimately show ?case by blast qed lemma insertion_sort_correct: "\h'. effect (insertion_sort a) h h' u \ mset (Array.get h a) = mset (Array.get h' a) \ sorted (Array.get h' a)" proof (cases "Array.length h a = 0") assume "Array.length h a = 0" then have "effect (insertion_sort a) h h ()" and "mset (Array.get h a) = mset (Array.get h a)" and "sorted (Array.get h a)" by (auto simp: insertion_sort_def length_def intro!: effect_intros) then show ?thesis by auto next assume "Array.length h a \ 0" then have "1 \ Array.length h a" by auto from for_insert_elt_correct [OF le_refl this] show ?thesis by (auto simp: insertion_sort_alt_def sort_upto_def) (metis One_nat_def effect_bindI effect_insertion_sort effect_lengthI insertion_sort_alt_def sort_upto_def) qed export_code insertion_sort in Haskell end diff --git a/thys/Iptables_Semantics/Common/Word_Upto.thy b/thys/Iptables_Semantics/Common/Word_Upto.thy --- a/thys/Iptables_Semantics/Common/Word_Upto.thy +++ b/thys/Iptables_Semantics/Common/Word_Upto.thy @@ -1,206 +1,206 @@ section\Word Upto\ theory Word_Upto imports Main IP_Addresses.Hs_Compat Word_Lib.Word_Lemmas begin text\Enumerate a range of machine words.\ text\enumerate from the back (inefficient)\ function word_upto :: "'a word \ 'a word \ ('a::len) word list" where "word_upto a b = (if a = b then [a] else word_upto a (b - 1) @ [b])" by pat_completeness auto (*by the way: does not terminate practically if b < a; will terminate after it reaches the word wrap-around!*) termination word_upto apply(relation "measure (unat \ uncurry (-) \ prod.swap)") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply(simp add: diff_right_commute; fail) apply(rule measure_unat) apply auto done declare word_upto.simps[simp del] text\enumerate from the front (more inefficient)\ function word_upto' :: "'a word \ 'a word \ ('a::len) word list" where "word_upto' a b = (if a = b then [a] else a # word_upto' (a + 1) b)" by pat_completeness auto termination word_upto' apply(relation "measure (\ (a, b). unat (b - a))") apply(rule wf_measure; fail) apply(simp) apply(subgoal_tac "unat (b - a - 1) < unat (b - a)") apply (simp add: diff_diff_add; fail) apply(rule measure_unat) apply auto done declare word_upto'.simps[simp del] lemma word_upto_cons_front[code]: "word_upto a b = word_upto' a b" proof(induction a b rule:word_upto'.induct) case (1 a b) have hlp1: "a \ b \ a # word_upto (a + 1) b = word_upto a b" apply(induction a b rule:word_upto.induct) apply simp apply(subst(1) word_upto.simps) apply(simp) apply safe apply(subst(1) word_upto.simps) apply (simp) apply(subst(1) word_upto.simps) apply (simp; fail) apply(case_tac "a \ b - 1") apply(simp) apply (metis Cons_eq_appendI word_upto.simps) apply(simp) done from 1[symmetric] show ?case apply(cases "a = b") subgoal apply(subst word_upto.simps) apply(subst word_upto'.simps) by(simp) apply(subst word_upto'.simps) by(simp add: hlp1) qed (* Most of the lemmas I show about word_upto hold without a \ b, but I don't need that right now and it's giving me a headache *) lemma word_upto_set_eq: "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" proof show "a \ b \ x \ set (word_upto a b) \ a \ x \ x \ b" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst(asm) word_upto.simps) apply(simp; fail) apply(subst(asm) word_upto.simps) apply(simp) apply(erule disjE) apply(simp; fail) proof(goal_cases) case (1 a b) from 1(2-3) have "b \ 0" by force from 1(2,3) have "a \ b - 1" by (simp add: word_le_minus_one_leq) from 1(1)[OF this 1(4)] show ?case by (metis dual_order.trans 1(2,3) less_imp_le measure_unat word_le_0_iff word_le_nat_alt) qed next show "a \ x \ x \ b \ x \ set (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(simp) apply(case_tac "x = b") apply(simp;fail) proof(goal_cases) case (1 a b) from 1(2-4) have "b \ 0" by force from 1(2,4) have "x \ b - 1" using le_step_down_word by auto from 1(1) this show ?case by simp qed qed lemma word_upto_distinct_hlp: "a \ b \ a \ b \ b \ set (word_upto a (b - 1))" apply(rule ccontr, unfold not_not) apply(subgoal_tac "a \ b - 1") apply(drule iffD1[OF word_upto_set_eq[of a "b -1" b]]) apply(simp add: word_upto.simps) apply(subgoal_tac "b \ 0") apply(meson leD measure_unat word_le_nat_alt) apply(blast intro: iffD1[OF word_le_0_iff]) using le_step_down_word apply blast done lemma distinct_word_upto: "a \ b \ distinct (word_upto a b)" apply(induction a b rule: word_upto.induct) apply(case_tac "a = b") apply(subst word_upto.simps) apply(simp; force) apply(subst word_upto.simps) apply(case_tac "a \ b - 1") apply(simp) apply(rule word_upto_distinct_hlp; simp) apply(simp) apply(rule ccontr) apply (simp add: not_le antisym word_minus_one_le_leq) done lemma word_upto_eq_upto: "s \ e \ e \ unat (max_word :: 'l word) \ word_upto ((of_nat :: nat \ ('l :: len) word) s) (of_nat e) = map of_nat (upt s (Suc e))" proof(induction e) let ?mwon = "of_nat :: nat \ 'l word" let ?mmw = "max_word :: 'l word" case (Suc e) show ?case proof(cases "?mwon s = ?mwon (Suc e)") case True have "s = Suc e" using le_unat_uoi Suc.prems True by metis with True show ?thesis by(subst word_upto.simps) (simp) next case False hence le: "s \ e" using le_SucE Suc.prems by blast have lm: "e \ unat ?mmw" using Suc.prems by simp have sucm: "(of_nat :: nat \ ('l :: len) word) (Suc e) - 1 = of_nat e" using Suc.prems(2) by simp note mIH = Suc.IH[OF le lm] show ?thesis by(subst word_upto.simps) (simp add: False[simplified] Suc.prems mIH sucm) qed qed(simp add: word_upto.simps) lemma word_upto_alt: "(a :: ('l :: len) word) \ b \ word_upto a b = map of_nat (upt (unat a) (Suc (unat b)))" proof - let ?mmw = "max_word :: 'l word" assume le: "a \ b" hence nle: "unat a \ unat b" by(unat_arith) have lem: "unat b \ unat ?mmw" by (simp add: word_unat_less_le) note word_upto_eq_upto[OF nle lem, unfolded word_unat.Rep_inverse] thus "word_upto a b = map of_nat [unat a.. b then map of_nat (upt (unat a) (Suc (unat b))) else word_upto a b)" using word_upto_alt by metis lemma sorted_word_upto: fixes a b :: "('l :: len) word" assumes "a \ b" shows "sorted (word_upto a b)" proof - define m and n where \m = unat a\ and \n = Suc (unat b)\ moreover have \sorted (map of_nat [m.. apply (simp add: sorted_map) apply (rule sorted_wrt_mono_rel [of _ \(\)\]) - apply (simp_all flip: sorted_sorted_wrt) + apply simp_all apply (simp add: le_unat_uoi less_Suc_eq_le n_def word_of_nat_le) apply transfer apply simp apply (subst take_bit_int_eq_self) apply (simp_all add: le_less_trans) apply (metis le_unat_uoi of_int_of_nat_eq of_nat_mono uint_word_of_int_eq unat_eq_nat_uint unsigned_of_int) done ultimately have \sorted (map of_nat [unat a.. by simp with assms show ?thesis by (simp only: word_upto_alt) qed end diff --git a/thys/KBPs/ODList.thy b/thys/KBPs/ODList.thy --- a/thys/KBPs/ODList.thy +++ b/thys/KBPs/ODList.thy @@ -1,657 +1,653 @@ (*<*) (* * Knowledge-based programs. * (C)opyright 2011, Peter Gammie, peteg42 at gmail.com. * License: BSD * * Based on Florian Haftmann's DList.thy and Tobias Nipkow's msort proofs. *) theory ODList imports "HOL-Library.Multiset" List_local begin (*>*) text\ Define a type of ordered distinct lists, intended to represent sets. The advantage of this representation is that it is isomorphic to the set of finite sets. Conversely it requires the carrier type to be a linear order. Note that this representation does not arise from a quotient on lists: all the unsorted lists are junk. \ context linorder begin text\ "Absorbing" msort, a variant of Tobias Nipkow's proofs from 1992. \ fun merge :: "'a list \ 'a list \ 'a list" where "merge [] ys = ys" | "merge xs [] = xs" | "merge (x#xs) (y#ys) = (if x = y then merge xs (y#ys) else if x < y then x # merge xs (y#ys) else y # merge (x#xs) ys)" (*<*) lemma set_merge[simp]: "set (merge xs ys) = set (xs @ ys)" by (induct xs ys rule: merge.induct) auto lemma distinct_sorted_merge[simp]: "\ distinct xs; distinct ys; sorted xs; sorted ys \ \ distinct (merge xs ys) \ sorted (merge xs ys)" by (induct xs ys rule: merge.induct) (auto) lemma mset_merge [simp]: "\ distinct (xs @ ys) \ \ mset (merge xs ys) = mset xs + mset ys" by (induct xs ys rule: merge.induct) (simp_all add: ac_simps) (*>*) text\The "absorbing" sort itself.\ fun msort :: "'a list \ 'a list" where "msort [] = []" | "msort [x] = [x]" | "msort xs = merge (msort (take (size xs div 2) xs)) (msort (drop (size xs div 2) xs))" (*<*) lemma msort_distinct_sorted[simp]: "distinct (msort xs) \ sorted (msort xs)" by (induct xs rule: msort.induct) simp_all lemma msort_set[simp]: "set (msort xs) = set xs" by (induct xs rule: msort.induct) (simp_all, metis List.set_simps(2) append_take_drop_id set_append) (* thankyou sledgehammer! *) lemma msort_remdups[simp]: "remdups (msort xs) = msort xs" by simp lemma msort_idle[simp]: "\ distinct xs; sorted xs \ \ msort xs = xs" by (rule map_sorted_distinct_set_unique[where f=id]) (auto simp: map.id) lemma mset_msort[simp]: "distinct xs \ mset (msort xs) = mset xs" by (rule iffD1[OF set_eq_iff_mset_eq_distinct]) simp_all lemma msort_sort[simp]: "distinct xs \ sort xs = msort xs" by (simp add: properties_for_sort) (*>*) end (* context linorder *) section \The @{term "odlist"} type\ typedef (overloaded) ('a :: linorder) odlist = "{ x::'a list . sorted x \ distinct x }" - morphisms toList odlist_Abs by (auto iff: sorted.simps(1)) + morphisms toList odlist_Abs by (auto iff: sorted_simps(1)) lemma distinct_toList[simp]: "distinct (toList xs)" using toList by auto lemma sorted_toList[simp]: "sorted (toList xs)" using toList by auto text\ Code generator voodoo: this is the constructor for the abstract type. \ definition ODList :: "('a :: linorder) list \ 'a odlist" where "ODList \ odlist_Abs \ msort" lemma toList_ODList: "toList (ODList xs) = msort xs" unfolding ODList_def by (simp add: odlist_Abs_inverse) lemma ODList_toList[simp, code abstype]: "ODList (toList xs) = xs" unfolding ODList_def by (cases xs) (simp add: odlist_Abs_inverse) text\ Runtime cast from @{typ "'a list"} into @{typ "'a odlist"}. This is just a renaming of @{term "ODList"} -- names are significant to the code generator's abstract type machinery. \ definition fromList :: "('a :: linorder) list \ 'a odlist" where "fromList \ ODList" lemma toList_fromList[code abstract]: "toList (fromList xs) = msort xs" unfolding fromList_def by (simp add: toList_ODList) subsection\Basic properties: equality, finiteness\ (*<*) declare toList_inject[iff] (*>*) instantiation odlist :: (linorder) equal (*<*) begin definition [code]: "HOL.equal A B \ odlist_equal (toList A) (toList B)" instance by standard (simp add: equal_odlist_def) end (*>*) instance odlist :: ("{finite, linorder}") finite (*<*) proof let ?ol = "UNIV :: 'a odlist set" let ?s = "UNIV :: 'a set set" have "finite ?s" by simp moreover have "?ol \ range (odlist_Abs \ sorted_list_of_set)" proof fix x show "x \ range (odlist_Abs \ sorted_list_of_set)" apply (cases x) apply (rule range_eqI[where x="set (toList x)"]) apply (clarsimp simp: odlist_Abs_inject sorted_list_of_set_sort_remdups odlist_Abs_inverse distinct_remdups_id) done qed ultimately show "finite ?ol" by (blast intro: finite_surj) qed (*>*) subsection\Constants\ definition empty :: "('a :: linorder) odlist" where "empty \ ODList []" lemma toList_empty[simp, code abstract]: "toList empty = []" unfolding empty_def by (simp add: toList_ODList) subsection\Operations\ subsubsection\toSet\ definition toSet :: "('a :: linorder) odlist \ 'a set" where "toSet X = set (toList X)" lemma toSet_empty[simp]: "toSet empty = {}" unfolding toSet_def empty_def by (simp add: toList_ODList) lemma toSet_ODList[simp]: "\ distinct xs; sorted xs \ \ toSet (ODList xs) = set xs" unfolding toSet_def by (simp add: toList_ODList) lemma toSet_fromList_set[simp]: "toSet (fromList xs) = set xs" unfolding toSet_def fromList_def by (simp add: toList_ODList) lemma toSet_inj[intro, simp]: "inj toSet" apply (rule injI) unfolding toSet_def apply (case_tac x) apply (case_tac y) apply (auto iff: odlist_Abs_inject odlist_Abs_inverse sorted_distinct_set_unique) done lemma toSet_eq_iff: "X = Y \ toSet X = toSet Y" by (blast dest: injD[OF toSet_inj]) subsubsection\head\ definition hd :: "('a :: linorder) odlist \ 'a" where [code]: "hd \ List.hd \ toList" lemma hd_toList: "toList xs = y # ys \ ODList.hd xs = y" unfolding hd_def by simp subsubsection\member\ definition member :: "('a :: linorder) odlist \ 'a \ bool" where [code]: "member xs x \ List.member (toList xs) x" lemma member_toSet[iff]: "member xs x \x \ toSet xs" unfolding member_def toSet_def by (simp add: in_set_member) subsubsection\Filter\ definition filter :: "(('a :: linorder) \ bool) \ 'a odlist \ 'a odlist" where "filter P xs \ ODList (List.filter P (toList xs))" lemma toList_filter[simp, code abstract]: "toList (filter P xs) = List.filter P (toList xs)" unfolding filter_def by (simp add: toList_ODList) lemma toSet_filter[simp]: "toSet (filter P xs) = { x \ toSet xs . P x }" unfolding filter_def apply simp apply (simp add: toSet_def) done subsubsection\All\ definition odlist_all :: "('a :: linorder \ bool) \ 'a odlist \ bool" where [code]: "odlist_all P xs \ list_all P (toList xs)" lemma odlist_all_iff: "odlist_all P xs \ (\x \ toSet xs. P x)" unfolding odlist_all_def toSet_def by (simp only: list_all_iff) lemma odlist_all_cong [fundef_cong]: "xs = ys \ (\x. x \ toSet ys \ f x = g x) \ odlist_all f xs = odlist_all g ys" by (simp add: odlist_all_iff) subsubsection\Difference\ definition difference :: "('a :: linorder) odlist \ 'a odlist \ 'a odlist" where "difference xs ys = ODList (List_local.difference (toList xs) (toList ys))" lemma toList_difference[simp, code abstract]: "toList (difference xs ys) = List_local.difference (toList xs) (toList ys)" unfolding difference_def by (simp add: toList_ODList) lemma toSet_difference[simp]: "toSet (difference xs ys) = toSet xs - toSet ys" unfolding difference_def apply simp apply (simp add: toSet_def) done subsubsection\Intersection\ definition intersect :: "('a :: linorder) odlist \ 'a odlist \ 'a odlist" where "intersect xs ys = ODList (List_local.intersection (toList xs) (toList ys))" lemma toList_intersect[simp, code abstract]: "toList (intersect xs ys) = List_local.intersection (toList xs) (toList ys)" unfolding intersect_def by (simp add: toList_ODList) lemma toSet_intersect[simp]: "toSet (intersect xs ys) = toSet xs \ toSet ys" unfolding intersect_def apply simp apply (simp add: toSet_def) done subsubsection\Union\ definition union :: "('a :: linorder) odlist \ 'a odlist \ 'a odlist" where "union xs ys = ODList (merge (toList xs) (toList ys))" lemma toList_union[simp, code abstract]: "toList (union xs ys) = merge (toList xs) (toList ys)" unfolding union_def by (simp add: toList_ODList) lemma toSet_union[simp]: "toSet (union xs ys) = toSet xs \ toSet ys" unfolding union_def apply simp apply (simp add: toSet_def) done definition big_union :: "('b \ ('a :: linorder) odlist) \ 'b list \ 'a odlist" where [code]: "big_union f X \ foldr (\a A. ODList.union (f a) A) X ODList.empty" lemma toSet_big_union[simp]: "toSet (big_union f X) = (\x \ set X. toSet (f x))" proof - { fix X Y have "toSet (foldr (\x A. ODList.union (f x) A) X Y) = toSet Y \ (\x \ set X. toSet (f x))" by (induct X arbitrary: Y) auto } thus ?thesis unfolding big_union_def by simp qed subsubsection\Case distinctions\ text\ We construct ODLists out of lists, so talk in terms of those, not a one-step constructor we don't use. \ lemma distinct_sorted_induct [consumes 2, case_names Nil insert]: assumes "distinct xs" assumes "sorted xs" assumes base: "P []" assumes step: "\x xs. \ distinct (x # xs); sorted (x # xs); P xs \ \ P (x # xs)" shows "P xs" using \distinct xs\ \sorted xs\ proof (induct xs) case Nil from \P []\ show ?case . next case (Cons x xs) then have "distinct (x # xs)" and "sorted (x # xs)" and "P xs" by (simp_all) with step show "P (x # xs)" . qed lemma odlist_induct [case_names empty insert, cases type: odlist]: assumes empty: "\dxs. dxs = empty \ P dxs" assumes insrt: "\dxs x xs. \ dxs = fromList (x # xs); distinct (x # xs); sorted (x # xs); P (fromList xs) \ \ P dxs" shows "P dxs" proof (cases dxs) case (odlist_Abs xs) then have dxs: "dxs = ODList xs" and distinct: "distinct xs" and sorted: "sorted xs" by (simp_all add: ODList_def) from \distinct xs\ and \sorted xs\ have "P (ODList xs)" proof (induct xs rule: distinct_sorted_induct) case Nil from empty show ?case by (simp add: empty_def) next case (insert x xs) thus ?case apply - apply (rule insrt) apply (auto simp: fromList_def) done qed with dxs show "P dxs" by simp qed lemma odlist_cases [case_names empty insert, cases type: odlist]: assumes empty: "dxs = empty \ P" assumes insert: "\x xs. \ dxs = fromList (x # xs); distinct (x # xs); sorted (x # xs) \ \ P" shows P proof (cases dxs) case (odlist_Abs xs) then have dxs: "dxs = ODList xs" and distinct: "distinct xs" and sorted: "sorted xs" by (simp_all add: ODList_def) show P proof (cases xs) case Nil with dxs have "dxs = empty" by (simp add: empty_def) with empty show P . next case (Cons y ys) with dxs distinct sorted insert show P by (simp add: fromList_def) qed qed subsubsection\Relations\ text\ Relations, represented as a list of pairs. \ type_synonym 'a odrelation = "('a \ 'a) odlist" subsubsection\Image\ text\ The output of @{term "List_local.image"} is not guaranteed to be ordered or distinct. Also the relation need not be monomorphic. \ definition image :: "('a :: linorder \ 'b :: linorder) odlist \ 'a odlist \ 'b odlist" where "image R xs = ODList (List_local.image (toList R) (toList xs))" lemma toList_image[simp, code abstract]: "toList (image R xs) = msort (List_local.image (toList R) (toList xs))" unfolding image_def by (simp add: toList_ODList) lemma toSet_image[simp]: "toSet (image R xs) = toSet R `` toSet xs" unfolding image_def by (simp add: toSet_def toList_ODList) subsubsection\Linear order\ text\ Lexicographic ordering on lists. Executable, unlike in List.thy. \ instantiation odlist :: (linorder) linorder begin print_context fun less_eq_list :: "'a list \ 'a list \ bool" where "less_eq_list [] ys = True" | "less_eq_list xs [] = False" | "less_eq_list (x # xs) (y # ys) = (x < y \ (x = y \ less_eq_list xs ys))" lemma less_eq_list_nil_inv: fixes xs :: "'a list" shows "less_eq_list xs [] \ xs = []" by (cases xs) simp_all lemma less_eq_list_cons_inv: fixes x :: 'a shows "less_eq_list (x # xs) yys \ \y ys. yys = y # ys \ (x < y \ (x = y \ less_eq_list xs ys))" by (cases yys) auto lemma less_eq_list_refl: fixes xs :: "'a list" shows "less_eq_list xs xs" by (induct xs) simp_all lemma less_eq_list_trans: fixes xs ys zs :: "'a list" shows "\ less_eq_list xs ys; less_eq_list ys zs \ \ less_eq_list xs zs" apply (induct xs ys arbitrary: zs rule: less_eq_list.induct) apply simp apply simp apply clarsimp apply (erule disjE) apply (drule less_eq_list_cons_inv) apply clarsimp apply (erule disjE) apply auto[1] apply auto[1] apply (auto dest: less_eq_list_cons_inv) done lemma less_eq_list_antisym: fixes xs ys :: "'a list" shows "\ less_eq_list xs ys; less_eq_list ys xs \ \ xs = ys" by (induct xs ys rule: less_eq_list.induct) (auto dest: less_eq_list_nil_inv) lemma less_eq_list_linear: fixes xs ys :: "'a list" shows "less_eq_list xs ys \ less_eq_list ys xs" by (induct xs ys rule: less_eq_list.induct) auto definition less_eq_odlist :: "'a odlist \ 'a odlist \ bool" where "xs \ ys \ less_eq_list (toList xs) (toList ys)" fun less_list :: "'a list \ 'a list \ bool" where "less_list [] [] = False" | "less_list [] ys = True" | "less_list xs [] = False" | "less_list (x # xs) (y # ys) = (x < y \ (x = y \ less_list xs ys))" definition less_odlist :: "'a odlist \ 'a odlist \ bool" where "xs < ys \ less_list (toList xs) (toList ys)" lemma less_eq_list_not_le: fixes xs ys :: "'a list" shows "(less_list xs ys) = (less_eq_list xs ys \ \ less_eq_list ys xs)" by (induct xs ys rule: less_list.induct) auto instance apply intro_classes unfolding less_eq_odlist_def less_odlist_def using less_eq_list_not_le less_eq_list_refl less_eq_list_trans less_eq_list_antisym apply blast using less_eq_list_not_le less_eq_list_refl less_eq_list_trans less_eq_list_antisym apply blast using less_eq_list_not_le less_eq_list_refl less_eq_list_trans less_eq_list_antisym apply blast using less_eq_list_not_le less_eq_list_refl less_eq_list_trans less_eq_list_antisym apply blast apply (rule less_eq_list_linear) done end subsubsection\Finite maps\ text\ A few operations on finite maps. Unlike the AssocList theory, ODLists give us canonical representations, so we can order them. Our tabulate has the wrong type (we want to take an odlist, not a list) so we can't use that part of the framework. \ definition lookup :: "('a :: linorder \ 'b :: linorder) odlist \ ('a \ 'b)" where [code]: "lookup = map_of \ toList" text\Specific to ODLists.\ definition tabulate :: "('a :: linorder) odlist \ ('a \ 'b :: linorder) \ ('a \ 'b) odlist" where "tabulate ks f = ODList (List.map (\k. (k, f k)) (toList ks))" definition (in order) mono_on :: "('a \ 'b::order) \ 'a set \ bool" where "mono_on f X \ (\x\X. \y\X. x \ y \ f x \ f y)" lemma (in order) mono_onI [intro?]: fixes f :: "'a \ 'b::order" shows "(\x y. x \ X \ y \ X \ x \ y \ f x \ f y) \ mono_on f X" unfolding mono_on_def by simp lemma (in order) mono_onD [dest?]: fixes f :: "'a \ 'b::order" shows "mono_on f X \ x \ X \ y \ X \ x \ y \ f x \ f y" unfolding mono_on_def by simp lemma (in order) mono_on_subset: fixes f :: "'a \ 'b::order" shows "mono_on f X \ Y \ X \ mono_on f Y" unfolding mono_on_def by auto lemma sorted_mono_map: "\ sorted xs; mono_on f (set xs) \ \ sorted (List.map f xs)" apply (induct xs) apply simp apply (simp) apply (cut_tac X="insert a (set xs)" and Y="set xs" in mono_on_subset) apply (auto dest: mono_onD) done lemma msort_map: "\ distinct xs; sorted xs; inj_on f (set xs); mono_on f (set xs) \ \ msort (List.map f xs) = List.map f xs" apply (rule msort_idle) apply (simp add: distinct_map) apply (simp add: sorted_mono_map) done lemma tabulate_toList[simp, code abstract]: "toList (tabulate ks f) = List.map (\k. (k, f k)) (toList ks)" unfolding tabulate_def apply (simp add: toList_ODList) apply (subst msort_map) apply simp_all apply (rule inj_onI) apply simp apply (rule mono_onI) apply (simp add: less_eq_prod_def less_le) done lemma lookup_tabulate[simp]: "lookup (tabulate ks f) = (Some o f) |` toSet ks" proof(induct ks rule: odlist_induct) case (empty dxs) thus ?case unfolding tabulate_def lookup_def by (simp add: toList_ODList) next case (insert dxs x xs) from insert have "map_of (List.map (\k. (k, f k)) xs) = map_of (msort (List.map (\k. (k, f k)) xs))" apply (subst msort_map) apply (auto intro: inj_onI) apply (rule mono_onI) apply (simp add: less_eq_prod_def less_le) done also from insert have "... = lookup (tabulate (fromList xs) f)" unfolding tabulate_def lookup_def by (simp add: toList_ODList toList_fromList) also from insert have "... = (Some \ f) |` toSet (fromList xs)" by (simp only: toSet_fromList_set) finally have IH: "map_of (List.map (\k. (k, f k)) xs) = (Some \ f) |` toSet (fromList xs)" . from insert have "lookup (tabulate dxs f) = map_of (toList (ODList (List.map (\k. (k, f k)) (x # xs))))" unfolding tabulate_def lookup_def by (simp add: toList_fromList) also have "... = map_of (msort (List.map (\k. (k, f k)) (x # xs)))" by (simp only: toList_ODList) also from insert have "... = map_of (List.map (\k. (k, f k)) (x # xs))" - apply (subst msort_map) - apply (auto intro: inj_onI) - apply (rule mono_onI) - apply (simp add: less_eq_prod_def less_le) - done + by (metis msort_idle tabulate_def tabulate_toList toList_ODList) also with insert IH have "... = (Some \ f) |` toSet dxs" by (auto simp add: restrict_map_def fun_eq_iff) finally show ?case . qed (*<*) end (*>*) diff --git a/thys/List-Index/List_Index.thy b/thys/List-Index/List_Index.thy --- a/thys/List-Index/List_Index.thy +++ b/thys/List-Index/List_Index.thy @@ -1,561 +1,561 @@ (* Author: Tobias Nipkow, with contributions from Dmitriy Traytel, Lukas Bulwahn, and Peter Lammich *) section \Index-based manipulation of lists\ theory List_Index imports Main begin text \\noindent This theory collects functions for index-based manipulation of lists. \ subsection \Finding an index\ text \ This subsection defines three functions for finding the index of items in a list: \begin{description} \item[\find_index P xs\] finds the index of the first element in \xs\ that satisfies \P\. \item[\index xs x\] finds the index of the first occurrence of \x\ in \xs\. \item[\last_index xs x\] finds the index of the last occurrence of \x\ in \xs\. \end{description} All functions return @{term "length xs"} if \xs\ does not contain a suitable element. The argument order of \find_index\ follows the function of the same name in the Haskell standard library. For \index\ (and \last_index\) the order is intentionally reversed: \index\ maps lists to a mapping from elements to their indices, almost the inverse of function \nth\.\ primrec find_index :: "('a \ bool) \ 'a list \ nat" where "find_index _ [] = 0" | "find_index P (x#xs) = (if P x then 0 else find_index P xs + 1)" definition index :: "'a list \ 'a \ nat" where "index xs = (\a. find_index (\x. x=a) xs)" definition last_index :: "'a list \ 'a \ nat" where "last_index xs x = (let i = index (rev xs) x; n = size xs in if i = n then i else n - (i+1))" lemma find_index_append: "find_index P (xs @ ys) = (if \x\set xs. P x then find_index P xs else size xs + find_index P ys)" by (induct xs) simp_all lemma find_index_le_size: "find_index P xs <= size xs" by(induct xs) simp_all lemma index_le_size: "index xs x <= size xs" by(simp add: index_def find_index_le_size) lemma last_index_le_size: "last_index xs x <= size xs" by(simp add: last_index_def Let_def index_le_size) lemma index_Nil[simp]: "index [] a = 0" by(simp add: index_def) lemma index_Cons[simp]: "index (x#xs) a = (if x=a then 0 else index xs a + 1)" by(simp add: index_def) lemma index_append: "index (xs @ ys) x = (if x : set xs then index xs x else size xs + index ys x)" by (induct xs) simp_all lemma index_conv_size_if_notin[simp]: "x \ set xs \ index xs x = size xs" by (induct xs) auto lemma find_index_eq_size_conv: "size xs = n \ (find_index P xs = n) = (\x \ set xs. ~ P x)" by(induct xs arbitrary: n) auto lemma size_eq_find_index_conv: "size xs = n \ (n = find_index P xs) = (\x \ set xs. ~ P x)" by(metis find_index_eq_size_conv) lemma index_size_conv: "size xs = n \ (index xs x = n) = (x \ set xs)" by(auto simp: index_def find_index_eq_size_conv) lemma size_index_conv: "size xs = n \ (n = index xs x) = (x \ set xs)" by (metis index_size_conv) lemma last_index_size_conv: "size xs = n \ (last_index xs x = n) = (x \ set xs)" apply(auto simp: last_index_def index_size_conv) apply(drule length_pos_if_in_set) apply arith done lemma size_last_index_conv: "size xs = n \ (n = last_index xs x) = (x \ set xs)" by (metis last_index_size_conv) lemma find_index_less_size_conv: "(find_index P xs < size xs) = (\x \ set xs. P x)" by (induct xs) auto lemma index_less_size_conv: "(index xs x < size xs) = (x \ set xs)" by(auto simp: index_def find_index_less_size_conv) lemma last_index_less_size_conv: "(last_index xs x < size xs) = (x : set xs)" by(simp add: last_index_def Let_def index_size_conv length_pos_if_in_set del:length_greater_0_conv) lemma index_less[simp]: "x : set xs \ size xs <= n \ index xs x < n" apply(induct xs) apply auto apply (metis index_less_size_conv less_eq_Suc_le less_trans_Suc) done lemma last_index_less[simp]: "x : set xs \ size xs <= n \ last_index xs x < n" by(simp add: last_index_less_size_conv[symmetric]) lemma last_index_Cons: "last_index (x#xs) y = (if x=y then if x \ set xs then last_index xs y + 1 else 0 else last_index xs y + 1)" using index_le_size[of "rev xs" y] apply(auto simp add: last_index_def index_append Let_def) apply(simp add: index_size_conv) done lemma last_index_append: "last_index (xs @ ys) x = (if x : set ys then size xs + last_index ys x else if x : set xs then last_index xs x else size xs + size ys)" by (induct xs) (simp_all add: last_index_Cons last_index_size_conv) lemma last_index_Snoc[simp]: "last_index (xs @ [x]) y = (if x=y then size xs else if y : set xs then last_index xs y else size xs + 1)" by(simp add: last_index_append last_index_Cons) lemma nth_find_index: "find_index P xs < size xs \ P(xs ! find_index P xs)" by (induct xs) auto lemma nth_index[simp]: "x \ set xs \ xs ! index xs x = x" by (induct xs) auto lemma nth_last_index[simp]: "x \ set xs \ xs ! last_index xs x = x" by(simp add:last_index_def index_size_conv Let_def rev_nth[symmetric]) lemma index_rev: "\ distinct xs; x \ set xs \ \ index (rev xs) x = length xs - index xs x - 1" by (induct xs) (auto simp: index_append) lemma index_nth_id: "\ distinct xs; n < length xs \ \ index xs (xs ! n) = n" by (metis in_set_conv_nth index_less_size_conv nth_eq_iff_index_eq nth_index) lemma index_upt[simp]: "m \ i \ i < n \ index [m.. set xs \ y \ set xs \ (index xs x = index xs y) = (x = y)" by (induct xs) auto lemma last_index_eq_index_conv[simp]: "x \ set xs \ y \ set xs \ (last_index xs x = last_index xs y) = (x = y)" by (induct xs) (auto simp:last_index_Cons) lemma inj_on_index: "inj_on (index xs) (set xs)" by (simp add:inj_on_def) lemma inj_on_index2: "I \ set xs \ inj_on (index xs) I" by (rule inj_onI) auto lemma inj_on_last_index: "inj_on (last_index xs) (set xs)" by (simp add:inj_on_def) lemma find_index_conv_takeWhile: "find_index P xs = size(takeWhile (Not o P) xs)" by(induct xs) auto lemma index_conv_takeWhile: "index xs x = size(takeWhile (\y. x\y) xs)" by(induct xs) auto lemma find_index_first: "i < find_index P xs \ \P (xs!i)" unfolding find_index_conv_takeWhile by (metis comp_apply nth_mem set_takeWhileD takeWhile_nth) lemma index_first: "i x\xs!i" using find_index_first unfolding index_def by blast lemma find_index_eqI: assumes "i\length xs" assumes "\jP (xs!j)" assumes "i P (xs!i)" shows "find_index P xs = i" by (metis (mono_tags, lifting) antisym_conv2 assms find_index_eq_size_conv find_index_first find_index_less_size_conv linorder_neqE_nat nth_find_index) lemma find_index_eq_iff: "find_index P xs = i \ (i\length xs \ (\jP (xs!j)) \ (i P (xs!i)))" by (auto intro: find_index_eqI simp: nth_find_index find_index_le_size find_index_first) lemma index_eqI: assumes "i\length xs" assumes "\j x" assumes "i xs!i = x" shows "index xs x = i" unfolding index_def by (simp add: find_index_eqI assms) lemma index_eq_iff: "index xs x = i \ (i\length xs \ (\j x) \ (i xs!i = x))" by (auto intro: index_eqI simp: index_le_size index_less_size_conv dest: index_first) lemma index_take: "index xs x >= i \ x \ set(take i xs)" apply(subst (asm) index_conv_takeWhile) apply(subgoal_tac "set(take i xs) <= set(takeWhile ((\) x) xs)") apply(blast dest: set_takeWhileD) apply(metis set_take_subset_set_take takeWhile_eq_take) done lemma last_index_drop: "last_index xs x < i \ x \ set(drop i xs)" apply(subgoal_tac "set(drop i xs) = set(take (size xs - i) (rev xs))") apply(simp add: last_index_def index_take Let_def split:if_split_asm) apply (metis rev_drop set_rev) done lemma set_take_if_index: assumes "index xs x < i" and "i \ length xs" shows "x \ set (take i xs)" proof - have "index (take i xs @ drop i xs) x < i" using append_take_drop_id[of i xs] assms(1) by simp thus ?thesis using assms(2) by(simp add:index_append del:append_take_drop_id split: if_splits) qed lemma index_take_if_index: assumes "index xs x \ n" shows "index (take n xs) x = index xs x" proof cases assume "x : set(take n xs)" with assms show ?thesis by (metis append_take_drop_id index_append) next assume "x \ set(take n xs)" with assms show ?thesis by (metis order_le_less set_take_if_index le_cases length_take min_def size_index_conv take_all) qed lemma index_take_if_set: "x : set(take n xs) \ index (take n xs) x = index xs x" by (metis index_take index_take_if_index linear) lemma index_last[simp]: "xs \ [] \ distinct xs \ index xs (last xs) = length xs - 1" by (induction xs) auto lemma index_update_if_diff2: "n < length xs \ x \ xs!n \ x \ y \ index (xs[n := y]) x = index xs x" by(subst (2) id_take_nth_drop[of n xs]) (auto simp: upd_conv_take_nth_drop index_append min_def) lemma set_drop_if_index: "distinct xs \ index xs x < i \ x \ set(drop i xs)" by (metis in_set_dropD index_nth_id last_index_drop last_index_less_size_conv nth_last_index) lemma index_swap_if_distinct: assumes "distinct xs" "i < size xs" "j < size xs" shows "index (xs[i := xs!j, j := xs!i]) x = (if x = xs!i then j else if x = xs!j then i else index xs x)" proof- have "distinct(xs[i := xs!j, j := xs!i])" using assms by simp with assms show ?thesis apply (auto simp: simp del: distinct_swap) apply (metis index_nth_id list_update_same_conv) apply (metis (erased, hide_lams) index_nth_id length_list_update list_update_swap nth_list_update_eq) apply (metis index_nth_id length_list_update nth_list_update_eq) by (metis index_update_if_diff2 length_list_update nth_list_update) qed lemma bij_betw_index: "distinct xs \ X = set xs \ l = size xs \ bij_betw (index xs) X {0.. set xs = X \ index xs ` X = {0.. inj_on f S; y \ S; set xs \ S \ \ index (map f xs) (f y) = index xs y" by (induct xs) (auto simp: inj_on_eq_iff) lemma index_map_inj: "inj f \ index (map f xs) (f y) = index xs y" by (simp add: index_map_inj_on[where S=UNIV]) subsection \Map with index\ primrec map_index' :: "nat \ (nat \ 'a \ 'b) \ 'a list \ 'b list" where "map_index' n f [] = []" | "map_index' n f (x#xs) = f n x # map_index' (Suc n) f xs" lemma length_map_index'[simp]: "length (map_index' n f xs) = length xs" by (induct xs arbitrary: n) auto lemma map_index'_map_zip: "map_index' n f xs = map (case_prod f) (zip [n ..< n + length xs] xs)" proof (induct xs arbitrary: n) case (Cons x xs) hence "map_index' n f (x#xs) = f n x # map (case_prod f) (zip [Suc n ..< n + length (x # xs)] xs)" by simp also have "\ = map (case_prod f) (zip (n # [Suc n ..< n + length (x # xs)]) (x # xs))" by simp also have "(n # [Suc n ..< n + length (x # xs)]) = [n ..< n + length (x # xs)]" by (induct xs) auto finally show ?case by simp qed simp abbreviation "map_index \ map_index' 0" lemmas map_index = map_index'_map_zip[of 0, simplified] lemma take_map_index: "take p (map_index f xs) = map_index f (take p xs)" unfolding map_index by (auto simp: min_def take_map take_zip) lemma drop_map_index: "drop p (map_index f xs) = map_index' p f (drop p xs)" unfolding map_index'_map_zip by (cases "p < length xs") (auto simp: drop_map drop_zip) lemma map_map_index[simp]: "map g (map_index f xs) = map_index (\n x. g (f n x)) xs" unfolding map_index by auto lemma map_index_map[simp]: "map_index f (map g xs) = map_index (\n x. f n (g x)) xs" unfolding map_index by (auto simp: map_zip_map2) lemma set_map_index[simp]: "x \ set (map_index f xs) = (\i < length xs. f i (xs ! i) = x)" unfolding map_index by (auto simp: set_zip intro!: image_eqI[of _ "case_prod f"]) lemma set_map_index'[simp]: "x\set (map_index' n f xs) \ (\i map_index f xs ! p = f p (xs ! p)" unfolding map_index by auto lemma map_index_cong: "\p < length xs. f p (xs ! p) = g p (xs ! p) \ map_index f xs = map_index g xs" unfolding map_index by (auto simp: set_zip) lemma map_index_id: "map_index (curry snd) xs = xs" unfolding map_index by auto lemma map_index_no_index[simp]: "map_index (\n x. f x) xs = map f xs" unfolding map_index by (induct xs rule: rev_induct) auto lemma map_index_congL: "\p < length xs. f p (xs ! p) = xs ! p \ map_index f xs = xs" by (rule trans[OF map_index_cong map_index_id]) auto lemma map_index'_is_NilD: "map_index' n f xs = [] \ xs = []" by (induct xs) auto declare map_index'_is_NilD[of 0, dest!] lemma map_index'_is_ConsD: "map_index' n f xs = y # ys \ \z zs. xs = z # zs \ f n z = y \ map_index' (n + 1) f zs = ys" by (induct xs arbitrary: n) auto lemma map_index'_eq_imp_length_eq: "map_index' n f xs = map_index' n g ys \ length xs = length ys" proof (induct ys arbitrary: xs n) case (Cons y ys) thus ?case by (cases xs) auto qed (auto dest!: map_index'_is_NilD) lemmas map_index_eq_imp_length_eq = map_index'_eq_imp_length_eq[of 0] lemma map_index'_comp[simp]: "map_index' n f (map_index' n g xs) = map_index' n (\n. f n o g n) xs" by (induct xs arbitrary: n) auto lemma map_index'_append[simp]: "map_index' n f (a @ b) = map_index' n f a @ map_index' (n + length a) f b" by (induct a arbitrary: n) auto lemma map_index_append[simp]: "map_index f (a @ b) = map_index f a @ map_index' (length a) f b" using map_index'_append[where n=0] by (simp del: map_index'_append) subsection \Insert at position\ primrec insert_nth :: "nat \ 'a \ 'a list \ 'a list" where "insert_nth 0 x xs = x # xs" | "insert_nth (Suc n) x xs = (case xs of [] \ [x] | y # ys \ y # insert_nth n x ys)" lemma insert_nth_take_drop[simp]: "insert_nth n x xs = take n xs @ [x] @ drop n xs" proof (induct n arbitrary: xs) case Suc thus ?case by (cases xs) auto qed simp lemma length_insert_nth: "length (insert_nth n x xs) = Suc (length xs)" by (induct xs) auto lemma set_insert_nth: "set (insert_nth i x xs) = insert x (set xs)" by (simp add: set_append[symmetric]) lemma distinct_insert_nth: assumes "distinct xs" assumes "x \ set xs" shows "distinct (insert_nth i x xs)" using assms proof (induct xs arbitrary: i) case Nil then show ?case by (cases i) auto next case (Cons a xs) then show ?case by (cases i) (auto simp add: set_insert_nth simp del: insert_nth_take_drop) qed lemma nth_insert_nth_front: assumes "i < j" "j \ length xs" shows "insert_nth j x xs ! i = xs ! i" using assms by (simp add: nth_append) lemma nth_insert_nth_index_eq: assumes "i \ length xs" shows "insert_nth i x xs ! i = x" using assms by (simp add: nth_append) lemma nth_insert_nth_back: assumes "j < i" "i \ length xs" shows "insert_nth j x xs ! i = xs ! (i - 1)" using assms by (cases i) (auto simp add: nth_append min_def) lemma nth_insert_nth: assumes "i \ length xs" "j \ length xs" shows "insert_nth j x xs ! i = (if i = j then x else if i < j then xs ! i else xs ! (i - 1))" using assms by (simp add: nth_insert_nth_front nth_insert_nth_index_eq nth_insert_nth_back del: insert_nth_take_drop) lemma insert_nth_inverse: assumes "j \ length xs" "j' \ length xs'" assumes "x \ set xs" "x \ set xs'" assumes "insert_nth j x xs = insert_nth j' x xs'" shows "j = j'" proof - from assms(1,3) have "\i\length xs. insert_nth j x xs ! i = x \ i = j" by (auto simp add: nth_insert_nth simp del: insert_nth_take_drop) moreover from assms(2,4) have "\i\length xs'. insert_nth j' x xs' ! i = x \ i = j'" by (auto simp add: nth_insert_nth simp del: insert_nth_take_drop) ultimately show "j = j'" using assms(1,2,5) by (metis dual_order.trans nat_le_linear) qed text \Insert several elements at given (ascending) positions\ lemma length_fold_insert_nth: "length (fold (\(p, b). insert_nth p b) pxs xs) = length xs + length pxs" by (induct pxs arbitrary: xs) auto lemma invar_fold_insert_nth: "\\x\set pxs. p < fst x; p < length xs; xs ! p = b\ \ fold (\(x, y). insert_nth x y) pxs xs ! p = b" by (induct pxs arbitrary: xs) (auto simp: nth_append) lemma nth_fold_insert_nth: "\sorted (map fst pxs); distinct (map fst pxs); \(p, b) \ set pxs. p < length xs + length pxs; i < length pxs; pxs ! i = (p, b)\ \ fold (\(p, b). insert_nth p b) pxs xs ! p = b" proof (induct pxs arbitrary: xs i p b) case (Cons pb pxs) show ?case proof (cases i) case 0 with Cons.prems have "p < Suc (length xs)" proof (induct pxs rule: rev_induct) case (snoc pb' pxs) then obtain p' b' where "pb' = (p', b')" by auto with snoc.prems have "\p \ fst ` set pxs. p < p'" "p' \ Suc (length xs + length pxs)" - by (auto simp: image_iff sorted_append le_eq_less_or_eq) + by (auto simp: image_iff sorted_wrt_append le_eq_less_or_eq) with snoc.prems show ?case by (intro snoc(1)) (auto simp: sorted_append) qed auto with 0 Cons.prems show ?thesis unfolding fold.simps o_apply by (intro invar_fold_insert_nth) (auto simp: image_iff le_eq_less_or_eq nth_append) next case (Suc n) with Cons.prems show ?thesis unfolding fold.simps by (auto intro!: Cons(1)) qed qed simp subsection \Remove at position\ fun remove_nth :: "nat \ 'a list \ 'a list" where "remove_nth i [] = []" | "remove_nth 0 (x # xs) = xs" | "remove_nth (Suc i) (x # xs) = x # remove_nth i xs" lemma remove_nth_take_drop: "remove_nth i xs = take i xs @ drop (Suc i) xs" proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) auto qed lemma remove_nth_insert_nth: assumes "i \ length xs" shows "remove_nth i (insert_nth i x xs) = xs" using assms proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) auto qed lemma insert_nth_remove_nth: assumes "i < length xs" shows "insert_nth i (xs ! i) (remove_nth i xs) = xs" using assms proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) auto qed lemma length_remove_nth: assumes "i < length xs" shows "length (remove_nth i xs) = length xs - 1" using assms unfolding remove_nth_take_drop by simp lemma set_remove_nth_subset: "set (remove_nth j xs) \ set xs" proof (induct xs arbitrary: j) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases j) auto qed lemma set_remove_nth: assumes "distinct xs" "j < length xs" shows "set (remove_nth j xs) = set xs - {xs ! j}" using assms proof (induct xs arbitrary: j) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases j) auto qed lemma distinct_remove_nth: assumes "distinct xs" shows "distinct (remove_nth i xs)" using assms proof (induct xs arbitrary: i) case Nil then show ?case by simp next case (Cons a xs) then show ?case by (cases i) (auto simp add: set_remove_nth_subset rev_subsetD) qed end diff --git a/thys/List_Inversions/List_Inversions.thy b/thys/List_Inversions/List_Inversions.thy --- a/thys/List_Inversions/List_Inversions.thy +++ b/thys/List_Inversions/List_Inversions.thy @@ -1,445 +1,444 @@ (* File: List_Inversions.thy Author: Manuel Eberl, TU München A formalisation of inversions of a list and the O(n log n) divide-and-conquer algorithm to count them. *) section \The Inversions of a List\ theory List_Inversions imports Main "HOL-Combinatorics.Permutations" begin subsection \Definition of inversions\ context preorder begin text \ We define inversions as pair of indices w.\,r.\,t.\ a preorder. \ inductive_set inversions :: "'a list \ (nat \ nat) set" for xs :: "'a list" where "i < j \ j < length xs \ less (xs ! j) (xs ! i) \ (i, j) \ inversions xs" lemma inversions_subset: "inversions xs \ Sigma {..i. {i<.. j < length xs \ less (xs ! j) (xs ! i)}" by (auto simp: inversions.simps) lemma inversions_code: "inversions xs = Sigma {..i. Set.filter (\j. less (xs ! j) (xs ! i)) {i<.. Suc 0 \ inversions xs = {}" by (auto simp: inversions_altdef) lemma inversions_imp_less: "z \ inversions xs \ fst z < snd z" "z \ inversions xs \ snd z < length xs" by (auto simp: inversions_altdef) lemma inversions_Nil [simp]: "inversions [] = {}" by (auto simp: inversions_altdef) lemma inversions_Cons: "inversions (x # xs) = (\j. (0, j + 1)) ` {j\{.. map_prod Suc Suc ` inversions xs" (is "_ = ?rhs") proof - have "z \ inversions (x # xs) \ z \ ?rhs" for z by (cases z) (auto simp: inversions_altdef map_prod_def nth_Cons split: nat.splits) thus ?thesis by blast qed text \ The following function returns the inversions between two lists, i.\,e.\ all pairs of an element in the first list with an element in the second list such that the former is greater than the latter. \ definition inversions_between :: "'a list \ 'a list \ (nat \ nat) set" where "inversions_between xs ys = {(i, j) \ {..{.. {.. We can now show the following equality for the inversions of the concatenation of two lists: \ proposition inversions_append: fixes xs ys defines "m \ length xs" and "n \ length ys" shows "inversions (xs @ ys) = inversions xs \ map_prod ((+) m) ((+) m) ` inversions ys \ map_prod id ((+) m) ` inversions_between xs ys" (is "_ = ?rhs") proof - note defs = inversions_altdef inversions_between_def m_def n_def map_prod_def have "z \ inversions (xs @ ys) \ z \ ?rhs" for z proof assume "z \ inversions (xs @ ys)" then obtain i j where [simp]: "z = (i, j)" and ij: "i < j" "j < m + n" "less ((xs @ ys) ! j) ((xs @ ys) ! i)" by (cases z) (auto simp: inversions_altdef m_def n_def) from ij consider "j < m" | "i \ m" | "i < m" "j \ m" by linarith thus "z \ ?rhs" proof cases assume "i < m" "j \ m" define j' where "j' = j - m" have [simp]: "j = m + j'" using \j \ m\ by (simp add: j'_def) from ij and \i < m\ show ?thesis by (auto simp: inversions_altdef map_prod_def inversions_between_def nth_append m_def n_def) next assume "i \ m" define i' j' where "i' = i - m" and "j' = j - m" have [simp]: "i = m + i'" "j = m + j'" using \i < j\ and \i \ m\ by (simp_all add: i'_def j'_def) from ij show ?thesis by (auto simp: inversions_altdef map_prod_def nth_append m_def n_def) qed (use ij in \auto simp: nth_append defs\) qed (auto simp: nth_append defs) thus ?thesis by blast qed subsection \Counting inversions\ text \ We now define versions of @{const inversions} and @{const inversions_between} that only return the \<^emph>\number\ of inversions. \ definition inversion_number :: "'a list \ nat" where "inversion_number xs = card (inversions xs)" definition inversion_number_between where "inversion_number_between xs ys = card (inversions_between xs ys)" lemma inversions_between_code: "inversions_between xs ys = Set.filter (\(i,j). less (ys ! j) (xs ! i)) ({..{.. Suc 0 \ inversion_number xs = 0" by (auto simp: inversion_number_def) lemma inversion_number_between_Nil [simp]: "inversion_number_between [] ys = 0" "inversion_number_between xs [] = 0" by (simp_all add: inversion_number_between_def) text \ We again get the following nice equation for the number of inversions of a concatenation: \ proposition inversion_number_append: "inversion_number (xs @ ys) = inversion_number xs + inversion_number ys + inversion_number_between xs ys" proof - define m n where "m = length xs" and "n = length ys" let ?A = "inversions xs" let ?B = "map_prod ((+) m) ((+) m) ` inversions ys" let ?C = "map_prod id ((+) m) ` inversions_between xs ys" have "inversion_number (xs @ ys) = card (?A \ ?B \ ?C)" by (simp add: inversion_number_def inversions_append m_def) also have "\ = card (?A \ ?B) + card ?C" by (intro card_Un_disjoint finite_inversions finite_inversions_between finite_UnI finite_imageI) (auto simp: inversions_altdef inversions_between_def m_def n_def) also have "card (?A \ ?B) = inversion_number xs + card ?B" unfolding inversion_number_def by (intro card_Un_disjoint finite_inversions finite_UnI finite_imageI) (auto simp: inversions_altdef m_def n_def) also have "card ?B = inversion_number ys" unfolding inversion_number_def by (intro card_image) (auto simp: map_prod_def inj_on_def) also have "card ?C = inversion_number_between xs ys" unfolding inversion_number_between_def by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis . qed subsection \Stability of inversions between lists under permutations\ text \ A crucial fact for counting list inversions with merge sort is that the number of inversions \<^emph>\between\ two lists does not change when the lists are permuted. This is true because the set of inversions commutes with the act of permuting the list: \ lemma inversions_between_permute1: assumes "\ permutes {.. xs) ys = map_prod (inv \) id ` inversions_between xs ys" proof - from assms have [simp]: "\ i < length xs" if "i < length xs" "\ permutes {.. using permutes_in_image[OF that(2)] that by auto have *: "inv \ permutes {.. i" for i]) qed lemma inversions_between_permute2: assumes "\ permutes {.. ys) = map_prod id (inv \) ` inversions_between xs ys" proof - from assms have [simp]: "\ i < length ys" if "i < length ys" "\ permutes {.. using permutes_in_image[OF that(2)] that by auto have *: "inv \ permutes {.. i" for i]) qed proposition inversions_between_permute: assumes "\1 permutes {..2 permutes {..1 xs) (permute_list \2 ys) = map_prod (inv \1) (inv \2) ` inversions_between xs ys" by (simp add: inversions_between_permute1 inversions_between_permute2 assms map_prod_def image_image case_prod_unfold) corollary inversion_number_between_permute: assumes "\1 permutes {..2 permutes {..1 xs) (permute_list \2 ys) = inversion_number_between xs ys" proof - have "inversion_number_between (permute_list \1 xs) (permute_list \2 ys) = card (map_prod (inv \1) (inv \2) ` inversions_between xs ys)" by (simp add: inversion_number_between_def inversions_between_permute assms) also have "\ = inversion_number_between xs ys" unfolding inversion_number_between_def using assms[THEN permutes_inj_on[OF permutes_inv]] by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis . qed text \ The following form of the above theorem is nicer to apply since it has the form of a congruence rule. \ corollary inversion_number_between_cong_mset: assumes "mset xs = mset xs'" and "mset ys = mset ys'" shows "inversion_number_between xs ys = inversion_number_between xs' ys'" proof - obtain \1 \2 where \12: "\1 permutes {..1 xs'" "\2 permutes {..2 ys'" using assms[THEN mset_eq_permutation] by metis thus ?thesis by (simp add: inversion_number_between_permute) qed subsection \Inversions between sorted lists\ text \ Another fact that is crucial to the efficient computation of the inversion number is this: If we have two sorted lists, we can reduce computing the inversions by inspecting the first elements and deleting one of them. \ lemma inversions_between_Cons_Cons: assumes "sorted_wrt less_eq (x # xs)" and "sorted_wrt less_eq (y # ys)" shows "inversions_between (x # xs) (y # ys) = (if \less y x then map_prod Suc id ` inversions_between xs (y # ys) else {.. {0} \ map_prod id Suc ` inversions_between (x # xs) ys)" using assms unfolding inversions_between_def map_prod_def by (auto, (auto simp: set_conv_nth nth_Cons less_le_not_le image_iff intro: order_trans split: nat.splits)?) (* A bit fragile, but doing this manually is annoying *) text \ This leads to the following analogous equation for counting the inversions between two sorted lists. Note that a single step of this only takes constant time (assuming we pre-computed the lengths of the lists) so that the entire function runs in linear time. \ lemma inversion_number_between_Cons_Cons: assumes "sorted_wrt less_eq (x # xs)" and "sorted_wrt less_eq (y # ys)" shows "inversion_number_between (x # xs) (y # ys) = (if \less y x then inversion_number_between xs (y # ys) else inversion_number_between (x # xs) ys + length (x # xs))" proof (cases "less y x") case False hence "inversion_number_between (x # xs) (y # ys) = card (map_prod Suc id ` inversions_between xs (y # ys))" by (simp add: inversion_number_between_def inversions_between_Cons_Cons[OF assms]) also have "\ = inversion_number_between xs (y # ys)" unfolding inversion_number_between_def by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis using False by simp next case True hence "inversion_number_between (x # xs) (y # ys) = card ({.. {0} \ map_prod id Suc ` inversions_between (x # xs) ys)" by (simp add: inversion_number_between_def inversions_between_Cons_Cons[OF assms]) also have "\ = length (x # xs) + card (map_prod id Suc ` inversions_between (x # xs) ys)" by (subst card_Un_disjoint) auto also have "card (map_prod id Suc ` inversions_between (x # xs) ys) = inversion_number_between (x # xs) ys" unfolding inversion_number_between_def by (intro card_image inj_onI) (auto simp: map_prod_def) finally show ?thesis using True by simp qed text \ We now define a function to compute the inversion number between two lists that are assumed to be sorted using the equalities we just derived. \ fun inversion_number_between_sorted :: "'a list \ 'a list \ nat" where "inversion_number_between_sorted [] ys = 0" | "inversion_number_between_sorted xs [] = 0" | "inversion_number_between_sorted (x # xs) (y # ys) = (if \less y x then inversion_number_between_sorted xs (y # ys) else inversion_number_between_sorted (x # xs) ys + length (x # xs))" theorem inversion_number_between_sorted_correct: "sorted_wrt less_eq xs \ sorted_wrt less_eq ys \ inversion_number_between_sorted xs ys = inversion_number_between xs ys" by (induction xs ys rule: inversion_number_between_sorted.induct) (simp_all add: inversion_number_between_Cons_Cons) end subsection \Merge sort\ (* TODO: Could be replaced by mergesort from HOL-Library in Isabelle 2019. *) text \ For convenience, we first define a simple merge sort that does not compute the inversions. At this point, we need to start assuming a linear ordering since the merging function does not work otherwise. \ context linorder begin definition split_list where "split_list xs = (let n = length xs div 2 in (take n xs, drop n xs))" fun merge_lists :: "'a list \ 'a list \ 'a list" where "merge_lists [] ys = ys" | "merge_lists xs [] = xs" | "merge_lists (x # xs) (y # ys) = (if less_eq x y then x # merge_lists xs (y # ys) else y # merge_lists (x # xs) ys)" lemma set_merge_lists [simp]: "set (merge_lists xs ys) = set xs \ set ys" by (induction xs ys rule: merge_lists.induct) auto lemma mset_merge_lists [simp]: "mset (merge_lists xs ys) = mset xs + mset ys" by (induction xs ys rule: merge_lists.induct) auto lemma sorted_merge_lists [simp, intro]: "sorted xs \ sorted ys \ sorted (merge_lists xs ys)" by (induction xs ys rule: merge_lists.induct) auto fun merge_sort :: "'a list \ 'a list" where "merge_sort xs = (if length xs \ 1 then xs else merge_lists (merge_sort (take (length xs div 2) xs)) (merge_sort (drop (length xs div 2) xs)))" lemmas [simp del] = merge_sort.simps lemma merge_sort_trivial [simp]: "length xs \ Suc 0 \ merge_sort xs = xs" by (subst merge_sort.simps) auto theorem mset_merge_sort [simp]: "mset (merge_sort xs) = mset xs" by (induction xs rule: merge_sort.induct) (subst merge_sort.simps, auto simp flip: mset_append) corollary set_merge_sort [simp]: "set (merge_sort xs) = set xs" by (rule mset_eq_setD) simp_all theorem sorted_merge_sort [simp, intro]: "sorted (merge_sort xs)" by (induction xs rule: merge_sort.induct) (subst merge_sort.simps, use sorted01 in auto) lemma inversion_number_between_code: "inversion_number_between xs ys = inversion_number_between_sorted (sort xs) (sort ys)" by (subst inversion_number_between_sorted_correct) - (simp_all add: sorted_sorted_wrt [symmetric] cong: inversion_number_between_cong_mset) + (simp_all add: cong: inversion_number_between_cong_mset) lemmas (in -) [code_unfold] = inversion_number_between_code subsection \Merge sort with inversion counting\ text \ Finally, we can put together all the components and define a variant of merge sort that counts the number of inversions in the original list: \ function sort_and_count_inversions :: "'a list \ 'a list \ nat" where "sort_and_count_inversions xs = (if length xs \ 1 then (xs, 0) else let (xs1, xs2) = split_list xs; (xs1', m) = sort_and_count_inversions xs1; (xs2', n) = sort_and_count_inversions xs2 in (merge_lists xs1' xs2', m + n + inversion_number_between_sorted xs1' xs2'))" by auto termination by (relation "measure length") (auto simp: split_list_def Let_def) lemmas [simp del] = sort_and_count_inversions.simps text \ The projection of this function to the first component is simply the standard merge sort algorithm that we defined and proved correct before. \ theorem fst_sort_and_count_inversions [simp]: "fst (sort_and_count_inversions xs) = merge_sort xs" by (induction xs rule: length_induct) (subst sort_and_count_inversions.simps, subst merge_sort.simps, simp_all add: split_list_def case_prod_unfold Let_def) text \ The projection to the second component is the inversion number. \ theorem snd_sort_and_count_inversions [simp]: "snd (sort_and_count_inversions xs) = inversion_number xs" proof (induction xs rule: length_induct) case (1 xs) show ?case proof (cases "length xs \ 1") case False have "xs = take (length xs div 2) xs @ drop (length xs div 2) xs" by simp also have "inversion_number \ = snd (sort_and_count_inversions xs)" by (subst inversion_number_append, subst sort_and_count_inversions.simps) (use False 1 in \auto simp: Let_def split_list_def case_prod_unfold inversion_number_between_sorted_correct - sorted_sorted_wrt [symmetric] cong: inversion_number_between_cong_mset\) finally show ?thesis .. qed (auto simp: sort_and_count_inversions.simps) qed lemmas (in -) [code_unfold] = snd_sort_and_count_inversions [symmetric] end end \ No newline at end of file diff --git a/thys/Polynomial_Factorization/Prime_Factorization.thy b/thys/Polynomial_Factorization/Prime_Factorization.thy --- a/thys/Polynomial_Factorization/Prime_Factorization.thy +++ b/thys/Polynomial_Factorization/Prime_Factorization.thy @@ -1,801 +1,801 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Prime Factorization\ text \This theory contains not-completely naive algorithms to test primality and to perform prime factorization. More precisely, it corresponds to prime factorization algorithm A in Knuth's textbook \cite{Knuth}.\ theory Prime_Factorization imports "HOL-Computational_Algebra.Primes" Missing_List Missing_Multiset begin subsection \Definitions\ definition primes_1000 :: "nat list" where "primes_1000 = [2, 3, 5, 7, 11, 13, 17, 19, 23, 29, 31, 37, 41, 43, 47, 53, 59, 61, 67, 71, 73, 79, 83, 89, 97, 101, 103, 107, 109, 113, 127, 131, 137, 139, 149, 151, 157, 163, 167, 173, 179, 181, 191, 193, 197, 199, 211, 223, 227, 229, 233, 239, 241, 251, 257, 263, 269, 271, 277, 281, 283, 293, 307, 311, 313, 317, 331, 337, 347, 349, 353, 359, 367, 373, 379, 383, 389, 397, 401, 409, 419, 421, 431, 433, 439, 443, 449, 457, 461, 463, 467, 479, 487, 491, 499, 503, 509, 521, 523, 541, 547, 557, 563, 569, 571, 577, 587, 593, 599, 601, 607, 613, 617, 619, 631, 641, 643, 647, 653, 659, 661, 673, 677, 683, 691, 701, 709, 719, 727, 733, 739, 743, 751, 757, 761, 769, 773, 787, 797, 809, 811, 821, 823, 827, 829, 839, 853, 857, 859, 863, 877, 881, 883, 887, 907, 911, 919, 929, 937, 941, 947, 953, 967, 971, 977, 983, 991, 997]" lemma primes_1000: "primes_1000 = filter prime [0..<1001]" by eval definition next_candidates :: "nat \ nat \ nat list" where "next_candidates n = (if n = 0 then (1001,primes_1000) else (n + 30, [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]))" definition "candidate_invariant n = (n = 0 \ n mod 30 = (11 :: nat))" partial_function (tailrec) remove_prime_factor :: "nat \ nat \ nat list \ nat \ nat list " where [code]: "remove_prime_factor p n ps = (case Divides.divmod_nat n p of (n',m) \ if m = 0 then remove_prime_factor p n' (p # ps) else (n,ps))" partial_function (tailrec) prime_factorization_nat_main :: "nat \ nat \ nat list \ nat list \ nat list" where [code]: "prime_factorization_nat_main n j is ps = (case is of [] \ (case next_candidates j of (j,is) \ prime_factorization_nat_main n j is ps) | (i # is) \ (case Divides.divmod_nat n i of (n',m) \ if m = 0 then case remove_prime_factor i n' (i # ps) of (n',ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' j is ps' else if i * i \ n then prime_factorization_nat_main n j is ps else (n # ps)))" partial_function (tailrec) prime_nat_main :: "nat \ nat \ nat list \ bool" where [code]: "prime_nat_main n j is = (case is of [] \ (case next_candidates j of (j,is) \ prime_nat_main n j is) | (i # is) \ (if i dvd n then i \ n else if i * i \ n then prime_nat_main n j is else True))" definition prime_nat :: "nat \ bool" where "prime_nat n \ if n < 2 then False else \ \TODO: integrate precomputed map\ case next_candidates 0 of (j,is) \ prime_nat_main n j is" definition prime_factorization_nat :: "nat \ nat list" where "prime_factorization_nat n \ rev (if n < 2 then [] else case next_candidates 0 of (j,is) \ prime_factorization_nat_main n j is [])" definition divisors_nat :: "nat \ nat list" where "divisors_nat n \ if n = 0 then [] else remdups_adj (sort (map prod_list (subseqs (prime_factorization_nat n))))" definition divisors_int_pos :: "int \ int list" where "divisors_int_pos x \ map int (divisors_nat (nat (abs x)))" definition divisors_int :: "int \ int list" where "divisors_int x \ let xs = divisors_int_pos x in xs @ (map uminus xs)" subsection \Proofs\ lemma remove_prime_factor: assumes res: "remove_prime_factor i n ps = (m,qs)" and i: "i > 1" and n: "n \ 0" shows "\ rs. qs = rs @ ps \ n = m * prod_list rs \ \ i dvd m \ set rs \ {i}" using res n proof (induct n arbitrary: ps rule: less_induct) case (less n ps) obtain n' mo where dm: "Divides.divmod_nat n i = (n',mo)" by force hence n': "n' = n div i" and mo: "mo = n mod i" by (auto simp: divmod_nat_def) from less(2)[unfolded remove_prime_factor.simps[of i n] dm] have res: "(if mo = 0 then remove_prime_factor i n' (i # ps) else (n, ps)) = (m, qs)" by auto from less(3) have n: "n \ 0" by auto with n' i have "n' < n" by auto note IH = less(1)[OF this] show ?case proof (cases "mo = 0") case True with mo n' have n: "n = n' * i" by auto with \n \ 0\ have n': "n' \ 0" by auto from True res have "remove_prime_factor i n' (i # ps) = (m,qs)" by auto from IH[OF this n'] obtain rs where "qs = rs @ i # ps" and "n' = m * prod_list rs \ \ i dvd m \ set rs \ {i}" by auto thus ?thesis by (intro exI[of _ "rs @ [i]"], unfold n, auto) next case False with mo have i_n: "\ i dvd n" by auto from False res have id: "m = n" "qs = ps" by auto show ?thesis unfolding id using i_n by auto qed qed lemma prime_sqrtI: assumes n: "n \ 2" and small: "\ j. 2 \ j \ j < i \ \ j dvd n" and i: "\ i * i \ n" shows "prime (n::nat)" unfolding prime_nat_iff proof (intro conjI impI allI) show "1 < n" using n by auto fix j assume jn: "j dvd n" from jn obtain k where njk: "n = j * k" unfolding dvd_def by auto with \1 < n\ have jn: "j \ n" by (metis dvd_imp_le jn neq0_conv not_less0) show "j = 1 \ j = n" proof (rule ccontr) assume "\ ?thesis" with njk n have j1: "j > 1 \ j \ n" by simp have "\ j k. 1 < j \ j \ k \ n = j * k" proof (cases "j \ k") case True thus ?thesis unfolding njk using j1 by blast next case False show ?thesis by (rule exI[of _ k], rule exI[of _ j], insert \1 < n\ j1 njk False, auto) (metis Suc_lessI mult_0_right neq0_conv) qed then obtain j k where j1: "1 < j" and jk: "j \ k" and njk: "n = j * k" by auto with small[of j] have ji: "j \ i" unfolding dvd_def by force from mult_mono[OF ji ji] have "i * i \ j * j" by auto with i have "j * j > n" by auto from this[unfolded njk] have "k < j" by auto with jk show False by auto qed qed lemma candidate_invariant_0: "candidate_invariant 0" unfolding candidate_invariant_def by auto lemma next_candidates: assumes res: "next_candidates n = (m,ps)" and n: "candidate_invariant n" shows "candidate_invariant m" "sorted ps" "{i. prime i \ n \ i \ i < m} \ set ps" "set ps \ {2..} \ {n.. []" "n < m" unfolding candidate_invariant_def proof - note res = res[unfolded next_candidates_def] note n = n[unfolded candidate_invariant_def] show "m = 0 \ m mod 30 = 11" using res n by (auto split: if_splits) - show "sorted ps" using res n by (auto split: if_splits simp: primes_1000_def sorted2_simps simp del: sorted.simps(2)) + show "sorted ps" using res n by (auto split: if_splits simp: primes_1000_def sorted2_simps simp del: sorted_wrt.simps(2)) show "set ps \ {2..} \ {n.. []" using res n by (auto split: if_splits simp: primes_1000_def) show "n < m" using res by (auto split: if_splits) show "{i. prime i \ n \ i \ i < m} \ set ps" proof (cases "n = 0") case True hence *: "m = 1001" "ps = primes_1000" using res by auto show ?thesis unfolding * True primes_1000 by auto next case False hence n: "n mod 30 = 11" and m: "m = n + 30" and ps: "ps = [n,n+2,n+6,n+8,n+12,n+18,n+20,n+26]" using res n by auto { fix i assume *: "prime i" "n \ i" "i < n + 30" "i \ set ps" from n * have i11: "i \ 11" by auto define j where "j = i - n" have i: "i = n + j" using \n \ i\ j_def by auto have "i mod 30 = (j + n) mod 30" using \n \ i\ unfolding j_def by simp also have "\ = (j mod 30 + n mod 30) mod 30" by (simp add: mod_simps) also have "\ = (j mod 30 + 11) mod 30" unfolding n by simp finally have i30: "i mod 30 = (j mod 30 + 11) mod 30" by simp have 2: "2 dvd (30 :: nat)" and 112: "11 mod (2 :: nat) = 1" by simp_all have "(j + 11) mod 2 = (j + 1) mod 2" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\j. j mod 2"] have 2: "i mod 2 = (j mod 2 + 1) mod 2" by (simp add: mod_simps mod_mod_cancel [OF 2]) have 3: "3 dvd (30 :: nat)" and 113: "11 mod (3 :: nat) = 2" by simp_all have "(j + 11) mod 3 = (j + 2) mod 3" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 3"] have 3: "i mod 3 = (j mod 3 + 2) mod 3" by (simp add: mod_simps mod_mod_cancel [OF 3]) have 5: "5 dvd (30 :: nat)" and 115: "11 mod (5 :: nat) = 1" by simp_all have "(j + 11) mod 5 = (j + 1) mod 5" by (rule mod_add_cong) simp_all with arg_cong [OF i30, of "\ j. j mod 5"] have 5: "i mod 5 = (j mod 5 + 1) mod 5" by (simp add: mod_simps mod_mod_cancel [OF 5]) from n *(2-)[unfolded ps i, simplified] have "j \ {1,3,5,7,9,11,13,15,17,19,21,23,25,27,29} \ j \ {4,10,16,22,28} \ j \ {14,24}" (is "j \ ?j2 \ j \ ?j3 \ j \ ?j5") by simp presburger moreover { assume "j \ ?j2" hence "j mod 2 = 1" by auto with 2 have "i mod 2 = 0" by auto with i11 have "2 dvd i" "i \ 2" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j3" hence "j mod 3 = 1" by auto with 3 have "i mod 3 = 0" by auto with i11 have "3 dvd i" "i \ 3" by auto with *(1) have False unfolding prime_nat_iff by auto } moreover { assume "j \ ?j5" hence "j mod 5 = 4" by auto with 5 have "i mod 5 = 0" by auto with i11 have "5 dvd i" "i \ 5" by auto with *(1) have False unfolding prime_nat_iff by auto } ultimately have False by blast } thus ?thesis unfolding m ps by auto qed qed lemma prime_test_iterate2: assumes small: "\ j. 2 \ j \ j < (i :: nat) \ \ j dvd n" and odd: "odd n" and n: "n \ 3" and i: "i \ 3" "odd i" and mod: "\ i dvd n" and j: "2 \ j" "j < i + 2" shows "\ j dvd n" proof assume dvd: "j dvd n" with small[OF j(1)] have "j \ i" by linarith with dvd mod have "j > i" by (cases "i = j", auto) with j have "j = Suc i" by simp with i have "even j" by auto with dvd j(1) have "2 dvd n" by (metis dvd_trans) with odd show False by auto qed lemma prime_divisor: assumes "j \ 2" and "j dvd n" shows "\ p :: nat. prime p \ p dvd j \ p dvd n" proof - let ?pf = "prime_factors j" from assms have "j > 0" by auto from prime_factorization_nat[OF this] have "j = (\p\?pf. p ^ multiplicity p j)" by auto with \j \ 2\ have "?pf \ {}" by auto then obtain p where p: "p \ ?pf" by auto hence pr: "prime p" by auto define rem where "rem = (\p\?pf - {p}. p ^ multiplicity p j)" from p have mult: "multiplicity p j \ 0" by (auto simp: prime_factors_multiplicity) have "finite ?pf" by simp have "j = (\p\?pf. p ^ multiplicity p j)" by fact also have "?pf = (insert p (?pf - {p}))" using p by auto also have "(\p\insert p (?pf - {p}). p ^ multiplicity p j) = p ^ multiplicity p j * rem" unfolding rem_def by (subst prod.insert, auto) also have "\ = p * (p ^ (multiplicity p j - 1) * rem)" using mult by (cases "multiplicity p j", auto) finally have pj: "p dvd j" unfolding dvd_def by blast with \j dvd n\ have "p dvd n" by (metis dvd_trans) with pj pr show ?thesis by blast qed lemma prime_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_nat_main n jj is \ res = prime n" proof (induct ni arbitrary: n i "is" jj res rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_nat_main n jjj iis" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) with res[unfolded simps] have res: "res = (if i' dvd n then n \ i' else if i' * i' \ n then prime_nat_main n jj iis else True)" by simp from 1(11) Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) have "set iis \ {i'..}" by(auto simp: Cons) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by(auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 n _ _ i'(3) sd_iis 1(10) iis] show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_nat_main n jj iis" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis by (rule IH[OF _ dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = True" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True have "i' \ 2" using i i' by auto from \i' dvd n\ obtain k where "n = i' * k" .. with n have "k \ 0" by (cases "k = 0", auto) with \n = i' * k\ have *: "i' < n \ i' = n" by auto with True res have "res \ i' = n" by auto also have "\ = prime n" using * proof assume "i' < n" with \i' \ 2\ \i' dvd n\ have "\ prime n" by (auto simp add: prime_nat_iff) with \i' < n\ show ?thesis by auto next assume "i' = n" with dvd n have "prime n" by (simp add: prime_nat_iff') with \i' = n\ show ?thesis by auto qed finally show ?thesis . qed qed qed lemma prime_factorization_nat_main: "ni = (n,i,is) \ i \ 2 \ n \ 2 \ (\ j. 2 \ j \ j < i \ \ (j dvd n)) \ (\ j. i \ j \ j < jj \ prime j \ j \ set is) \ i \ jj \ sorted is \ distinct is \ candidate_invariant jj \ set is \ {i.. res = prime_factorization_nat_main n jj is ps \ \ qs. res = qs @ ps \ Ball (set qs) prime \ n = prod_list qs" proof (induct ni arbitrary: n i "is" jj res ps rule: wf_induct[OF wf_measures[of "[\ (n,i,is). n - i, \ (n,i,is). if is = [] then 1 else 0]"]]) case (1 ni n i "is" jj res ps) note res = 1(12) from 1(3-4) have i: "i \ 2" and i2: "Suc i \ 2" and n: "n \ 2" by auto from 1(5) have dvd: "\ j. 2 \ j \ j < i \ \ j dvd n" . from 1(7) have ijj: "i \ jj" . note sort_dist = 1(8-9) have "is": "\ j. i \ j \ j < jj \ prime j \ j \ set is" by (rule 1(6)) note simps = prime_factorization_nat_main.simps[of n jj "is"] note IH = 1(1)[rule_format, unfolded 1(2), OF _ refl] show ?case proof (cases "is") case Nil obtain jjj iis where can: "next_candidates jj = (jjj,iis)" by force from res[unfolded simps, unfolded Nil can split] have res: "res = prime_factorization_nat_main n jjj iis ps" by auto from next_candidates[OF can 1(10)] have can: "sorted iis" "distinct iis" "candidate_invariant jjj" "{i. prime i \ jj \ i \ i < jjj} \ set iis" "set iis \ {2..} \ {jj.. []" "jj < jjj" by blast+ from can ijj have "i \ jjj" by auto note IH = IH[OF _ i n dvd _ this can(1-3) _ res] show ?thesis proof (rule IH, force simp: Nil can(6)) fix x assume ix: "i \ x" and xj: "x < jjj" and px: "prime x" from "is"[OF ix _ px] Nil have jx: "jj \ x" by force with can(4) xj px show "x \ set iis" by auto qed (insert can(5) ijj, auto) next case (Cons i' iis) obtain n' m where dm: "Divides.divmod_nat n i' = (n',m)" by force hence n': "n' = n div i'" and m: "m = n mod i'" by (auto simp: divmod_nat_def) have m: "(m = 0) = (i' dvd n)" unfolding m by auto from Cons res[unfolded simps] dm m n' have res: "res = ( if i' dvd n then case remove_prime_factor i' (n div i') (i' # ps) of (n', ps') \ if n' = 1 then ps' else prime_factorization_nat_main n' jj iis ps' else if i' * i' \ n then prime_factorization_nat_main n jj iis ps else n # ps)" by simp from 1(11) i Cons have iis: "set iis \ {i.. i'" "i' < jj" "Suc i' \ jj" "i' > 1" by auto from sort_dist have sd_iis: "sorted iis" "distinct iis" and "i' \ set iis" by(auto simp: Cons) from sort_dist(1) Cons have "set iis \ {i'..}" by(auto) with iis have "set iis \ {i'..i' \ set iis\ have iis: "set iis \ {Suc i'.. j" "j < i'" have "\ j dvd n" proof assume "j dvd n" from prime_divisor[OF j(1) this] obtain p where p: "prime p" "p dvd j" "p dvd n" by auto have pj: "p \ j" by (rule dvd_imp_le[OF p(2)], insert j, auto) have p2: "2 \ p" using p(1) by (rule prime_ge_2_nat) from dvd[OF p2] p(3) have pi: "p \ i" by force from pj j(2) i' "is"[OF pi _ p(1)] have "p \ set is" by auto with \sorted is\ have "i' \ p" by (auto simp: Cons) with pj j(2) show False by arith qed } note dvd = this from i' i have i'2: "2 \ Suc i'" by auto note IH = IH[OF _ i'2 _ _ _ i'(3) sd_iis 1(10) iis] { fix x assume "Suc i' \ x" "x < jj" "prime x" hence "i \ x" "x < jj" "prime x" using i' by auto from "is"[OF this] have "x \ set is" . with \Suc i' \ x\ have "x \ set iis" unfolding Cons by auto } note iis = this show ?thesis proof (cases "i' dvd n") case False note dvdi = this { fix j assume j: "2 \ j" "j < Suc i'" have "\ j dvd n" proof (cases "j = i'") case False with j have "j < i'" by auto from dvd[OF j(1) this] show ?thesis . qed (insert False, auto) } note dvds = this show ?thesis proof (cases "i' * i' \ n") case True note iin = this with res False have res: "res = prime_factorization_nat_main n jj iis ps" by auto from iin have i_n: "i' < n" using dvd dvdi n nat_neq_iff dvd_refl by blast show ?thesis by (rule IH[OF _ n dvds iis res], insert i_n i', auto) next case False with res dvdi have res: "res = n # ps" by auto have n: "prime n" by (rule prime_sqrtI[OF n dvd False]) thus ?thesis unfolding res by auto qed next case True note i_n = this obtain n'' qs where rp: "remove_prime_factor i' (n div i') (i' # ps) = (n'',qs)" by force with res True have res: "res = (if n'' = 1 then qs else prime_factorization_nat_main n'' jj iis qs)" by auto have pi: "prime i'" unfolding prime_nat_iff proof (intro conjI allI impI) show "1 < i'" using i' i by auto fix j assume ji: "j dvd i'" with i' i have j0: "j \ 0" by (cases "j = 0", auto) from ji i_n have jn: "j dvd n" by (metis dvd_trans) with dvd[of j] have j: "2 > j \ j \ i'" by linarith from ji \1 < i'\ have "j \ i'" unfolding dvd_def by (simp add: dvd_imp_le ji) with j j0 show "j = 1 \ j = i'" by linarith qed from True n' have id: "n = n' * i'" by auto from n id have "n' \ 0" by (cases "n = 0", auto) with id have "i' \ n" by auto from remove_prime_factor[OF rp[folded n'] \1 < i'\ \n' \ 0\] obtain rs where qs: "qs = rs @ i' # ps" and n': "n' = n'' * prod_list rs" and i_n'': "\ i' dvd n''" and rs: "set rs \ {i'}" by auto { fix j assume "j dvd n''" hence "j dvd n" unfolding id n' by auto } note dvd' = this show ?thesis proof (cases "n'' = 1") case False with res have res: "res = prime_factorization_nat_main n'' jj iis qs" by simp from i i' have "i' \ 2" by simp from False n' \n' \ 0\ have n2: "n'' \ 2" by (cases "n'' = 0"; auto) have lrs: "prod_list rs \ 0" using n' \n' \ 0\ by (cases "prod_list rs = 0", auto) with \i' \ 2\ have "prod_list rs * i' \ 2" by (cases "prod_list rs", auto) hence nn'': "n > n''" unfolding id n' using n2 by simp have "i' \ n" unfolding id n' using pi False by fastforce with \i' \ n\ i' have "n > i" by auto with nn'' i i' have less: "n - i > n'' - Suc i'" by simp { fix j assume 2: "2 \ j" and ji: "j < Suc i'" have "\ j dvd n''" proof (cases "j = i'") case False with ji have "j < i'" by auto from dvd' dvd[OF 2 this] show ?thesis by blast qed (insert i_n'', auto) } from IH[OF _ n2 this iis res] less obtain ss where res: "res = ss @ qs \ Ball (set ss) prime \ n'' = prod_list ss" by auto thus ?thesis unfolding id n' qs using pi rs by auto next case True with res have res: "res = qs" by auto show ?thesis unfolding id n' res qs True using rs \prime i'\ by (intro exI[of _ "rs @ [i']"], auto) qed qed qed qed lemma prime_nat[simp]: "prime_nat n = prime n" proof (cases "n < 2") case True thus ?thesis unfolding prime_nat_def prime_nat_iff by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) from n can have res: "prime_nat n = prime_nat_main n jj is" unfolding prime_nat_def by auto show ?thesis using prime_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub res] cann(4) by auto qed lemma prime_factorization_nat: fixes n :: nat defines "pf \ prime_factorization_nat n" shows "Ball (set pf) prime" and "n \ 0 \ prod_list pf = n" and "n = 0 \ pf = []" proof - note pf = pf_def[unfolded prime_factorization_nat_def] have "Ball (set pf) prime \ (n \ 0 \ prod_list pf = n) \ (n = 0 \ pf = [])" proof (cases "n < 2") case True thus ?thesis using pf by auto next case False hence n: "n \ 2" by auto obtain jj "is" where can: "next_candidates 0 = (jj,is)" by force from next_candidates[OF this candidate_invariant_0] have cann: "sorted is" "distinct is" "candidate_invariant jj" "{i. prime i \ 0 \ i \ i < jj} \ set is" "set is \ {2..} \ {0.. []" by auto from cann have sub: "set is \ {2..is \ []\ have jj: "jj \ 2" by (cases "is", auto) let ?pfm = "prime_factorization_nat_main n jj is []" from pf[unfolded can] False have res: "pf = rev ?pfm" by simp from prime_factorization_nat_main[OF refl le_refl n _ _ jj cann(1-3) sub refl, of Nil] cann(4) have "Ball (set ?pfm) prime" "n = prod_list ?pfm" by auto thus ?thesis unfolding res using n by auto qed thus "Ball (set pf) prime" "n \ 0 \ prod_list pf = n" "n = 0 \ pf = []" by auto qed lemma prod_mset_multiset_prime_factorization_nat [simp]: "(x::nat) \ 0 \ prod_mset (prime_factorization x) = x" by simp (* TODO Move *) lemma prime_factorization_unique'': fixes A :: "'a :: {factorial_semiring_multiplicative} multiset" assumes "\p. p \# A \ prime p" assumes "prod_mset A = normalize x" shows "prime_factorization x = A" proof - have "prod_mset A \ 0" by (auto dest: assms(1)) with assms(2) have "x \ 0" by simp hence "prod_mset (prime_factorization x) = prod_mset A" by (simp add: assms prod_mset_prime_factorization) with assms show ?thesis by (intro prime_factorization_unique') auto qed lemma multiset_prime_factorization_nat_correct: "prime_factorization n = mset (prime_factorization_nat n)" proof - note pf = prime_factorization_nat[of n] show ?thesis proof (cases "n = 0") case True thus ?thesis using pf(3) by simp next case False note pf = pf(1) pf(2)[OF False] show ?thesis proof (rule prime_factorization_unique'') show "prime p" if "p \# mset (prime_factorization_nat n)" for p using pf(1) that by simp let ?l = "\i\#prime_factorization n. i" let ?r = "\i\#mset (prime_factorization_nat n). i" show "prod_mset (mset (prime_factorization_nat n)) = normalize n" by (simp add: pf(2) prod_mset_prod_list) qed qed qed lemma multiset_prime_factorization_code[code_unfold]: "prime_factorization = (\n. mset (prime_factorization_nat n))" by (intro ext multiset_prime_factorization_nat_correct) lemma divisors_nat: "n \ 0 \ set (divisors_nat n) = {p. p dvd n}" "distinct (divisors_nat n)" "divisors_nat 0 = []" proof - show "distinct (divisors_nat n)" "divisors_nat 0 = []" unfolding divisors_nat_def by auto assume n: "n \ 0" from n have "n > 0" by auto { fix x have "(x dvd n) = (x \ 0 \ (\p. multiplicity p x \ multiplicity p n))" proof (cases "x = 0") case False with \n > 0\ show ?thesis by (auto simp: dvd_multiplicity_eq) next case True with n show ?thesis by auto qed } note dvd = this let ?dn = "set (divisors_nat n)" let ?mf = "\ (n :: nat). prime_factorization n" have "?dn = prod_list ` set (subseqs (prime_factorization_nat n))" unfolding divisors_nat_def using n by auto also have "\ = prod_mset ` mset ` set (subseqs (prime_factorization_nat n))" by (force simp: prod_mset_prod_list) also have "mset ` set (subseqs (prime_factorization_nat n)) = { ps. ps \# mset (prime_factorization_nat n)}" unfolding multiset_of_subseqs by simp also have "\ = { ps. ps \# ?mf n}" thm multiset_prime_factorization_code[symmetric] unfolding multiset_prime_factorization_nat_correct[symmetric] by auto also have "prod_mset ` \ = {p. p dvd n}" (is "?l = ?r") proof - { fix x assume "x dvd n" from this[unfolded dvd] have x: "x \ 0" by auto from \x dvd n\ \x \ 0\ \n \ 0\ have sub: "?mf x \# ?mf n" by (subst prime_factorization_subset_iff_dvd) auto have "prod_mset (?mf x) = x" using x by (simp add: prime_factorization_nat) hence "x \ ?l" using sub by force } moreover { fix x assume "x \ ?l" then obtain ps where x: "x = prod_mset ps" and sub: "ps \# ?mf n" by auto have "x dvd n" using prod_mset_subset_imp_dvd[OF sub] n x by simp } ultimately show ?thesis by blast qed finally show "set (divisors_nat n) = {p. p dvd n}" . qed lemma divisors_int_pos: "x \ 0 \ set (divisors_int_pos x) = {i. i dvd x \ i > 0}" "distinct (divisors_int_pos x)" "divisors_int_pos 0 = []" proof - show "divisors_int_pos 0 = []" by code_simp show "distinct (divisors_int_pos x)" unfolding divisors_int_pos_def using divisors_nat(2)[of "nat (abs x)"] by (simp add: distinct_map inj_on_def) assume x: "x \ 0" let ?x = "nat (abs x)" from x have xx: "?x \ 0" by auto from x have 0: "\ y. y dvd x \ y \ 0" by auto have id: "int ` {p. int p dvd x} = {i. i dvd x \ 0 < i}" (is "?l = ?r") proof - { fix y assume "y \ ?l" then obtain p where y: "y = int p" and dvd: "int p dvd x" by auto have "y \ ?r" unfolding y using dvd 0[OF dvd] by auto } moreover { fix y assume "y \ ?r" hence dvd: "y dvd x" and y0: "y > 0" by auto define n where "n = nat y" from y0 have y: "y = int n" unfolding n_def by auto with dvd have "y \ ?l" by auto } ultimately show ?thesis by blast qed from xx show "set (divisors_int_pos x) = {i. i dvd x \ i > 0}" by (simp add: divisors_int_pos_def divisors_nat id) qed lemma divisors_int: "x \ 0 \ set (divisors_int x) = {i. i dvd x}" "distinct (divisors_int x)" "divisors_int 0 = []" proof - show "divisors_int 0 = []" by code_simp show "distinct (divisors_int x)" proof (cases "x = 0") case True show ?thesis unfolding True by code_simp next case False from divisors_int_pos(1)[OF False] divisors_int_pos(2) show ?thesis unfolding divisors_int_def Let_def distinct_append distinct_map inj_on_def by auto qed assume x: "x \ 0" show "set (divisors_int x) = {i. i dvd x}" unfolding divisors_int_def Let_def set_append set_map divisors_int_pos(1)[OF x] using x by auto (metis (no_types, lifting) dvd_mult_div_cancel image_eqI linorder_neqE_linordered_idom mem_Collect_eq minus_dvd_iff minus_minus mult_zero_left neg_less_0_iff_less) qed definition divisors_fun :: "('a \ ('a :: {comm_monoid_mult,zero}) list) \ bool" where "divisors_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x }) \ (\ x. distinct (df x))" lemma divisors_funD: "divisors_fun df \ x \ 0 \ d dvd x \ d \ set (df x)" unfolding divisors_fun_def by auto definition divisors_pos_fun :: "('a \ ('a :: {comm_monoid_mult,zero,ord}) list) \ bool" where "divisors_pos_fun df \ (\ x. x \ 0 \ set (df x) = { d. d dvd x \ d > 0}) \ (\ x. distinct (df x))" lemma divisors_pos_funD: "divisors_pos_fun df \ x \ 0 \ d dvd x \ d > 0 \ d \ set (df x)" unfolding divisors_pos_fun_def by auto lemma divisors_fun_nat: "divisors_fun divisors_nat" unfolding divisors_fun_def using divisors_nat by auto lemma divisors_fun_int: "divisors_fun divisors_int" unfolding divisors_fun_def using divisors_int by auto lemma divisors_pos_fun_int: "divisors_pos_fun divisors_int_pos" unfolding divisors_pos_fun_def using divisors_int_pos by auto end diff --git a/thys/Polynomials/MPoly_Type_Class_Ordered.thy b/thys/Polynomials/MPoly_Type_Class_Ordered.thy --- a/thys/Polynomials/MPoly_Type_Class_Ordered.thy +++ b/thys/Polynomials/MPoly_Type_Class_Ordered.thy @@ -1,3676 +1,3676 @@ (* Author: Fabian Immler, Alexander Maletzky *) section \Type-Class-Multivariate Polynomials in Ordered Terms\ theory MPoly_Type_Class_Ordered imports MPoly_Type_Class begin class the_min = linorder + fixes the_min::'a assumes the_min_min: "the_min \ x" text \Type class @{class the_min} guarantees that a least element exists. Instances of @{class the_min} should provide @{emph \computable\} definitions of that element.\ instantiation nat :: the_min begin definition "the_min_nat = (0::nat)" instance by (standard, simp add: the_min_nat_def) end instantiation unit :: the_min begin definition "the_min_unit = ()" instance by (standard, simp add: the_min_unit_def) end locale ordered_term = term_powerprod pair_of_term term_of_pair + ordered_powerprod ord ord_strict + ord_term_lin: linorder ord_term ord_term_strict for pair_of_term::"'t \ ('a::comm_powerprod \ 'k::{the_min,wellorder})" and term_of_pair::"('a \ 'k) \ 't" and ord::"'a \ 'a \ bool" (infixl "\" 50) and ord_strict (infixl "\" 50) and ord_term::"'t \ 't \ bool" (infixl "\\<^sub>t" 50) and ord_term_strict::"'t \ 't \ bool" (infixl "\\<^sub>t" 50) + assumes splus_mono: "v \\<^sub>t w \ t \ v \\<^sub>t t \ w" assumes ord_termI: "pp_of_term v \ pp_of_term w \ component_of_term v \ component_of_term w \ v \\<^sub>t w" begin abbreviation ord_term_conv (infixl "\\<^sub>t" 50) where "ord_term_conv \ (\\<^sub>t)\\" abbreviation ord_term_strict_conv (infixl "\\<^sub>t" 50) where "ord_term_strict_conv \ (\\<^sub>t)\\" text \The definition of @{locale ordered_term} only covers TOP and POT orderings. These two types of orderings are the only interesting ones.\ definition "min_term \ term_of_pair (0, the_min)" lemma min_term_min: "min_term \\<^sub>t v" proof (rule ord_termI) show "pp_of_term min_term \ pp_of_term v" by (simp add: min_term_def zero_min term_simps) next show "component_of_term min_term \ component_of_term v" by (simp add: min_term_def the_min_min term_simps) qed lemma splus_mono_strict: assumes "v \\<^sub>t w" shows "t \ v \\<^sub>t t \ w" proof - from assms have "v \\<^sub>t w" and "v \ w" by simp_all from this(1) have "t \ v \\<^sub>t t \ w" by (rule splus_mono) moreover from \v \ w\ have "t \ v \ t \ w" by (simp add: term_simps) ultimately show ?thesis using ord_term_lin.antisym_conv1 by blast qed lemma splus_mono_left: assumes "s \ t" shows "s \ v \\<^sub>t t \ v" proof (rule ord_termI, simp_all add: term_simps) from assms show "s + pp_of_term v \ t + pp_of_term v" by (rule plus_monotone) qed lemma splus_mono_strict_left: assumes "s \ t" shows "s \ v \\<^sub>t t \ v" proof - from assms have "s \ t" and "s \ t" by simp_all from this(1) have "s \ v \\<^sub>t t \ v" by (rule splus_mono_left) moreover from \s \ t\ have "s \ v \ t \ v" by (simp add: term_simps) ultimately show ?thesis using ord_term_lin.antisym_conv1 by blast qed lemma ord_term_canc: assumes "t \ v \\<^sub>t t \ w" shows "v \\<^sub>t w" proof (rule ccontr) assume "\ v \\<^sub>t w" hence "w \\<^sub>t v" by simp hence "t \ w \\<^sub>t t \ v" by (rule splus_mono_strict) with assms show False by simp qed lemma ord_term_strict_canc: assumes "t \ v \\<^sub>t t \ w" shows "v \\<^sub>t w" proof (rule ccontr) assume "\ v \\<^sub>t w" hence "w \\<^sub>t v" by simp hence "t \ w \\<^sub>t t \ v" by (rule splus_mono) with assms show False by simp qed lemma ord_term_canc_left: assumes "t \ v \\<^sub>t s \ v" shows "t \ s" proof (rule ccontr) assume "\ t \ s" hence "s \ t" by simp hence "s \ v \\<^sub>t t \ v" by (rule splus_mono_strict_left) with assms show False by simp qed lemma ord_term_strict_canc_left: assumes "t \ v \\<^sub>t s \ v" shows "t \ s" proof (rule ccontr) assume "\ t \ s" hence "s \ t" by simp hence "s \ v \\<^sub>t t \ v" by (rule splus_mono_left) with assms show False by simp qed lemma ord_adds_term: assumes "u adds\<^sub>t v" shows "u \\<^sub>t v" proof - from assms have *: "component_of_term u \ component_of_term v" and "pp_of_term u adds pp_of_term v" by (simp_all add: adds_term_def) from this(2) have "pp_of_term u \ pp_of_term v" by (rule ord_adds) from this * show ?thesis by (rule ord_termI) qed end subsection \Interpretations\ context ordered_powerprod begin subsubsection \Unit\ sublocale punit: ordered_term to_pair_unit fst "(\)" "(\)" "(\)" "(\)" apply standard subgoal by (simp, fact plus_monotone_left) subgoal by (simp only: punit_pp_of_term punit_component_of_term) done lemma punit_min_term [simp]: "punit.min_term = 0" by (simp add: punit.min_term_def) end subsection \Definitions\ context ordered_term begin definition higher :: "('t \\<^sub>0 'b) \ 't \ ('t \\<^sub>0 'b::zero)" where "higher p t = except p {s. s \\<^sub>t t}" definition lower :: "('t \\<^sub>0 'b) \ 't \ ('t \\<^sub>0 'b::zero)" where "lower p t = except p {s. t \\<^sub>t s}" definition lt :: "('t \\<^sub>0 'b::zero) \ 't" where "lt p = (if p = 0 then min_term else ord_term_lin.Max (keys p))" abbreviation "lp p \ pp_of_term (lt p)" definition lc :: "('t \\<^sub>0 'b::zero) \ 'b" where "lc p = lookup p (lt p)" definition tt :: "('t \\<^sub>0 'b::zero) \ 't" where "tt p = (if p = 0 then min_term else ord_term_lin.Min (keys p))" abbreviation "tp p \ pp_of_term (tt p)" definition tc :: "('t \\<^sub>0 'b::zero) \ 'b" where "tc p \ lookup p (tt p)" definition tail :: "('t \\<^sub>0 'b) \ ('t \\<^sub>0 'b::zero)" where "tail p \ lower p (lt p)" subsection \Leading Term and Leading Coefficient: @{const lt} and @{const lc}\ lemma lt_zero [simp]: "lt 0 = min_term" by (simp add: lt_def) lemma lc_zero [simp]: "lc 0 = 0" by (simp add: lc_def) lemma lt_uminus [simp]: "lt (- p) = lt p" by (simp add: lt_def keys_uminus) lemma lc_uminus [simp]: "lc (- p) = - lc p" by (simp add: lc_def) lemma lt_alt: assumes "p \ 0" shows "lt p = ord_term_lin.Max (keys p)" using assms unfolding lt_def by simp lemma lt_max: assumes "lookup p v \ 0" shows "v \\<^sub>t lt p" proof - from assms have t_in: "v \ keys p" by (simp add: in_keys_iff) hence "keys p \ {}" by auto hence "p \ 0" using keys_zero by blast from lt_alt[OF this] ord_term_lin.Max_ge[OF finite_keys t_in] show ?thesis by simp qed lemma lt_eqI: assumes "lookup p v \ 0" and "\u. lookup p u \ 0 \ u \\<^sub>t v" shows "lt p = v" proof - from assms(1) have "v \ keys p" by (simp add: in_keys_iff) hence "keys p \ {}" by auto hence "p \ 0" using keys_zero by blast have "u \\<^sub>t v" if "u \ keys p" for u proof - from that have "lookup p u \ 0" by (simp add: in_keys_iff) thus "u \\<^sub>t v" by (rule assms(2)) qed from lt_alt[OF \p \ 0\] ord_term_lin.Max_eqI[OF finite_keys this \v \ keys p\] show ?thesis by simp qed lemma lt_less: assumes "p \ 0" and "\u. v \\<^sub>t u \ lookup p u = 0" shows "lt p \\<^sub>t v" proof - from \p \ 0\ have "keys p \ {}" by simp have "\u\keys p. u \\<^sub>t v" proof fix u::'t assume "u \ keys p" hence "lookup p u \ 0" by (simp add: in_keys_iff) hence "\ v \\<^sub>t u" using assms(2)[of u] by auto thus "u \\<^sub>t v" by simp qed with lt_alt[OF assms(1)] ord_term_lin.Max_less_iff[OF finite_keys \keys p \ {}\] show ?thesis by simp qed lemma lt_le: assumes "\u. v \\<^sub>t u \ lookup p u = 0" shows "lt p \\<^sub>t v" proof (cases "p = 0") case True show ?thesis by (simp add: True min_term_min) next case False hence "keys p \ {}" by simp have "\u\keys p. u \\<^sub>t v" proof fix u::'t assume "u \ keys p" hence "lookup p u \ 0" unfolding keys_def by simp hence "\ v \\<^sub>t u" using assms[of u] by auto thus "u \\<^sub>t v" by simp qed with lt_alt[OF False] ord_term_lin.Max_le_iff[OF finite_keys[of p] \keys p \ {}\] show ?thesis by simp qed lemma lt_gr: assumes "lookup p s \ 0" and "t \\<^sub>t s" shows "t \\<^sub>t lt p" using assms lt_max ord_term_lin.order.strict_trans2 by blast lemma lc_not_0: assumes "p \ 0" shows "lc p \ 0" proof - from keys_zero assms have "keys p \ {}" by auto from lt_alt[OF assms] ord_term_lin.Max_in[OF finite_keys this] show ?thesis by (simp add: in_keys_iff lc_def) qed lemma lc_eq_zero_iff: "lc p = 0 \ p = 0" using lc_not_0 lc_zero by blast lemma lt_in_keys: assumes "p \ 0" shows "lt p \ (keys p)" by (metis assms in_keys_iff lc_def lc_not_0) lemma lt_monomial: "lt (monomial c t) = t" if "c \ 0" using that by (auto simp add: lt_def dest: monomial_0D) lemma lc_monomial [simp]: "lc (monomial c t) = c" proof (cases "c = 0") case True thus ?thesis by simp next case False thus ?thesis by (simp add: lc_def lt_monomial) qed lemma lt_le_iff: "lt p \\<^sub>t v \ (\u. v \\<^sub>t u \ lookup p u = 0)" (is "?L \ ?R") proof assume ?L show ?R proof (intro allI impI) fix u note \lt p \\<^sub>t v\ also assume "v \\<^sub>t u" finally have "lt p \\<^sub>t u" . hence "\ u \\<^sub>t lt p" by simp with lt_max[of p u] show "lookup p u = 0" by blast qed next assume ?R thus ?L using lt_le by auto qed lemma lt_plus_eqI: assumes "lt p \\<^sub>t lt q" shows "lt (p + q) = lt q" proof (cases "q = 0") case True with assms have "lt p \\<^sub>t min_term" by (simp add: lt_def) with min_term_min[of "lt p"] show ?thesis by simp next case False show ?thesis proof (intro lt_eqI) from lt_gr[of p "lt q" "lt p"] assms have "lookup p (lt q) = 0" by blast with lookup_add[of p q "lt q"] lc_not_0[OF False] show "lookup (p + q) (lt q) \ 0" unfolding lc_def by simp next fix u assume "lookup (p + q) u \ 0" show "u \\<^sub>t lt q" proof (rule ccontr) assume "\ u \\<^sub>t lt q" hence qs: "lt q \\<^sub>t u" by simp with assms have "lt p \\<^sub>t u" by simp with lt_gr[of p u "lt p"] have "lookup p u = 0" by blast moreover from qs lt_gr[of q u "lt q"] have "lookup q u = 0" by blast ultimately show False using \lookup (p + q) u \ 0\ lookup_add[of p q u] by auto qed qed qed lemma lt_plus_eqI_2: assumes "lt q \\<^sub>t lt p" shows "lt (p + q) = lt p" proof (cases "p = 0") case True with assms have "lt q \\<^sub>t min_term" by (simp add: lt_def) with min_term_min[of "lt q"] show ?thesis by simp next case False show ?thesis proof (intro lt_eqI) from lt_gr[of q "lt p" "lt q"] assms have "lookup q (lt p) = 0" by blast with lookup_add[of p q "lt p"] lc_not_0[OF False] show "lookup (p + q) (lt p) \ 0" unfolding lc_def by simp next fix u assume "lookup (p + q) u \ 0" show "u \\<^sub>t lt p" proof (rule ccontr) assume "\ u \\<^sub>t lt p" hence ps: "lt p \\<^sub>t u" by simp with assms have "lt q \\<^sub>t u" by simp with lt_gr[of q u "lt q"] have "lookup q u = 0" by blast also from ps lt_gr[of p u "lt p"] have "lookup p u = 0" by blast ultimately show False using \lookup (p + q) u \ 0\ lookup_add[of p q u] by auto qed qed qed lemma lt_plus_eqI_3: assumes "lt q = lt p" and "lc p + lc q \ 0" shows "lt (p + q) = lt (p::'t \\<^sub>0 'b::monoid_add)" proof (rule lt_eqI) from assms(2) show "lookup (p + q) (lt p) \ 0" by (simp add: lookup_add lc_def assms(1)) next fix u assume "lookup (p + q) u \ 0" hence "lookup p u + lookup q u \ 0" by (simp add: lookup_add) hence "lookup p u \ 0 \ lookup q u \ 0" by auto thus "u \\<^sub>t lt p" proof assume "lookup p u \ 0" thus ?thesis by (rule lt_max) next assume "lookup q u \ 0" hence "u \\<^sub>t lt q" by (rule lt_max) thus ?thesis by (simp only: assms(1)) qed qed lemma lt_plus_lessE: assumes "lt p \\<^sub>t lt (p + q)" shows "lt p \\<^sub>t lt q" proof (rule ccontr) assume "\ lt p \\<^sub>t lt q" hence "lt p = lt q \ lt q \\<^sub>t lt p" by auto thus False proof assume lt_eq: "lt p = lt q" have "lt (p + q) \\<^sub>t lt p" proof (rule lt_le) fix u assume "lt p \\<^sub>t u" with lt_gr[of p u "lt p"] have "lookup p u = 0" by blast from \lt p \\<^sub>t u\ have "lt q \\<^sub>t u" using lt_eq by simp with lt_gr[of q u "lt q"] have "lookup q u = 0" by blast with \lookup p u = 0\ show "lookup (p + q) u = 0" by (simp add: lookup_add) qed with assms show False by simp next assume "lt q \\<^sub>t lt p" from lt_plus_eqI_2[OF this] assms show False by simp qed qed lemma lt_plus_lessE_2: assumes "lt q \\<^sub>t lt (p + q)" shows "lt q \\<^sub>t lt p" proof (rule ccontr) assume "\ lt q \\<^sub>t lt p" hence "lt q = lt p \ lt p \\<^sub>t lt q" by auto thus False proof assume lt_eq: "lt q = lt p" have "lt (p + q) \\<^sub>t lt q" proof (rule lt_le) fix u assume "lt q \\<^sub>t u" with lt_gr[of q u "lt q"] have "lookup q u = 0" by blast from \lt q \\<^sub>t u\ have "lt p \\<^sub>t u" using lt_eq by simp with lt_gr[of p u "lt p"] have "lookup p u = 0" by blast with \lookup q u = 0\ show "lookup (p + q) u = 0" by (simp add: lookup_add) qed with assms show False by simp next assume "lt p \\<^sub>t lt q" from lt_plus_eqI[OF this] assms show False by simp qed qed lemma lt_plus_lessI': fixes p q :: "'t \\<^sub>0 'b::monoid_add" assumes "p + q \ 0" and lt_eq: "lt q = lt p" and lc_eq: "lc p + lc q = 0" shows "lt (p + q) \\<^sub>t lt p" proof (rule ccontr) assume "\ lt (p + q) \\<^sub>t lt p" hence "lt (p + q) = lt p \ lt p \\<^sub>t lt (p + q)" by auto thus False proof assume "lt (p + q) = lt p" have "lookup (p + q) (lt p) = (lookup p (lt p)) + (lookup q (lt q))" unfolding lt_eq lookup_add .. also have "... = lc p + lc q" unfolding lc_def .. also have "... = 0" unfolding lc_eq by simp finally have "lookup (p + q) (lt p) = 0" . hence "lt (p + q) \ lt p" using lc_not_0[OF \p + q \ 0\] unfolding lc_def by auto with \lt (p + q) = lt p\ show False by simp next assume "lt p \\<^sub>t lt (p + q)" have "lt p \\<^sub>t lt q" by (rule lt_plus_lessE, fact+) hence "lt p \ lt q" by simp with lt_eq show False by simp qed qed corollary lt_plus_lessI: fixes p q :: "'t \\<^sub>0 'b::group_add" assumes "p + q \ 0" and "lt q = lt p" and "lc q = - lc p" shows "lt (p + q) \\<^sub>t lt p" using assms(1, 2) proof (rule lt_plus_lessI') from assms(3) show "lc p + lc q = 0" by simp qed lemma lt_plus_distinct_eq_max: assumes "lt p \ lt q" shows "lt (p + q) = ord_term_lin.max (lt p) (lt q)" proof (rule ord_term_lin.linorder_cases) assume a: "lt p \\<^sub>t lt q" hence "lt (p + q) = lt q" by (rule lt_plus_eqI) also from a have "... = ord_term_lin.max (lt p) (lt q)" by (simp add: ord_term_lin.max.absorb2) finally show ?thesis . next assume a: "lt q \\<^sub>t lt p" hence "lt (p + q) = lt p" by (rule lt_plus_eqI_2) also from a have "... = ord_term_lin.max (lt p) (lt q)" by (simp add: ord_term_lin.max.absorb1) finally show ?thesis . next assume "lt p = lt q" with assms show ?thesis .. qed lemma lt_plus_le_max: "lt (p + q) \\<^sub>t ord_term_lin.max (lt p) (lt q)" proof (cases "lt p = lt q") case True show ?thesis proof (rule lt_le) fix u assume "ord_term_lin.max (lt p) (lt q) \\<^sub>t u" hence "lt p \\<^sub>t u" and "lt q \\<^sub>t u" by simp_all hence "lookup p u = 0" and "lookup q u = 0" using lt_max ord_term_lin.leD by blast+ thus "lookup (p + q) u = 0" by (simp add: lookup_add) qed next case False hence "lt (p + q) = ord_term_lin.max (lt p) (lt q)" by (rule lt_plus_distinct_eq_max) thus ?thesis by simp qed lemma lt_minus_eqI: "lt p \\<^sub>t lt q \ lt (p - q) = lt q" for p q :: "'t \\<^sub>0 'b::ab_group_add" by (metis lt_plus_eqI_2 lt_uminus uminus_add_conv_diff) lemma lt_minus_eqI_2: "lt q \\<^sub>t lt p \ lt (p - q) = lt p" for p q :: "'t \\<^sub>0 'b::ab_group_add" by (metis lt_minus_eqI lt_uminus minus_diff_eq) lemma lt_minus_eqI_3: assumes "lt q = lt p" and "lc q \ lc p" shows "lt (p - q) = lt (p::'t \\<^sub>0 'b::ab_group_add)" proof (rule lt_eqI) from assms(2) show "lookup (p - q) (lt p) \ 0" by (simp add: lookup_minus lc_def assms(1)) next fix u assume "lookup (p - q) u \ 0" hence "lookup p u \ lookup q u" by (simp add: lookup_minus) hence "lookup p u \ 0 \ lookup q u \ 0" by auto thus "u \\<^sub>t lt p" proof assume "lookup p u \ 0" thus ?thesis by (rule lt_max) next assume "lookup q u \ 0" hence "u \\<^sub>t lt q" by (rule lt_max) thus ?thesis by (simp only: assms(1)) qed qed lemma lt_minus_distinct_eq_max: assumes "lt p \ lt (q::'t \\<^sub>0 'b::ab_group_add)" shows "lt (p - q) = ord_term_lin.max (lt p) (lt q)" proof (rule ord_term_lin.linorder_cases) assume a: "lt p \\<^sub>t lt q" hence "lt (p - q) = lt q" by (rule lt_minus_eqI) also from a have "... = ord_term_lin.max (lt p) (lt q)" by (simp add: ord_term_lin.max.absorb2) finally show ?thesis . next assume a: "lt q \\<^sub>t lt p" hence "lt (p - q) = lt p" by (rule lt_minus_eqI_2) also from a have "... = ord_term_lin.max (lt p) (lt q)" by (simp add: ord_term_lin.max.absorb1) finally show ?thesis . next assume "lt p = lt q" with assms show ?thesis .. qed lemma lt_minus_lessE: "lt p \\<^sub>t lt (p - q) \ lt p \\<^sub>t lt q" for p q :: "'t \\<^sub>0 'b::ab_group_add" using lt_minus_eqI_2 by fastforce lemma lt_minus_lessE_2: "lt q \\<^sub>t lt (p - q) \ lt q \\<^sub>t lt p" for p q :: "'t \\<^sub>0 'b::ab_group_add" using lt_plus_eqI_2 by fastforce lemma lt_minus_lessI: "p - q \ 0 \ lt q = lt p \ lc q = lc p \ lt (p - q) \\<^sub>t lt p" for p q :: "'t \\<^sub>0 'b::ab_group_add" by (metis (no_types, hide_lams) diff_diff_eq2 diff_self group_eq_aux lc_def lc_not_0 lookup_minus lt_minus_eqI ord_term_lin.antisym_conv3) lemma lt_max_keys: assumes "v \ keys p" shows "v \\<^sub>t lt p" proof (rule lt_max) from assms show "lookup p v \ 0" by (simp add: in_keys_iff) qed lemma lt_eqI_keys: assumes "v \ keys p" and a2: "\u. u \ keys p \ u \\<^sub>t v" shows "lt p = v" by (rule lt_eqI, simp_all only: in_keys_iff[symmetric], fact+) lemma lt_gr_keys: assumes "u \ keys p" and "v \\<^sub>t u" shows "v \\<^sub>t lt p" proof (rule lt_gr) from assms(1) show "lookup p u \ 0" by (simp add: in_keys_iff) qed fact lemma lt_plus_eq_maxI: assumes "lt p = lt q \ lc p + lc q \ 0" shows "lt (p + q) = ord_term_lin.max (lt p) (lt q)" proof (cases "lt p = lt q") case True show ?thesis proof (rule lt_eqI_keys) from True have "lc p + lc q \ 0" by (rule assms) thus "ord_term_lin.max (lt p) (lt q) \ keys (p + q)" by (simp add: in_keys_iff lc_def lookup_add True) next fix u assume "u \ keys (p + q)" hence "u \\<^sub>t lt (p + q)" by (rule lt_max_keys) also have "... \\<^sub>t ord_term_lin.max (lt p) (lt q)" by (fact lt_plus_le_max) finally show "u \\<^sub>t ord_term_lin.max (lt p) (lt q)" . qed next case False thus ?thesis by (rule lt_plus_distinct_eq_max) qed lemma lt_monom_mult: assumes "c \ (0::'b::semiring_no_zero_divisors)" and "p \ 0" shows "lt (monom_mult c t p) = t \ lt p" proof (intro lt_eqI) from assms(1) show "lookup (monom_mult c t p) (t \ lt p) \ 0" proof (simp add: lookup_monom_mult_plus) show "lookup p (lt p) \ 0" using assms(2) lt_in_keys by auto qed next fix u::'t assume "lookup (monom_mult c t p) u \ 0" hence "u \ keys (monom_mult c t p)" by (simp add: in_keys_iff) also have "... \ (\) t ` keys p" by (fact keys_monom_mult_subset) finally obtain v where "v \ keys p" and "u = t \ v" .. show "u \\<^sub>t t \ lt p" unfolding \u = t \ v\ proof (rule splus_mono) from \v \ keys p\ show "v \\<^sub>t lt p" by (rule lt_max_keys) qed qed lemma lt_monom_mult_zero: assumes "c \ (0::'b::semiring_no_zero_divisors)" shows "lt (monom_mult c 0 p) = lt p" proof (cases "p = 0") case True show ?thesis by (simp add: True) next case False with assms show ?thesis by (simp add: lt_monom_mult term_simps) qed corollary lt_map_scale: "c \ (0::'b::semiring_no_zero_divisors) \ lt (c \ p) = lt p" by (simp add: map_scale_eq_monom_mult lt_monom_mult_zero) lemma lc_monom_mult [simp]: "lc (monom_mult c t p) = (c::'b::semiring_no_zero_divisors) * lc p" proof (cases "c = 0") case True thus ?thesis by simp next case False show ?thesis proof (cases "p = 0") case True thus ?thesis by simp next case False with \c \ 0\ show ?thesis by (simp add: lc_def lt_monom_mult lookup_monom_mult_plus) qed qed corollary lc_map_scale [simp]: "lc (c \ p) = (c::'b::semiring_no_zero_divisors) * lc p" by (simp add: map_scale_eq_monom_mult) lemma (in ordered_term) lt_mult_scalar_monomial_right: assumes "c \ (0::'b::semiring_no_zero_divisors)" and "p \ 0" shows "lt (p \ monomial c v) = punit.lt p \ v" proof (intro lt_eqI) from assms(1) show "lookup (p \ monomial c v) (punit.lt p \ v) \ 0" proof (simp add: lookup_mult_scalar_monomial_right_plus) from assms(2) show "lookup p (punit.lt p) \ 0" using in_keys_iff punit.lt_in_keys by fastforce qed next fix u::'t assume "lookup (p \ monomial c v) u \ 0" hence "u \ keys (p \ monomial c v)" by (simp add: in_keys_iff) also have "... \ (\t. t \ v) ` keys p" by (fact keys_mult_scalar_monomial_right_subset) finally obtain t where "t \ keys p" and "u = t \ v" .. show "u \\<^sub>t punit.lt p \ v" unfolding \u = t \ v\ proof (rule splus_mono_left) from \t \ keys p\ show "t \ punit.lt p" by (rule punit.lt_max_keys) qed qed lemma lc_mult_scalar_monomial_right: "lc (p \ monomial c v) = punit.lc p * (c::'b::semiring_no_zero_divisors)" proof (cases "c = 0") case True thus ?thesis by simp next case False show ?thesis proof (cases "p = 0") case True thus ?thesis by simp next case False with \c \ 0\ show ?thesis by (simp add: punit.lc_def lc_def lt_mult_scalar_monomial_right lookup_mult_scalar_monomial_right_plus) qed qed lemma lookup_monom_mult_eq_zero: assumes "s \ lt p \\<^sub>t v" shows "lookup (monom_mult (c::'b::semiring_no_zero_divisors) s p) v = 0" by (metis assms aux lt_gr lt_monom_mult monom_mult_zero_left monom_mult_zero_right ord_term_lin.order.strict_implies_not_eq) lemma in_keys_monom_mult_le: assumes "v \ keys (monom_mult c t p)" shows "v \\<^sub>t t \ lt p" proof - from keys_monom_mult_subset assms have "v \ (\) t ` (keys p)" .. then obtain u where "u \ keys p" and "v = t \ u" .. from \u \ keys p\ have "u \\<^sub>t lt p" by (rule lt_max_keys) thus "v \\<^sub>t t \ lt p" unfolding \v = t \ u\ by (rule splus_mono) qed lemma lt_monom_mult_le: "lt (monom_mult c t p) \\<^sub>t t \ lt p" by (metis aux in_keys_monom_mult_le lt_in_keys lt_le_iff) lemma monom_mult_inj_2: assumes "monom_mult c t1 p = monom_mult c t2 p" and "c \ 0" and "(p::'t \\<^sub>0 'b::semiring_no_zero_divisors) \ 0" shows "t1 = t2" proof - from assms(1) have "lt (monom_mult c t1 p) = lt (monom_mult c t2 p)" by simp with \c \ 0\ \p \ 0\ have "t1 \ lt p = t2 \ lt p" by (simp add: lt_monom_mult) thus ?thesis by (simp add: term_simps) qed subsection \Trailing Term and Trailing Coefficient: @{const tt} and @{const tc}\ lemma tt_zero [simp]: "tt 0 = min_term" by (simp add: tt_def) lemma tc_zero [simp]: "tc 0 = 0" by (simp add: tc_def) lemma tt_alt: assumes "p \ 0" shows "tt p = ord_term_lin.Min (keys p)" using assms unfolding tt_def by simp lemma tt_min_keys: assumes "v \ keys p" shows "tt p \\<^sub>t v" proof - from assms have "keys p \ {}" by auto hence "p \ 0" by simp from tt_alt[OF this] ord_term_lin.Min_le[OF finite_keys assms] show ?thesis by simp qed lemma tt_min: assumes "lookup p v \ 0" shows "tt p \\<^sub>t v" proof - from assms have "v \ keys p" unfolding keys_def by simp thus ?thesis by (rule tt_min_keys) qed lemma tt_in_keys: assumes "p \ 0" shows "tt p \ keys p" unfolding tt_alt[OF assms] by (rule ord_term_lin.Min_in, fact finite_keys, simp add: assms) lemma tt_eqI: assumes "v \ keys p" and "\u. u \ keys p \ v \\<^sub>t u" shows "tt p = v" proof - from assms(1) have "keys p \ {}" by auto hence "p \ 0" by simp from assms(1) have "tt p \\<^sub>t v" by (rule tt_min_keys) moreover have "v \\<^sub>t tt p" by (rule assms(2), rule tt_in_keys, fact \p \ 0\) ultimately show ?thesis by simp qed lemma tt_gr: assumes "\u. u \ keys p \ v \\<^sub>t u" and "p \ 0" shows "v \\<^sub>t tt p" proof - from \p \ 0\ have "keys p \ {}" by simp show ?thesis by (rule assms(1), rule tt_in_keys, fact \p \ 0\) qed lemma tt_less: assumes "u \ keys p" and "u \\<^sub>t v" shows "tt p \\<^sub>t v" proof - from \u \ keys p\ have "tt p \\<^sub>t u" by (rule tt_min_keys) also have "... \\<^sub>t v" by fact finally show "tt p \\<^sub>t v" . qed lemma tt_ge: assumes "\u. u \\<^sub>t v \ lookup p u = 0" and "p \ 0" shows "v \\<^sub>t tt p" proof - from \p \ 0\ have "keys p \ {}" by simp have "\u\keys p. v \\<^sub>t u" proof fix u::'t assume "u \ keys p" hence "lookup p u \ 0" unfolding keys_def by simp hence "\ u \\<^sub>t v" using assms(1)[of u] by auto thus "v \\<^sub>t u" by simp qed with tt_alt[OF \p \ 0\] ord_term_lin.Min_ge_iff[OF finite_keys[of p] \keys p \ {}\] show ?thesis by simp qed lemma tt_ge_keys: assumes "\u. u \ keys p \ v \\<^sub>t u" and "p \ 0" shows "v \\<^sub>t tt p" by (rule assms(1), rule tt_in_keys, fact) lemma tt_ge_iff: "v \\<^sub>t tt p \ ((p \ 0 \ v = min_term) \ (\u. u \\<^sub>t v \ lookup p u = 0))" (is "?L \ (?A \ ?B)") proof assume ?L show "?A \ ?B" proof (intro conjI allI impI) show "p \ 0 \ v = min_term" proof (cases "p = 0") case True show ?thesis proof (rule disjI2) from \?L\ True have "v \\<^sub>t min_term" by (simp add: tt_def) with min_term_min[of v] show "v = min_term" by simp qed next case False thus ?thesis .. qed next fix u assume "u \\<^sub>t v" also note \v \\<^sub>t tt p\ finally have "u \\<^sub>t tt p" . hence "\ tt p \\<^sub>t u" by simp with tt_min[of p u] show "lookup p u = 0" by blast qed next assume "?A \ ?B" hence ?A and ?B by simp_all show ?L proof (cases "p = 0") case True with \?A\ have "v = min_term" by simp with True show ?thesis by (simp add: tt_def) next case False from \?B\ show ?thesis using tt_ge[OF _ False] by auto qed qed lemma tc_not_0: assumes "p \ 0" shows "tc p \ 0" unfolding tc_def in_keys_iff[symmetric] using assms by (rule tt_in_keys) lemma tt_monomial: assumes "c \ 0" shows "tt (monomial c v) = v" proof (rule tt_eqI) from keys_of_monomial[OF assms, of v] show "v \ keys (monomial c v)" by simp next fix u assume "u \ keys (monomial c v)" with keys_of_monomial[OF assms, of v] have "u = v" by simp thus "v \\<^sub>t u" by simp qed lemma tc_monomial [simp]: "tc (monomial c t) = c" proof (cases "c = 0") case True thus ?thesis by simp next case False thus ?thesis by (simp add: tc_def tt_monomial) qed lemma tt_plus_eqI: assumes "p \ 0" and "tt p \\<^sub>t tt q" shows "tt (p + q) = tt p" proof (intro tt_eqI) from tt_less[of "tt p" q "tt q"] \tt p \\<^sub>t tt q\ have "tt p \ keys q" by blast with lookup_add[of p q "tt p"] tc_not_0[OF \p \ 0\] show "tt p \ keys (p + q)" unfolding in_keys_iff tc_def by simp next fix u assume "u \ keys (p + q)" show "tt p \\<^sub>t u" proof (rule ccontr) assume "\ tt p \\<^sub>t u" hence sp: "u \\<^sub>t tt p" by simp hence "u \\<^sub>t tt q" using \tt p \\<^sub>t tt q\ by simp with tt_less[of u q "tt q"] have "u \ keys q" by blast moreover from sp tt_less[of u p "tt p"] have "u \ keys p" by blast ultimately show False using \u \ keys (p + q)\ Poly_Mapping.keys_add[of p q] by auto qed qed lemma tt_plus_lessE: fixes p q assumes "p + q \ 0" and tt: "tt (p + q) \\<^sub>t tt p" shows "tt q \\<^sub>t tt p" proof (cases "p = 0") case True with tt show ?thesis by simp next case False show ?thesis proof (rule ccontr) assume "\ tt q \\<^sub>t tt p" hence "tt p = tt q \ tt p \\<^sub>t tt q" by auto thus False proof assume tt_eq: "tt p = tt q" have "tt p \\<^sub>t tt (p + q)" proof (rule tt_ge_keys) fix u assume "u \ keys (p + q)" hence "u \ keys p \ keys q" proof show "keys (p + q) \ keys p \ keys q" by (fact Poly_Mapping.keys_add) qed thus "tt p \\<^sub>t u" proof assume "u \ keys p" thus ?thesis by (rule tt_min_keys) next assume "u \ keys q" thus ?thesis unfolding tt_eq by (rule tt_min_keys) qed qed (fact \p + q \ 0\) with tt show False by simp next assume "tt p \\<^sub>t tt q" from tt_plus_eqI[OF False this] tt show False by (simp add: ac_simps) qed qed qed lemma tt_plus_lessI: fixes p q :: "_ \\<^sub>0 'b::ring" assumes "p + q \ 0" and tt_eq: "tt q = tt p" and tc_eq: "tc q = - tc p" shows "tt p \\<^sub>t tt (p + q)" proof (rule ccontr) assume "\ tt p \\<^sub>t tt (p + q)" hence "tt p = tt (p + q) \ tt (p + q) \\<^sub>t tt p" by auto thus False proof assume "tt p = tt (p + q)" have "lookup (p + q) (tt p) = (lookup p (tt p)) + (lookup q (tt q))" unfolding tt_eq lookup_add .. also have "... = tc p + tc q" unfolding tc_def .. also have "... = 0" unfolding tc_eq by simp finally have "lookup (p + q) (tt p) = 0" . hence "tt (p + q) \ tt p" using tc_not_0[OF \p + q \ 0\] unfolding tc_def by auto with \tt p = tt (p + q)\ show False by simp next assume "tt (p + q) \\<^sub>t tt p" have "tt q \\<^sub>t tt p" by (rule tt_plus_lessE, fact+) hence "tt q \ tt p" by simp with tt_eq show False by simp qed qed lemma tt_uminus [simp]: "tt (- p) = tt p" by (simp add: tt_def keys_uminus) lemma tc_uminus [simp]: "tc (- p) = - tc p" by (simp add: tc_def) lemma tt_monom_mult: assumes "c \ (0::'b::semiring_no_zero_divisors)" and "p \ 0" shows "tt (monom_mult c t p) = t \ tt p" proof (intro tt_eqI, rule keys_monom_multI, rule tt_in_keys, fact, fact) fix u assume "u \ keys (monom_mult c t p)" then obtain v where "v \ keys p" and u: "u = t \ v" by (rule keys_monom_multE) show "t \ tt p \\<^sub>t u" unfolding u add.commute[of t] by (rule splus_mono, rule tt_min_keys, fact) qed lemma tt_map_scale: "c \ (0::'b::semiring_no_zero_divisors) \ tt (c \ p) = tt p" by (cases "p = 0") (simp_all add: map_scale_eq_monom_mult tt_monom_mult term_simps) lemma tc_monom_mult [simp]: "tc (monom_mult c t p) = (c::'b::semiring_no_zero_divisors) * tc p" proof (cases "c = 0") case True thus ?thesis by simp next case False show ?thesis proof (cases "p = 0") case True thus ?thesis by simp next case False with \c \ 0\ show ?thesis by (simp add: tc_def tt_monom_mult lookup_monom_mult_plus) qed qed corollary tc_map_scale [simp]: "tc (c \ p) = (c::'b::semiring_no_zero_divisors) * tc p" by (simp add: map_scale_eq_monom_mult) lemma in_keys_monom_mult_ge: assumes "v \ keys (monom_mult c t p)" shows "t \ tt p \\<^sub>t v" proof - from keys_monom_mult_subset assms have "v \ (\) t ` (keys p)" .. then obtain u where "u \ keys p" and "v = t \ u" .. from \u \ keys p\ have "tt p \\<^sub>t u" by (rule tt_min_keys) thus "t \ tt p \\<^sub>t v" unfolding \v = t \ u\ by (rule splus_mono) qed lemma lt_ge_tt: "tt p \\<^sub>t lt p" proof (cases "p = 0") case True show ?thesis unfolding True lt_def tt_def by simp next case False show ?thesis by (rule lt_max_keys, rule tt_in_keys, fact False) qed lemma lt_eq_tt_monomial: assumes "is_monomial p" shows "lt p = tt p" proof - from assms obtain c v where "c \ 0" and p: "p = monomial c v" by (rule is_monomial_monomial) from \c \ 0\ have "lt p = v" and "tt p = v" unfolding p by (rule lt_monomial, rule tt_monomial) thus ?thesis by simp qed subsection \@{const higher} and @{const lower}\ lemma lookup_higher: "lookup (higher p u) v = (if u \\<^sub>t v then lookup p v else 0)" by (auto simp add: higher_def lookup_except) lemma lookup_higher_when: "lookup (higher p u) v = (lookup p v when u \\<^sub>t v)" by (auto simp add: lookup_higher when_def) lemma higher_plus: "higher (p + q) v = higher p v + higher q v" by (rule poly_mapping_eqI, simp add: lookup_add lookup_higher) lemma higher_uminus [simp]: "higher (- p) v = -(higher p v)" by (rule poly_mapping_eqI, simp add: lookup_higher) lemma higher_minus: "higher (p - q) v = higher p v - higher q v" by (auto intro!: poly_mapping_eqI simp: lookup_minus lookup_higher) lemma higher_zero [simp]: "higher 0 t = 0" by (rule poly_mapping_eqI, simp add: lookup_higher) lemma higher_eq_iff: "higher p v = higher q v \ (\u. v \\<^sub>t u \ lookup p u = lookup q u)" (is "?L \ ?R") proof assume ?L show ?R proof (intro allI impI) fix u assume "v \\<^sub>t u" moreover from \?L\ have "lookup (higher p v) u = lookup (higher q v) u" by simp ultimately show "lookup p u = lookup q u" by (simp add: lookup_higher) qed next assume ?R show ?L proof (rule poly_mapping_eqI, simp add: lookup_higher, rule) fix u assume "v \\<^sub>t u" with \?R\ show "lookup p u = lookup q u" by simp qed qed lemma higher_eq_zero_iff: "higher p v = 0 \ (\u. v \\<^sub>t u \ lookup p u = 0)" proof - have "higher p v = higher 0 v \ (\u. v \\<^sub>t u \ lookup p u = lookup 0 u)" by (rule higher_eq_iff) thus ?thesis by simp qed lemma keys_higher: "keys (higher p v) = {u\keys p. v \\<^sub>t u}" by (rule set_eqI, simp only: in_keys_iff, simp add: lookup_higher) lemma higher_higher: "higher (higher p u) v = higher p (ord_term_lin.max u v)" by (rule poly_mapping_eqI, simp add: lookup_higher) lemma lookup_lower: "lookup (lower p u) v = (if v \\<^sub>t u then lookup p v else 0)" by (auto simp add: lower_def lookup_except) lemma lookup_lower_when: "lookup (lower p u) v = (lookup p v when v \\<^sub>t u)" by (auto simp add: lookup_lower when_def) lemma lower_plus: "lower (p + q) v = lower p v + lower q v" by (rule poly_mapping_eqI, simp add: lookup_add lookup_lower) lemma lower_uminus [simp]: "lower (- p) v = - lower p v" by (rule poly_mapping_eqI, simp add: lookup_lower) lemma lower_minus: "lower (p - (q::_ \\<^sub>0 'b::ab_group_add)) v = lower p v - lower q v" by (auto intro!: poly_mapping_eqI simp: lookup_minus lookup_lower) lemma lower_zero [simp]: "lower 0 v = 0" by (rule poly_mapping_eqI, simp add: lookup_lower) lemma lower_eq_iff: "lower p v = lower q v \ (\u. u \\<^sub>t v \ lookup p u = lookup q u)" (is "?L \ ?R") proof assume ?L show ?R proof (intro allI impI) fix u assume "u \\<^sub>t v" moreover from \?L\ have "lookup (lower p v) u = lookup (lower q v) u" by simp ultimately show "lookup p u = lookup q u" by (simp add: lookup_lower) qed next assume ?R show ?L proof (rule poly_mapping_eqI, simp add: lookup_lower, rule) fix u assume "u \\<^sub>t v" with \?R\ show "lookup p u = lookup q u" by simp qed qed lemma lower_eq_zero_iff: "lower p v = 0 \ (\u. u \\<^sub>t v \ lookup p u = 0)" proof - have "lower p v = lower 0 v \ (\u. u \\<^sub>t v \ lookup p u = lookup 0 u)" by (rule lower_eq_iff) thus ?thesis by simp qed lemma keys_lower: "keys (lower p v) = {u\keys p. u \\<^sub>t v}" by (rule set_eqI, simp only: in_keys_iff, simp add: lookup_lower) lemma lower_lower: "lower (lower p u) v = lower p (ord_term_lin.min u v)" by (rule poly_mapping_eqI, simp add: lookup_lower) lemma lt_higher: assumes "v \\<^sub>t lt p" shows "lt (higher p v) = lt p" proof (rule lt_eqI_keys, simp_all add: keys_higher, rule conjI, rule lt_in_keys, rule) assume "p = 0" hence "lt p = min_term" by (simp add: lt_def) with min_term_min[of v] assms show False by simp next fix u assume "u \ keys p \ v \\<^sub>t u" hence "u \ keys p" .. thus "u \\<^sub>t lt p" by (rule lt_max_keys) qed fact lemma lc_higher: assumes "v \\<^sub>t lt p" shows "lc (higher p v) = lc p" by (simp add: lc_def lt_higher assms lookup_higher) lemma higher_eq_zero_iff': "higher p v = 0 \ lt p \\<^sub>t v" by (simp add: higher_eq_zero_iff lt_le_iff) lemma higher_id_iff: "higher p v = p \ (p = 0 \ v \\<^sub>t tt p)" (is "?L \ ?R") proof assume ?L show ?R proof (cases "p = 0") case True thus ?thesis .. next case False show ?thesis proof (rule disjI2, rule tt_gr) fix u assume "u \ keys p" hence "lookup p u \ 0" by (simp add: in_keys_iff) from \?L\ have "lookup (higher p v) u = lookup p u" by simp hence "lookup p u = (if v \\<^sub>t u then lookup p u else 0)" by (simp only: lookup_higher) hence "\ v \\<^sub>t u \ lookup p u = 0" by simp with \lookup p u \ 0\ show "v \\<^sub>t u" by auto qed fact qed next assume ?R show ?L proof (cases "p = 0") case True thus ?thesis by simp next case False with \?R\ have "v \\<^sub>t tt p" by simp show ?thesis proof (rule poly_mapping_eqI, simp add: lookup_higher, intro impI) fix u assume "\ v \\<^sub>t u" hence "u \\<^sub>t v" by simp from this \v \\<^sub>t tt p\ have "u \\<^sub>t tt p" by simp hence "\ tt p \\<^sub>t u" by simp with tt_min[of p u] show "lookup p u = 0" by blast qed qed qed lemma tt_lower: assumes "tt p \\<^sub>t v" shows "tt (lower p v) = tt p" proof (cases "p = 0") case True thus ?thesis by simp next case False show ?thesis proof (rule tt_eqI, simp_all add: keys_lower, rule, rule tt_in_keys) fix u assume "u \ keys p \ u \\<^sub>t v" hence "u \ keys p" .. thus "tt p \\<^sub>t u" by (rule tt_min_keys) qed fact+ qed lemma tc_lower: assumes "tt p \\<^sub>t v" shows "tc (lower p v) = tc p" by (simp add: tc_def tt_lower assms lookup_lower) lemma lt_lower: "lt (lower p v) \\<^sub>t lt p" proof (cases "lower p v = 0") case True thus ?thesis by (simp add: lt_def min_term_min) next case False show ?thesis proof (rule lt_le, simp add: lookup_lower, rule impI, rule ccontr) fix u assume "lookup p u \ 0" hence "u \\<^sub>t lt p" by (rule lt_max) moreover assume "lt p \\<^sub>t u" ultimately show False by simp qed qed lemma lt_lower_less: assumes "lower p v \ 0" shows "lt (lower p v) \\<^sub>t v" using assms proof (rule lt_less) fix u assume "v \\<^sub>t u" thus "lookup (lower p v) u = 0" by (simp add: lookup_lower_when) qed lemma lt_lower_eq_iff: "lt (lower p v) = lt p \ (lt p = min_term \ lt p \\<^sub>t v)" (is "?L \ ?R") proof assume ?L show ?R proof (rule ccontr, simp, elim conjE) assume "lt p \ min_term" hence "min_term \\<^sub>t lt p" using min_term_min ord_term_lin.dual_order.not_eq_order_implies_strict by blast assume "\ lt p \\<^sub>t v" hence "v \\<^sub>t lt p" by simp have "lt (lower p v) \\<^sub>t lt p" proof (cases "lower p v = 0") case True thus ?thesis using \min_term \\<^sub>t lt p\ by (simp add: lt_def) next case False show ?thesis proof (rule lt_less) fix u assume "lt p \\<^sub>t u" with \v \\<^sub>t lt p\ have "\ u \\<^sub>t v" by simp thus "lookup (lower p v) u = 0" by (simp add: lookup_lower) qed fact qed with \?L\ show False by simp qed next assume ?R show ?L proof (cases "lt p = min_term") case True hence "lt p \\<^sub>t lt (lower p v)" by (simp add: min_term_min) with lt_lower[of p v] show ?thesis by simp next case False with \?R\ have "lt p \\<^sub>t v" by simp show ?thesis proof (rule lt_eqI_keys, simp_all add: keys_lower, rule conjI, rule lt_in_keys, rule) assume "p = 0" hence "lt p = min_term" by (simp add: lt_def) with False show False .. next fix u assume "u \ keys p \ u \\<^sub>t v" hence "u \ keys p" .. thus "u \\<^sub>t lt p" by (rule lt_max_keys) qed fact qed qed lemma tt_higher: assumes "v \\<^sub>t lt p" shows "tt p \\<^sub>t tt (higher p v)" proof (rule tt_ge_keys, simp add: keys_higher) fix u assume "u \ keys p \ v \\<^sub>t u" hence "u \ keys p" .. thus "tt p \\<^sub>t u" by (rule tt_min_keys) next show "higher p v \ 0" proof (simp add: higher_eq_zero_iff, intro exI conjI) have "p \ 0" proof assume "p = 0" hence "lt p \\<^sub>t v" by (simp add: lt_def min_term_min) with assms show False by simp qed thus "lookup p (lt p) \ 0" using lt_in_keys by auto qed fact qed lemma tt_higher_eq_iff: "tt (higher p v) = tt p \ ((lt p \\<^sub>t v \ tt p = min_term) \ v \\<^sub>t tt p)" (is "?L \ ?R") proof assume ?L show ?R proof (rule ccontr, simp, elim conjE) assume a: "lt p \\<^sub>t v \ tt p \ min_term" assume "\ v \\<^sub>t tt p" hence "tt p \\<^sub>t v" by simp have "tt p \\<^sub>t tt (higher p v)" proof (cases "higher p v = 0") case True with \?L\ have "tt p = min_term" by (simp add: tt_def) with a have "v \\<^sub>t lt p" by auto have "lt p \ min_term" proof assume "lt p = min_term" with \v \\<^sub>t lt p\ show False using min_term_min[of v] by auto qed hence "p \ 0" by (auto simp add: lt_def) from \v \\<^sub>t lt p\ have "higher p v \ 0" by (simp add: higher_eq_zero_iff') from this True show ?thesis .. next case False show ?thesis proof (rule tt_gr) fix u assume "u \ keys (higher p v)" hence "v \\<^sub>t u" by (simp add: keys_higher) with \tt p \\<^sub>t v\ show "tt p \\<^sub>t u" by simp qed fact qed with \?L\ show False by simp qed next assume ?R show ?L proof (cases "lt p \\<^sub>t v \ tt p = min_term") case True hence "lt p \\<^sub>t v" and "tt p = min_term" by simp_all from \lt p \\<^sub>t v\ have "higher p v = 0" by (simp add: higher_eq_zero_iff') with \tt p = min_term\ show ?thesis by (simp add: tt_def) next case False with \?R\ have "v \\<^sub>t tt p" by simp show ?thesis proof (rule tt_eqI, simp_all add: keys_higher, rule conjI, rule tt_in_keys, rule) assume "p = 0" hence "tt p = min_term" by (simp add: tt_def) with \v \\<^sub>t tt p\ min_term_min[of v] show False by simp next fix u assume "u \ keys p \ v \\<^sub>t u" hence "u \ keys p" .. thus "tt p \\<^sub>t u" by (rule tt_min_keys) qed fact qed qed lemma lower_eq_zero_iff': "lower p v = 0 \ (p = 0 \ v \\<^sub>t tt p)" by (auto simp add: lower_eq_zero_iff tt_ge_iff) lemma lower_id_iff: "lower p v = p \ (p = 0 \ lt p \\<^sub>t v)" (is "?L \ ?R") proof assume ?L show ?R proof (cases "p = 0") case True thus ?thesis .. next case False show ?thesis proof (rule disjI2, rule lt_less) fix u assume "v \\<^sub>t u" from \?L\ have "lookup (lower p v) u = lookup p u" by simp hence "lookup p u = (if u \\<^sub>t v then lookup p u else 0)" by (simp only: lookup_lower) hence "v \\<^sub>t u \ lookup p u = 0" by (meson ord_term_lin.leD) with \v \\<^sub>t u\ show "lookup p u = 0" by simp qed fact qed next assume ?R show ?L proof (cases "p = 0", simp) case False with \?R\ have "lt p \\<^sub>t v" by simp show ?thesis proof (rule poly_mapping_eqI, simp add: lookup_lower, intro impI) fix u assume "\ u \\<^sub>t v" hence "v \\<^sub>t u" by simp with \lt p \\<^sub>t v\ have "lt p \\<^sub>t u" by simp hence "\ u \\<^sub>t lt p" by simp with lt_max[of p u] show "lookup p u = 0" by blast qed qed qed lemma lower_higher_commute: "higher (lower p s) t = lower (higher p t) s" by (rule poly_mapping_eqI, simp add: lookup_higher lookup_lower) lemma lt_lower_higher: assumes "v \\<^sub>t lt (lower p u)" shows "lt (lower (higher p v) u) = lt (lower p u)" by (simp add: lower_higher_commute[symmetric] lt_higher[OF assms]) lemma lc_lower_higher: assumes "v \\<^sub>t lt (lower p u)" shows "lc (lower (higher p v) u) = lc (lower p u)" using assms by (simp add: lc_def lt_lower_higher lookup_lower lookup_higher) lemma trailing_monomial_higher: assumes "p \ 0" shows "p = (higher p (tt p)) + monomial (tc p) (tt p)" proof (rule poly_mapping_eqI, simp only: lookup_add) fix v show "lookup p v = lookup (higher p (tt p)) v + lookup (monomial (tc p) (tt p)) v" proof (cases "tt p \\<^sub>t v") case True show ?thesis proof (cases "v = tt p") assume "v = tt p" hence "\ tt p \\<^sub>t v" by simp hence "lookup (higher p (tt p)) v = 0" by (simp add: lookup_higher) moreover from \v = tt p\ have "lookup (monomial (tc p) (tt p)) v = tc p" by (simp add: lookup_single) moreover from \v = tt p\ have "lookup p v = tc p" by (simp add: tc_def) ultimately show ?thesis by simp next assume "v \ tt p" from this True have "tt p \\<^sub>t v" by simp hence "lookup (higher p (tt p)) v = lookup p v" by (simp add: lookup_higher) moreover from \v \ tt p\ have "lookup (monomial (tc p) (tt p)) v = 0" by (simp add: lookup_single) ultimately show ?thesis by simp qed next case False hence "v \\<^sub>t tt p" by simp hence "tt p \ v" by simp from False have "\ tt p \\<^sub>t v" by simp have "lookup p v = 0" proof (rule ccontr) assume "lookup p v \ 0" from tt_min[OF this] False show False by simp qed moreover from \tt p \ v\ have "lookup (monomial (tc p) (tt p)) v = 0" by (simp add: lookup_single) moreover from \\ tt p \\<^sub>t v\ have "lookup (higher p (tt p)) v = 0" by (simp add: lookup_higher) ultimately show ?thesis by simp qed qed lemma higher_lower_decomp: "higher p v + monomial (lookup p v) v + lower p v = p" proof (rule poly_mapping_eqI) fix u show "lookup (higher p v + monomial (lookup p v) v + lower p v) u = lookup p u" proof (rule ord_term_lin.linorder_cases) assume "u \\<^sub>t v" thus ?thesis by (simp add: lookup_add lookup_higher_when lookup_single lookup_lower_when) next assume "u = v" thus ?thesis by (simp add: lookup_add lookup_higher_when lookup_single lookup_lower_when) next assume "v \\<^sub>t u" thus ?thesis by (simp add: lookup_add lookup_higher_when lookup_single lookup_lower_when) qed qed subsection \@{const tail}\ lemma lookup_tail: "lookup (tail p) v = (if v \\<^sub>t lt p then lookup p v else 0)" by (simp add: lookup_lower tail_def) lemma lookup_tail_when: "lookup (tail p) v = (lookup p v when v \\<^sub>t lt p)" by (simp add: lookup_lower_when tail_def) lemma lookup_tail_2: "lookup (tail p) v = (if v = lt p then 0 else lookup p v)" proof (rule ord_term_lin.linorder_cases[of v "lt p"]) assume "v \\<^sub>t lt p" hence "v \ lt p" by simp from this \v \\<^sub>t lt p\ lookup_tail[of p v] show ?thesis by simp next assume "v = lt p" hence "\ v \\<^sub>t lt p" by simp from \v = lt p\ this lookup_tail[of p v] show ?thesis by simp next assume "lt p \\<^sub>t v" hence "\ v \\<^sub>t lt p" by simp hence cp: "lookup p v = 0" using lt_max by blast from \\ v \\<^sub>t lt p\ have "\ v = lt p" and "\ v \\<^sub>t lt p" by simp_all thus ?thesis using cp lookup_tail[of p v] by simp qed lemma leading_monomial_tail: "p = monomial (lc p) (lt p) + tail p" for p::"_ \\<^sub>0 'b::comm_monoid_add" proof (rule poly_mapping_eqI) fix v have "lookup p v = lookup (monomial (lc p) (lt p)) v + lookup (tail p) v" proof (cases "v \\<^sub>t lt p") case True show ?thesis proof (cases "v = lt p") assume "v = lt p" hence "\ v \\<^sub>t lt p" by simp hence c3: "lookup (tail p) v = 0" unfolding lookup_tail[of p v] by simp from \v = lt p\ have c2: "lookup (monomial (lc p) (lt p)) v = lc p" by simp from \v = lt p\ have c1: "lookup p v = lc p" by (simp add: lc_def) from c1 c2 c3 show ?thesis by simp next assume "v \ lt p" from this True have "v \\<^sub>t lt p" by simp hence c2: "lookup (tail p) v = lookup p v" unfolding lookup_tail[of p v] by simp from \v \ lt p\ have c1: "lookup (monomial (lc p) (lt p)) v = 0" by (simp add: lookup_single) from c1 c2 show ?thesis by simp qed next case False hence "lt p \\<^sub>t v" by simp hence "lt p \ v" by simp from False have "\ v \\<^sub>t lt p" by simp have c1: "lookup p v = 0" proof (rule ccontr) assume "lookup p v \ 0" from lt_max[OF this] False show False by simp qed from \lt p \ v\ have c2: "lookup (monomial (lc p) (lt p)) v = 0" by (simp add: lookup_single) from \\ v \\<^sub>t lt p\ lookup_tail[of p v] have c3: "lookup (tail p) v = 0" by simp from c1 c2 c3 show ?thesis by simp qed thus "lookup p v = lookup (monomial (lc p) (lt p) + tail p) v" by (simp add: lookup_add) qed lemma tail_alt: "tail p = except p {lt p}" by (rule poly_mapping_eqI, simp add: lookup_tail_2 lookup_except) corollary tail_alt_2: "tail p = p - monomial (lc p) (lt p)" proof - have "p = monomial (lc p) (lt p) + tail p" by (fact leading_monomial_tail) also have "... = tail p + monomial (lc p) (lt p)" by (simp only: add.commute) finally have "p - monomial (lc p) (lt p) = (tail p + monomial (lc p) (lt p)) - monomial (lc p) (lt p)" by simp thus ?thesis by simp qed lemma tail_zero [simp]: "tail 0 = 0" by (simp only: tail_alt except_zero) lemma lt_tail: assumes "tail p \ 0" shows "lt (tail p) \\<^sub>t lt p" proof (intro lt_less) fix u assume "lt p \\<^sub>t u" hence "\ u \\<^sub>t lt p" by simp thus "lookup (tail p) u = 0" unfolding lookup_tail[of p u] by simp qed fact lemma keys_tail: "keys (tail p) = keys p - {lt p}" by (simp add: tail_alt keys_except) lemma tail_monomial: "tail (monomial c v) = 0" by (metis (no_types, lifting) lookup_tail_2 lookup_single_not_eq lt_less lt_monomial ord_term_lin.dual_order.strict_implies_not_eq single_zero tail_zero) lemma (in ordered_term) mult_scalar_tail_rec_left: "p \ q = monom_mult (punit.lc p) (punit.lt p) q + (punit.tail p) \ q" unfolding punit.lc_def punit.tail_alt by (fact mult_scalar_rec_left) lemma mult_scalar_tail_rec_right: "p \ q = p \ monomial (lc q) (lt q) + p \ tail q" unfolding tail_alt lc_def by (rule mult_scalar_rec_right) lemma lt_tail_max: assumes "tail p \ 0" and "v \ keys p" and "v \\<^sub>t lt p" shows "v \\<^sub>t lt (tail p)" proof (rule lt_max_keys, simp add: keys_tail assms(2)) from assms(3) show "v \ lt p" by auto qed lemma keys_tail_less_lt: assumes "v \ keys (tail p)" shows "v \\<^sub>t lt p" using assms by (meson in_keys_iff lookup_tail) lemma tt_tail: assumes "tail p \ 0" shows "tt (tail p) = tt p" proof (rule tt_eqI, simp_all add: keys_tail) from assms have "p \ 0" using tail_zero by auto show "tt p \ keys p \ tt p \ lt p" proof (rule conjI, rule tt_in_keys, fact) have "tt p \\<^sub>t lt p" by (metis assms lower_eq_zero_iff' tail_def ord_term_lin.le_less_linear) thus "tt p \ lt p" by simp qed next fix u assume "u \ keys p \ u \ lt p" hence "u \ keys p" .. thus "tt p \\<^sub>t u" by (rule tt_min_keys) qed lemma tc_tail: assumes "tail p \ 0" shows "tc (tail p) = tc p" proof (simp add: tc_def tt_tail[OF assms] lookup_tail_2, rule) assume "tt p = lt p" moreover have "tt p \\<^sub>t lt p" by (metis assms lower_eq_zero_iff' tail_def ord_term_lin.le_less_linear) ultimately show "lookup p (lt p) = 0" by simp qed lemma tt_tail_min: assumes "s \ keys p" shows "tt (tail p) \\<^sub>t s" proof (cases "tail p = 0") case True hence "tt (tail p) = min_term" by (simp add: tt_def) thus ?thesis by (simp add: min_term_min) next case False from assms show ?thesis by (simp add: tt_tail[OF False], rule tt_min_keys) qed lemma tail_monom_mult: "tail (monom_mult c t p) = monom_mult (c::'b::semiring_no_zero_divisors) t (tail p)" proof (cases "p = 0") case True hence "tail p = 0" and "monom_mult c t p = 0" by simp_all thus ?thesis by simp next case False show ?thesis proof (cases "c = 0") case True hence "monom_mult c t p = 0" and "monom_mult c t (tail p) = 0" by simp_all thus ?thesis by simp next case False let ?a = "monom_mult c t p" let ?b = "monom_mult c t (tail p)" from \p \ 0\ False have "?a \ 0" by (simp add: monom_mult_eq_zero_iff) from False \p \ 0\ have lt_a: "lt ?a = t \ lt p" by (rule lt_monom_mult) show ?thesis proof (rule poly_mapping_eqI, simp add: lookup_tail lt_a, intro conjI impI) fix u assume "u \\<^sub>t t \ lt p" show "lookup (monom_mult c t p) u = lookup (monom_mult c t (tail p)) u" proof (cases "t adds\<^sub>p u") case True then obtain v where "u = t \ v" by (rule adds_ppE) from \u \\<^sub>t t \ lt p\ have "v \\<^sub>t lt p" unfolding \u = t \ v\ by (rule ord_term_strict_canc) hence "lookup p v = lookup (tail p) v" by (simp add: lookup_tail) thus ?thesis by (simp add: \u = t \ v\ lookup_monom_mult_plus) next case False hence "lookup ?a u = 0" by (simp add: lookup_monom_mult) moreover have "lookup ?b u = 0" proof (rule ccontr, simp only: in_keys_iff[symmetric] keys_monom_mult[OF \c \ 0\]) assume "u \ (\) t ` keys (tail p)" then obtain v where "u = t \ v" by auto hence "t adds\<^sub>p u" by (simp add: term_simps) with False show False .. qed ultimately show ?thesis by simp qed next fix u assume "\ u \\<^sub>t t \ lt p" hence "t \ lt p \\<^sub>t u" by simp show "lookup (monom_mult c t (tail p)) u = 0" proof (rule ccontr, simp only: in_keys_iff[symmetric] keys_monom_mult[OF False]) assume "u \ (\) t ` keys (tail p)" then obtain v where "v \ keys (tail p)" and "u = t \ v" by auto from \t \ lt p \\<^sub>t u\ have "lt p \\<^sub>t v" unfolding \u = t \ v\ by (rule ord_term_canc) from \v \ keys (tail p)\ have "v \ keys p" and "v \ lt p" by (simp_all add: keys_tail) from \v \ keys p\ have "v \\<^sub>t lt p" by (rule lt_max_keys) with \lt p \\<^sub>t v\ have "v = lt p " by simp with \v \ lt p\ show False .. qed qed qed qed lemma keys_plus_eq_lt_tt_D: assumes "keys (p + q) = {lt p, tt q}" and "lt q \\<^sub>t lt p" and "tt q \\<^sub>t tt (p::_ \\<^sub>0 'b::comm_monoid_add)" shows "tail p + higher q (tt q) = 0" proof - note assms(3) also have "... \\<^sub>t lt p" by (rule lt_ge_tt) finally have "tt q \\<^sub>t lt p" . hence "lt p \ tt q" by simp have "q \ 0" proof assume "q = 0" hence "tt q = min_term" by (simp add: tt_def) with \q = 0\ assms(1) have "keys p = {lt p, min_term}" by simp hence "min_term \ keys p" by simp hence "tt p \\<^sub>t tt q" unfolding \tt q = min_term\ by (rule tt_min_keys) with assms(3) show False by simp qed hence "tc q \ 0" by (rule tc_not_0) have "p = monomial (lc p) (lt p) + tail p" by (rule leading_monomial_tail) moreover from \q \ 0\ have "q = higher q (tt q) + monomial (tc q) (tt q)" by (rule trailing_monomial_higher) ultimately have pq: "p + q = (monomial (lc p) (lt p) + monomial (tc q) (tt q)) + (tail p + higher q (tt q))" (is "_ = (?m1 + ?m2) + ?b") by (simp add: algebra_simps) have keys_m1: "keys ?m1 = {lt p}" proof (rule keys_of_monomial, rule lc_not_0, rule) assume "p = 0" with assms(2) have "lt q \\<^sub>t min_term" by (simp add: lt_def) with min_term_min[of "lt q"] show False by simp qed moreover from \tc q \ 0\ have keys_m2: "keys ?m2 = {tt q}" by (rule keys_of_monomial) ultimately have keys_m1_m2: "keys (?m1 + ?m2) = {lt p, tt q}" using \lt p \ tt q\ keys_plus_eqI[of ?m1 ?m2] by auto show ?thesis proof (rule ccontr) assume "?b \ 0" hence "keys ?b \ {}" by simp then obtain t where "t \ keys ?b" by blast hence t_in: "t \ keys (tail p) \ keys (higher q (tt q))" using Poly_Mapping.keys_add[of "tail p" "higher q (tt q)"] by blast hence "t \ lt p" proof (rule, simp add: keys_tail, simp add: keys_higher, elim conjE) assume "t \ keys q" hence "t \\<^sub>t lt q" by (rule lt_max_keys) from this assms(2) show ?thesis by simp qed moreover from t_in have "t \ tt q" proof (rule, simp add: keys_tail, elim conjE) assume "t \ keys p" hence "tt p \\<^sub>t t" by (rule tt_min_keys) with assms(3) show ?thesis by simp next assume "t \ keys (higher q (tt q))" thus ?thesis by (auto simp only: keys_higher) qed ultimately have "t \ keys (?m1 + ?m2)" by (simp add: keys_m1_m2) moreover from in_keys_plusI2[OF \t \ keys ?b\ this] have "t \ keys (?m1 + ?m2)" by (simp only: keys_m1_m2 pq[symmetric] assms(1)) ultimately show False .. qed qed subsection \Order Relation on Polynomials\ definition ord_strict_p :: "('t \\<^sub>0 'b::zero) \ ('t \\<^sub>0 'b) \ bool" (infixl "\\<^sub>p" 50) where "p \\<^sub>p q \ (\v. lookup p v = 0 \ lookup q v \ 0 \ (\u. v \\<^sub>t u \ lookup p u = lookup q u))" definition ord_p :: "('t \\<^sub>0 'b::zero) \ ('t \\<^sub>0 'b) \ bool" (infixl "\\<^sub>p" 50) where "ord_p p q \ (p \\<^sub>p q \ p = q)" lemma ord_strict_pI: assumes "lookup p v = 0" and "lookup q v \ 0" and "\u. v \\<^sub>t u \ lookup p u = lookup q u" shows "p \\<^sub>p q" unfolding ord_strict_p_def using assms by blast lemma ord_strict_pE: assumes "p \\<^sub>p q" obtains v where "lookup p v = 0" and "lookup q v \ 0" and "\u. v \\<^sub>t u \ lookup p u = lookup q u" using assms unfolding ord_strict_p_def by blast lemma not_ord_pI: assumes "lookup p v \ lookup q v" and "lookup p v \ 0" and "\u. v \\<^sub>t u \ lookup p u = lookup q u" shows "\ p \\<^sub>p q" proof assume "p \\<^sub>p q" hence "p \\<^sub>p q \ p = q" by (simp only: ord_p_def) thus False proof assume "p \\<^sub>p q" then obtain v' where 1: "lookup p v' = 0" and 2: "lookup q v' \ 0" and 3: "\u. v' \\<^sub>t u \ lookup p u = lookup q u" by (rule ord_strict_pE, blast) from 1 2 have "lookup p v' \ lookup q v'" by simp hence "\ v \\<^sub>t v'" using assms(3) by blast hence "v' \\<^sub>t v \ v' = v" by auto thus ?thesis proof assume "v' \\<^sub>t v" hence "lookup p v = lookup q v" by (rule 3) with assms(1) show ?thesis .. next assume "v' = v" with assms(2) 1 show ?thesis by auto qed next assume "p = q" hence "lookup p v = lookup q v" by simp with assms(1) show ?thesis .. qed qed corollary not_ord_strict_pI: assumes "lookup p v \ lookup q v" and "lookup p v \ 0" and "\u. v \\<^sub>t u \ lookup p u = lookup q u" shows "\ p \\<^sub>p q" proof - from assms have "\ p \\<^sub>p q" by (rule not_ord_pI) thus ?thesis by (simp add: ord_p_def) qed lemma ord_strict_higher: "p \\<^sub>p q \ (\v. lookup p v = 0 \ lookup q v \ 0 \ higher p v = higher q v)" unfolding ord_strict_p_def higher_eq_iff .. lemma ord_strict_p_asymmetric: assumes "p \\<^sub>p q" shows "\ q \\<^sub>p p" using assms unfolding ord_strict_p_def proof fix v1::'t assume "lookup p v1 = 0 \ lookup q v1 \ 0 \ (\u. v1 \\<^sub>t u \ lookup p u = lookup q u)" hence "lookup p v1 = 0" and "lookup q v1 \ 0" and v1: "\u. v1 \\<^sub>t u \ lookup p u = lookup q u" by auto show "\ (\v. lookup q v = 0 \ lookup p v \ 0 \ (\u. v \\<^sub>t u \ lookup q u = lookup p u))" proof (intro notI, erule exE) fix v2::'t assume "lookup q v2 = 0 \ lookup p v2 \ 0 \ (\u. v2 \\<^sub>t u \ lookup q u = lookup p u)" hence "lookup q v2 = 0" and "lookup p v2 \ 0" and v2: "\u. v2 \\<^sub>t u \ lookup q u = lookup p u" by auto show False proof (rule ord_term_lin.linorder_cases) assume "v1 \\<^sub>t v2" from v1[rule_format, OF this] \lookup q v2 = 0\ \lookup p v2 \ 0\ show ?thesis by simp next assume "v1 = v2" thus ?thesis using \lookup p v1 = 0\ \lookup p v2 \ 0\ by simp next assume "v2 \\<^sub>t v1" from v2[rule_format, OF this] \lookup p v1 = 0\ \lookup q v1 \ 0\ show ?thesis by simp qed qed qed lemma ord_strict_p_irreflexive: "\ p \\<^sub>p p" unfolding ord_strict_p_def proof (intro notI, erule exE) fix v::'t assume "lookup p v = 0 \ lookup p v \ 0 \ (\u. v \\<^sub>t u \ lookup p u = lookup p u)" hence "lookup p v = 0" and "lookup p v \ 0" by auto thus False by simp qed lemma ord_strict_p_transitive: assumes "a \\<^sub>p b" and "b \\<^sub>p c" shows "a \\<^sub>p c" proof - from \a \\<^sub>p b\ obtain v1 where "lookup a v1 = 0" and "lookup b v1 \ 0" and v1[rule_format]: "(\u. v1 \\<^sub>t u \ lookup a u = lookup b u)" unfolding ord_strict_p_def by auto from \b \\<^sub>p c\ obtain v2 where "lookup b v2 = 0" and "lookup c v2 \ 0" and v2[rule_format]: "(\u. v2 \\<^sub>t u \ lookup b u = lookup c u)" unfolding ord_strict_p_def by auto show "a \\<^sub>p c" proof (rule ord_term_lin.linorder_cases) assume "v1 \\<^sub>t v2" show ?thesis unfolding ord_strict_p_def proof show "lookup a v2 = 0 \ lookup c v2 \ 0 \ (\u. v2 \\<^sub>t u \ lookup a u = lookup c u)" proof (intro conjI allI impI) from \lookup b v2 = 0\ v1[OF \v1 \\<^sub>t v2\] show "lookup a v2 = 0" by simp next from \lookup c v2 \ 0\ show "lookup c v2 \ 0" . next fix u assume "v2 \\<^sub>t u" from ord_term_lin.less_trans[OF \v1 \\<^sub>t v2\ this] have "v1 \\<^sub>t u" . from v2[OF \v2 \\<^sub>t u\] v1[OF this] show "lookup a u = lookup c u" by simp qed qed next assume "v2 \\<^sub>t v1" show ?thesis unfolding ord_strict_p_def proof show "lookup a v1 = 0 \ lookup c v1 \ 0 \ (\u. v1 \\<^sub>t u \ lookup a u = lookup c u)" proof (intro conjI allI impI) from \lookup a v1 = 0\ show "lookup a v1 = 0" . next from \lookup b v1 \ 0\ v2[OF \v2 \\<^sub>t v1\] show "lookup c v1 \ 0" by simp next fix u assume "v1 \\<^sub>t u" from ord_term_lin.less_trans[OF \v2 \\<^sub>t v1\ this] have "v2 \\<^sub>t u" . from v1[OF \v1 \\<^sub>t u\] v2[OF this] show "lookup a u = lookup c u" by simp qed qed next assume "v1 = v2" thus ?thesis using \lookup b v1 \ 0\ \lookup b v2 = 0\ by simp qed qed sublocale order ord_p ord_strict_p proof (intro order_strictI) fix p q :: "'t \\<^sub>0 'b" show "(p \\<^sub>p q) = (p \\<^sub>p q \ p = q)" unfolding ord_p_def .. next fix p q :: "'t \\<^sub>0 'b" assume "p \\<^sub>p q" thus "\ q \\<^sub>p p" by (rule ord_strict_p_asymmetric) next fix p::"'t \\<^sub>0 'b" show "\ p \\<^sub>p p" by (fact ord_strict_p_irreflexive) next fix a b c :: "'t \\<^sub>0 'b" assume "a \\<^sub>p b" and "b \\<^sub>p c" thus "a \\<^sub>p c" by (rule ord_strict_p_transitive) qed lemma ord_p_zero_min: "0 \\<^sub>p p" unfolding ord_p_def ord_strict_p_def proof (cases "p = 0") case True thus "(\v. lookup 0 v = 0 \ lookup p v \ 0 \ (\u. v \\<^sub>t u \ lookup 0 u = lookup p u)) \ 0 = p" by auto next case False show "(\v. lookup 0 v = 0 \ lookup p v \ 0 \ (\u. v \\<^sub>t u \ lookup 0 u = lookup p u)) \ 0 = p" proof show "(\v. lookup 0 v = 0 \ lookup p v \ 0 \ (\u. v \\<^sub>t u \ lookup 0 u = lookup p u))" proof show "lookup 0 (lt p) = 0 \ lookup p (lt p) \ 0 \ (\u. (lt p) \\<^sub>t u \ lookup 0 u = lookup p u)" proof (intro conjI allI impI) show "lookup 0 (lt p) = 0" by (transfer, simp) next from lc_not_0[OF False] show "lookup p (lt p) \ 0" unfolding lc_def . next fix u assume "lt p \\<^sub>t u" hence "\ u \\<^sub>t lt p" by simp hence "lookup p u = 0" using lt_max[of p u] by metis thus "lookup 0 u = lookup p u" by simp qed qed qed qed lemma lt_ord_p: assumes "lt p \\<^sub>t lt q" shows "p \\<^sub>p q" proof - have "q \ 0" proof assume "q = 0" with assms have "lt p \\<^sub>t min_term" by (simp add: lt_def) with min_term_min[of "lt p"] show False by simp qed show ?thesis unfolding ord_strict_p_def proof (intro exI conjI allI impI) show "lookup p (lt q) = 0" proof (rule ccontr) assume "lookup p (lt q) \ 0" from lt_max[OF this] \lt p \\<^sub>t lt q\ show False by simp qed next from lc_not_0[OF \q \ 0\] show "lookup q (lt q) \ 0" unfolding lc_def . next fix u assume "lt q \\<^sub>t u" hence "lt p \\<^sub>t u" using \lt p \\<^sub>t lt q\ by simp have c1: "lookup q u = 0" proof (rule ccontr) assume "lookup q u \ 0" from lt_max[OF this] \lt q \\<^sub>t u\ show False by simp qed have c2: "lookup p u = 0" proof (rule ccontr) assume "lookup p u \ 0" from lt_max[OF this] \lt p \\<^sub>t u\ show False by simp qed from c1 c2 show "lookup p u = lookup q u" by simp qed qed lemma ord_p_lt: assumes "p \\<^sub>p q" shows "lt p \\<^sub>t lt q" proof (rule ccontr) assume "\ lt p \\<^sub>t lt q" hence "lt q \\<^sub>t lt p" by simp from lt_ord_p[OF this] \p \\<^sub>p q\ show False by simp qed lemma ord_p_tail: assumes "p \ 0" and "lt p = lt q" and "p \\<^sub>p q" shows "tail p \\<^sub>p tail q" using assms unfolding ord_strict_p_def proof - assume "p \ 0" and "lt p = lt q" and "\v. lookup p v = 0 \ lookup q v \ 0 \ (\u. v \\<^sub>t u \ lookup p u = lookup q u)" then obtain v where "lookup p v = 0" and "lookup q v \ 0" and a: "\u. v \\<^sub>t u \ lookup p u = lookup q u" by auto from lt_max[OF \lookup q v \ 0\] \lt p = lt q\ have "v \\<^sub>t lt p \ v = lt p" by auto hence "v \\<^sub>t lt p" proof assume "v \\<^sub>t lt p" thus ?thesis . next assume "v = lt p" thus ?thesis using lc_not_0[OF \p \ 0\] \lookup p v = 0\ unfolding lc_def by auto qed have pt: "lookup (tail p) v = lookup p v" using lookup_tail[of p v] \v \\<^sub>t lt p\ by simp have "q \ 0" proof assume "q = 0" hence "p \\<^sub>p 0" using \p \\<^sub>p q\ by simp hence "\ 0 \\<^sub>p p" by auto thus False using ord_p_zero_min[of p] by simp qed have qt: "lookup (tail q) v = lookup q v" using lookup_tail[of q v] \v \\<^sub>t lt p\ \lt p = lt q\ by simp show "\w. lookup (tail p) w = 0 \ lookup (tail q) w \ 0 \ (\u. w \\<^sub>t u \ lookup (tail p) u = lookup (tail q) u)" proof (intro exI conjI allI impI) from pt \lookup p v = 0\ show "lookup (tail p) v = 0" by simp next from qt \lookup q v \ 0\ show "lookup (tail q) v \ 0" by simp next fix u assume "v \\<^sub>t u" from a[rule_format, OF \v \\<^sub>t u\] lookup_tail[of p u] lookup_tail[of q u] \lt p = lt q\ show "lookup (tail p) u = lookup (tail q) u" by simp qed qed lemma tail_ord_p: assumes "p \ 0" shows "tail p \\<^sub>p p" proof (cases "tail p = 0") case True with ord_p_zero_min[of p] \p \ 0\ show ?thesis by simp next case False from lt_tail[OF False] show ?thesis by (rule lt_ord_p) qed lemma higher_lookup_eq_zero: assumes pt: "lookup p v = 0" and hp: "higher p v = 0" and le: "q \\<^sub>p p" shows "(lookup q v = 0) \ (higher q v) = 0" using le unfolding ord_p_def proof assume "q \\<^sub>p p" thus ?thesis unfolding ord_strict_p_def proof fix w assume "lookup q w = 0 \ lookup p w \ 0 \ (\u. w \\<^sub>t u \ lookup q u = lookup p u)" hence qs: "lookup q w = 0" and ps: "lookup p w \ 0" and u: "\u. w \\<^sub>t u \ lookup q u = lookup p u" by auto from hp have pu: "\u. v \\<^sub>t u \ lookup p u = 0" by (simp only: higher_eq_zero_iff) from pu[rule_format, of w] ps have "\ v \\<^sub>t w" by auto hence "w \\<^sub>t v" by simp hence "w \\<^sub>t v \ w = v" by auto hence st: "w \\<^sub>t v" proof (rule disjE, simp_all) assume "w = v" from this pt ps show False by simp qed show ?thesis proof from u[rule_format, OF st] pt show "lookup q v = 0" by simp next have "\u. v \\<^sub>t u \ lookup q u = 0" proof (intro allI, intro impI) fix u assume "v \\<^sub>t u" from this st have "w \\<^sub>t u" by simp from u[rule_format, OF this] pu[rule_format, OF \v \\<^sub>t u\] show "lookup q u = 0" by simp qed thus "higher q v = 0" by (simp only: higher_eq_zero_iff) qed qed next assume "q = p" thus ?thesis using assms by simp qed lemma ord_strict_p_recI: assumes "lt p = lt q" and "lc p = lc q" and tail: "tail p \\<^sub>p tail q" shows "p \\<^sub>p q" proof - from tail obtain v where pt: "lookup (tail p) v = 0" and qt: "lookup (tail q) v \ 0" and a: "\u. v \\<^sub>t u \ lookup (tail p) u = lookup (tail q) u" unfolding ord_strict_p_def by auto from qt lookup_zero[of v] have "tail q \ 0" by auto from lt_max[OF qt] lt_tail[OF this] have "v \\<^sub>t lt q" by simp hence "v \\<^sub>t lt p" using \lt p = lt q\ by simp show ?thesis unfolding ord_strict_p_def proof (rule exI[of _ v], intro conjI allI impI) from lookup_tail[of p v] \v \\<^sub>t lt p\ pt show "lookup p v = 0" by simp next from lookup_tail[of q v] \v \\<^sub>t lt q\ qt show "lookup q v \ 0" by simp next fix u assume "v \\<^sub>t u" from this a have s: "lookup (tail p) u = lookup (tail q) u" by simp show "lookup p u = lookup q u" proof (cases "u = lt p") case True from True \lc p = lc q\ \lt p = lt q\ show ?thesis unfolding lc_def by simp next case False from False s lookup_tail_2[of p u] lookup_tail_2[of q u] \lt p = lt q\ show ?thesis by simp qed qed qed lemma ord_strict_p_recE1: assumes "p \\<^sub>p q" shows "q \ 0" proof assume "q = 0" from this assms ord_p_zero_min[of p] show False by simp qed lemma ord_strict_p_recE2: assumes "p \ 0" and "p \\<^sub>p q" and "lt p = lt q" shows "lc p = lc q" proof - from \p \\<^sub>p q\ obtain v where pt: "lookup p v = 0" and qt: "lookup q v \ 0" and a: "\u. v \\<^sub>t u \ lookup p u = lookup q u" unfolding ord_strict_p_def by auto show ?thesis proof (cases "v \\<^sub>t lt p") case True from this a have "lookup p (lt p) = lookup q (lt p)" by simp thus ?thesis using \lt p = lt q\ unfolding lc_def by simp next case False from this lt_max[OF qt] \lt p = lt q\ have "v = lt p" by simp from this lc_not_0[OF \p \ 0\] pt show ?thesis unfolding lc_def by auto qed qed lemma ord_strict_p_rec [code]: "p \\<^sub>p q = (q \ 0 \ (p = 0 \ (let v1 = lt p; v2 = lt q in (v1 \\<^sub>t v2 \ (v1 = v2 \ lookup p v1 = lookup q v2 \ lower p v1 \\<^sub>p lower q v2)) ) ) )" (is "?L = ?R") proof assume ?L show ?R proof (intro conjI, rule ord_strict_p_recE1, fact) have "((lt p = lt q \ lc p = lc q \ tail p \\<^sub>p tail q) \ lt p \\<^sub>t lt q) \ p = 0" proof (intro disjCI) assume "p \ 0" and nl: "\ lt p \\<^sub>t lt q" from \?L\ have "p \\<^sub>p q" by simp from ord_p_lt[OF this] nl have "lt p = lt q" by simp show "lt p = lt q \ lc p = lc q \ tail p \\<^sub>p tail q" by (intro conjI, fact, rule ord_strict_p_recE2, fact+, rule ord_p_tail, fact+) qed thus "p = 0 \ (let v1 = lt p; v2 = lt q in (v1 \\<^sub>t v2 \ v1 = v2 \ lookup p v1 = lookup q v2 \ lower p v1 \\<^sub>p lower q v2) )" unfolding lc_def tail_def by auto qed next assume ?R hence "q \ 0" and dis: "p = 0 \ (let v1 = lt p; v2 = lt q in (v1 \\<^sub>t v2 \ v1 = v2 \ lookup p v1 = lookup q v2 \ lower p v1 \\<^sub>p lower q v2) )" by simp_all show ?L proof (cases "p = 0") assume "p = 0" hence "p \\<^sub>p q" using ord_p_zero_min[of q] by simp thus ?thesis using \p = 0\ \q \ 0\ by simp next assume "p \ 0" hence "let v1 = lt p; v2 = lt q in (v1 \\<^sub>t v2 \ v1 = v2 \ lookup p v1 = lookup q v2 \ lower p v1 \\<^sub>p lower q v2)" using dis by simp hence "lt p \\<^sub>t lt q \ (lt p = lt q \ lc p = lc q \ tail p \\<^sub>p tail q)" unfolding lc_def tail_def by (simp add: Let_def) thus ?thesis proof assume "lt p \\<^sub>t lt q" from lt_ord_p[OF this] show ?thesis . next assume "lt p = lt q \ lc p = lc q \ tail p \\<^sub>p tail q" hence "lt p = lt q" and "lc p = lc q" and "tail p \\<^sub>p tail q" by simp_all thus ?thesis by (rule ord_strict_p_recI) qed qed qed lemma ord_strict_p_monomial_iff: "p \\<^sub>p monomial c v \ (c \ 0 \ (p = 0 \ lt p \\<^sub>t v))" proof - from ord_p_zero_min[of "tail p"] have *: "\ tail p \\<^sub>p 0" by auto show ?thesis by (simp add: ord_strict_p_rec[of p] Let_def tail_def[symmetric] lc_def[symmetric] monomial_0_iff tail_monomial *, simp add: lt_monomial cong: conj_cong) qed corollary ord_strict_p_monomial_plus: assumes "p \\<^sub>p monomial c v" and "q \\<^sub>p monomial c v" shows "p + q \\<^sub>p monomial c v" proof - from assms(1) have "c \ 0" and "p = 0 \ lt p \\<^sub>t v" by (simp_all add: ord_strict_p_monomial_iff) from this(2) show ?thesis proof assume "p = 0" with assms(2) show ?thesis by simp next assume "lt p \\<^sub>t v" from assms(2) have "q = 0 \ lt q \\<^sub>t v" by (simp add: ord_strict_p_monomial_iff) thus ?thesis proof assume "q = 0" with assms(1) show ?thesis by simp next assume "lt q \\<^sub>t v" with \lt p \\<^sub>t v\ have "lt (p + q) \\<^sub>t v" using lt_plus_le_max ord_term_lin.dual_order.strict_trans2 ord_term_lin.max_less_iff_conj by blast with \c \ 0\ show ?thesis by (simp add: ord_strict_p_monomial_iff) qed qed qed lemma ord_strict_p_monom_mult: assumes "p \\<^sub>p q" and "c \ (0::'b::semiring_no_zero_divisors)" shows "monom_mult c t p \\<^sub>p monom_mult c t q" proof - from assms(1) obtain v where 1: "lookup p v = 0" and 2: "lookup q v \ 0" and 3: "\u. v \\<^sub>t u \ lookup p u = lookup q u" unfolding ord_strict_p_def by auto show ?thesis unfolding ord_strict_p_def proof (intro exI conjI allI impI) from 1 show "lookup (monom_mult c t p) (t \ v) = 0" by (simp add: lookup_monom_mult_plus) next from 2 assms(2) show "lookup (monom_mult c t q) (t \ v) \ 0" by (simp add: lookup_monom_mult_plus) next fix u assume "t \ v \\<^sub>t u" show "lookup (monom_mult c t p) u = lookup (monom_mult c t q) u" proof (cases "t adds\<^sub>p u") case True then obtain w where u: "u = t \ w" .. from \t \ v \\<^sub>t u\ have "v \\<^sub>t w" unfolding u by (rule ord_term_strict_canc) hence "lookup p w = lookup q w" by (rule 3) thus ?thesis by (simp add: u lookup_monom_mult_plus) next case False thus ?thesis by (simp add: lookup_monom_mult) qed qed qed lemma ord_strict_p_plus: assumes "p \\<^sub>p q" and "keys r \ keys q = {}" shows "p + r \\<^sub>p q + r" proof - from assms(1) obtain v where 1: "lookup p v = 0" and 2: "lookup q v \ 0" and 3: "\u. v \\<^sub>t u \ lookup p u = lookup q u" unfolding ord_strict_p_def by auto have eq: "lookup r v = 0" by (meson "2" assms(2) disjoint_iff_not_equal in_keys_iff) show ?thesis unfolding ord_strict_p_def proof (intro exI conjI allI impI, simp_all add: lookup_add) from 1 show "lookup p v + lookup r v = 0" by (simp add: eq) next from 2 show "lookup q v + lookup r v \ 0" by (simp add: eq) next fix u assume "v \\<^sub>t u" hence "lookup p u = lookup q u" by (rule 3) thus "lookup p u + lookup r u = lookup q u + lookup r u" by simp qed qed lemma poly_mapping_tail_induct [case_names 0 tail]: assumes "P 0" and "\p. p \ 0 \ P (tail p) \ P p" shows "P p" proof (induct "card (keys p)" arbitrary: p) case 0 with finite_keys[of p] have "keys p = {}" by simp hence "p = 0" by simp from \P 0\ show ?case unfolding \p = 0\ . next case ind: (Suc n) from ind(2) have "keys p \ {}" by auto hence "p \ 0" by simp thus ?case proof (rule assms(2)) show "P (tail p)" proof (rule ind(1)) from \p \ 0\ have "lt p \ keys p" by (rule lt_in_keys) hence "card (keys (tail p)) = card (keys p) - 1" by (simp add: keys_tail) also have "... = n" unfolding ind(2)[symmetric] by simp finally show "n = card (keys (tail p))" by simp qed qed qed lemma poly_mapping_neqE: assumes "p \ q" obtains v where "v \ keys p \ keys q" and "lookup p v \ lookup q v" and "\u. v \\<^sub>t u \ lookup p u = lookup q u" proof - let ?A = "{v. lookup p v \ lookup q v}" define v where "v = ord_term_lin.Max ?A" have "?A \ keys p \ keys q" using UnI2 in_keys_iff by fastforce also have "finite ..." by (rule finite_UnI) (fact finite_keys)+ finally(finite_subset) have fin: "finite ?A" . moreover have "?A \ {}" proof assume "?A = {}" hence "p = q" using poly_mapping_eqI by fastforce with assms show False .. qed ultimately have "v \ ?A" unfolding v_def by (rule ord_term_lin.Max_in) show ?thesis proof from \?A \ keys p \ keys q\ \v \ ?A\ show "v \ keys p \ keys q" .. next from \v \ ?A\ show "lookup p v \ lookup q v" by simp next fix u assume "v \\<^sub>t u" show "lookup p u = lookup q u" proof (rule ccontr) assume "lookup p u \ lookup q u" hence "u \ ?A" by simp with fin have "u \\<^sub>t v" unfolding v_def by (rule ord_term_lin.Max_ge) with \v \\<^sub>t u\ show False by simp qed qed qed subsection \Monomials\ lemma keys_monomial: assumes "is_monomial p" shows "keys p = {lt p}" using assms by (metis is_monomial_monomial lt_monomial keys_of_monomial) lemma monomial_eq_itself: assumes "is_monomial p" shows "monomial (lc p) (lt p) = p" proof - from assms have "p \ 0" by (rule monomial_not_0) hence "lc p \ 0" by (rule lc_not_0) hence keys1: "keys (monomial (lc p) (lt p)) = {lt p}" by (rule keys_of_monomial) show ?thesis by (rule poly_mapping_keys_eqI, simp only: keys_monomial[OF assms] keys1, simp only: keys1 lookup_single Poly_Mapping.when_def, auto simp add: lc_def) qed lemma lt_eq_min_term_monomial: assumes "lt p = min_term" shows "monomial (lc p) min_term = p" proof (rule poly_mapping_eqI) fix v from min_term_min[of v] have "v = min_term \ min_term \\<^sub>t v" by auto thus "lookup (monomial (lc p) min_term) v = lookup p v" proof assume "v = min_term" thus ?thesis by (simp add: lookup_single lc_def assms) next assume "min_term \\<^sub>t v" moreover have "v \ keys p" proof assume "v \ keys p" hence "v \\<^sub>t lt p" by (rule lt_max_keys) with \min_term \\<^sub>t v\ show False by (simp add: assms) qed ultimately show ?thesis by (simp add: lookup_single in_keys_iff) qed qed lemma is_monomial_monomial_ordered: assumes "is_monomial p" obtains c v where "c \ 0" and "lc p = c" and "lt p = v" and "p = monomial c v" proof - from assms obtain c v where "c \ 0" and p_eq: "p = monomial c v" by (rule is_monomial_monomial) note this(1) moreover have "lc p = c" unfolding p_eq by (rule lc_monomial) moreover from \c \ 0\ have "lt p = v" unfolding p_eq by (rule lt_monomial) ultimately show ?thesis using p_eq .. qed lemma monomial_plus_not_0: assumes "c \ 0" and "lt p \\<^sub>t v" shows "monomial c v + p \ 0" proof assume "monomial c v + p = 0" hence "0 = lookup (monomial c v + p) v" by simp also have "... = c + lookup p v" by (simp add: lookup_add) also have "... = c" proof - from assms(2) have "\ v \\<^sub>t lt p" by simp with lt_max[of p v] have "lookup p v = 0" by blast thus ?thesis by simp qed finally show False using \c \ 0\ by simp qed lemma lt_monomial_plus: assumes "c \ (0::'b::comm_monoid_add)" and "lt p \\<^sub>t v" shows "lt (monomial c v + p) = v" proof - have eq: "lt (monomial c v) = v" by (simp only: lt_monomial[OF \c \ 0\]) moreover have "lt (p + monomial c v) = lt (monomial c v)" by (rule lt_plus_eqI, simp only: eq, fact) ultimately show ?thesis by (simp add: add.commute) qed lemma lc_monomial_plus: assumes "c \ (0::'b::comm_monoid_add)" and "lt p \\<^sub>t v" shows "lc (monomial c v + p) = c" proof - from assms(2) have "\ v \\<^sub>t lt p" by simp with lt_max[of p v] have "lookup p v = 0" by blast thus ?thesis by (simp add: lc_def lt_monomial_plus[OF assms] lookup_add) qed lemma tt_monomial_plus: assumes "p \ (0::_ \\<^sub>0 'b::comm_monoid_add)" and "lt p \\<^sub>t v" shows "tt (monomial c v + p) = tt p" proof (cases "c = 0") case True thus ?thesis by (simp add: monomial_0I) next case False have eq: "tt (monomial c v) = v" by (simp only: tt_monomial[OF \c \ 0\]) moreover have "tt (p + monomial c v) = tt p" proof (rule tt_plus_eqI, fact, simp only: eq) from lt_ge_tt[of p] assms(2) show "tt p \\<^sub>t v" by simp qed ultimately show ?thesis by (simp add: ac_simps) qed lemma tc_monomial_plus: assumes "p \ (0::_ \\<^sub>0 'b::comm_monoid_add)" and "lt p \\<^sub>t v" shows "tc (monomial c v + p) = tc p" proof (simp add: tc_def tt_monomial_plus[OF assms] lookup_add lookup_single Poly_Mapping.when_def, rule impI) assume "v = tt p" with assms(2) have "lt p \\<^sub>t tt p" by simp with lt_ge_tt[of p] show "c + lookup p (tt p) = lookup p (tt p)" by simp qed lemma tail_monomial_plus: assumes "c \ (0::'b::comm_monoid_add)" and "lt p \\<^sub>t v" shows "tail (monomial c v + p) = p" (is "tail ?q = _") proof - from assms have "lt ?q = v" by (rule lt_monomial_plus) moreover have "lower (monomial c v) v = 0" by (simp add: lower_eq_zero_iff', rule disjI2, simp add: tt_monomial[OF \c \ 0\]) ultimately show ?thesis by (simp add: tail_def lower_plus lower_id_iff, intro disjI2 assms(2)) qed subsection \Lists of Keys\ text \In algorithms one very often needs to compute the sorted list of all terms appearing in a list of polynomials.\ definition pps_to_list :: "'t set \ 't list" where "pps_to_list S = rev (ord_term_lin.sorted_list_of_set S)" definition keys_to_list :: "('t \\<^sub>0 'b::zero) \ 't list" where "keys_to_list p = pps_to_list (keys p)" definition Keys_to_list :: "('t \\<^sub>0 'b::zero) list \ 't list" where "Keys_to_list ps = fold (\p ts. merge_wrt (\\<^sub>t) (keys_to_list p) ts) ps []" text \Function @{const pps_to_list} turns finite sets of terms into sorted lists, where the lists are sorted descending (i.\,e. greater elements come before smaller ones).\ lemma distinct_pps_to_list: "distinct (pps_to_list S)" unfolding pps_to_list_def distinct_rev by (rule ord_term_lin.distinct_sorted_list_of_set) lemma set_pps_to_list: assumes "finite S" shows "set (pps_to_list S) = S" unfolding pps_to_list_def set_rev using assms by simp lemma length_pps_to_list: "length (pps_to_list S) = card S" proof (cases "finite S") case True from distinct_card[OF distinct_pps_to_list] have "length (pps_to_list S) = card (set (pps_to_list S))" by simp also from True have "... = card S" by (simp only: set_pps_to_list) finally show ?thesis . next case False thus ?thesis by (simp add: pps_to_list_def) qed lemma pps_to_list_sorted_wrt: "sorted_wrt (\\<^sub>t) (pps_to_list S)" proof - have "sorted_wrt (\\<^sub>t) (pps_to_list S)" proof - have tr: "transp (\\<^sub>t)" using transp_def by fastforce have *: "(\x y. y \\<^sub>t x) = (\\<^sub>t)" by simp show ?thesis - by (simp only: * pps_to_list_def sorted_wrt_rev ord_term_lin.sorted_sorted_wrt[symmetric], + by (simp only: * pps_to_list_def sorted_wrt_rev, rule ord_term_lin.sorted_sorted_list_of_set) qed with distinct_pps_to_list have "sorted_wrt (\x y. x \\<^sub>t y \ x \ y) (pps_to_list S)" by (rule distinct_sorted_wrt_imp_sorted_wrt_strict) moreover have "(\\<^sub>t) = (\x y. x \\<^sub>t y \ x \ y)" using ord_term_lin.dual_order.order_iff_strict by auto ultimately show ?thesis by simp qed lemma pps_to_list_nth_leI: assumes "j \ i" and "i < card S" shows "(pps_to_list S) ! i \\<^sub>t (pps_to_list S) ! j" proof (cases "j = i") case True show ?thesis by (simp add: True) next case False with assms(1) have "j < i" by simp let ?ts = "pps_to_list S" from pps_to_list_sorted_wrt \j < i\ have "(\\<^sub>t)\\ (?ts ! j) (?ts ! i)" proof (rule sorted_wrt_nth_less) from assms(2) show "i < length ?ts" by (simp only: length_pps_to_list) qed thus ?thesis by simp qed lemma pps_to_list_nth_lessI: assumes "j < i" and "i < card S" shows "(pps_to_list S) ! i \\<^sub>t (pps_to_list S) ! j" proof - let ?ts = "pps_to_list S" from assms(1) have "j \ i" and "i \ j" by simp_all with assms(2) have "i < length ?ts" and "j < length ?ts" by (simp_all only: length_pps_to_list) show ?thesis proof (rule ord_term_lin.neq_le_trans) from \i \ j\ show "?ts ! i \ ?ts ! j" by (simp add: nth_eq_iff_index_eq[OF distinct_pps_to_list \i < length ?ts\ \j < length ?ts\]) next from \j \ i\ assms(2) show "?ts ! i \\<^sub>t ?ts ! j" by (rule pps_to_list_nth_leI) qed qed lemma pps_to_list_nth_leD: assumes "(pps_to_list S) ! i \\<^sub>t (pps_to_list S) ! j" and "j < card S" shows "j \ i" proof (rule ccontr) assume "\ j \ i" hence "i < j" by simp from this \j < card S\ have "(pps_to_list S) ! j \\<^sub>t (pps_to_list S) ! i" by (rule pps_to_list_nth_lessI) with assms(1) show False by simp qed lemma pps_to_list_nth_lessD: assumes "(pps_to_list S) ! i \\<^sub>t (pps_to_list S) ! j" and "j < card S" shows "j < i" proof (rule ccontr) assume "\ j < i" hence "i \ j" by simp from this \j < card S\ have "(pps_to_list S) ! j \\<^sub>t (pps_to_list S) ! i" by (rule pps_to_list_nth_leI) with assms(1) show False by simp qed lemma set_keys_to_list: "set (keys_to_list p) = keys p" by (simp add: keys_to_list_def set_pps_to_list) lemma length_keys_to_list: "length (keys_to_list p) = card (keys p)" by (simp only: keys_to_list_def length_pps_to_list) lemma keys_to_list_zero [simp]: "keys_to_list 0 = []" by (simp add: keys_to_list_def pps_to_list_def) lemma Keys_to_list_Nil [simp]: "Keys_to_list [] = []" by (simp add: Keys_to_list_def) lemma set_Keys_to_list: "set (Keys_to_list ps) = Keys (set ps)" proof - have "set (Keys_to_list ps) = (\p\set ps. set (keys_to_list p)) \ set []" unfolding Keys_to_list_def by (rule set_fold, simp only: set_merge_wrt) also have "... = Keys (set ps)" by (simp add: Keys_def set_keys_to_list) finally show ?thesis . qed lemma Keys_to_list_sorted_wrt_aux: assumes "sorted_wrt (\\<^sub>t) ts" shows "sorted_wrt (\\<^sub>t) (fold (\p ts. merge_wrt (\\<^sub>t) (keys_to_list p) ts) ps ts)" using assms proof (induct ps arbitrary: ts) case Nil thus ?case by simp next case (Cons p ps) show ?case proof (simp only: fold.simps o_def, rule Cons(1), rule sorted_merge_wrt) show "transp (\\<^sub>t)" unfolding transp_def by fastforce next fix x y :: 't assume "x \ y" thus "x \\<^sub>t y \ y \\<^sub>t x" by auto next show "sorted_wrt (\\<^sub>t) (keys_to_list p)" unfolding keys_to_list_def by (fact pps_to_list_sorted_wrt) qed fact qed corollary Keys_to_list_sorted_wrt: "sorted_wrt (\\<^sub>t) (Keys_to_list ps)" unfolding Keys_to_list_def proof (rule Keys_to_list_sorted_wrt_aux) show "sorted_wrt (\\<^sub>t) []" by simp qed corollary distinct_Keys_to_list: "distinct (Keys_to_list ps)" proof (rule distinct_sorted_wrt_irrefl) show "irreflp (\\<^sub>t)" by (simp add: irreflp_def) next show "transp (\\<^sub>t)" unfolding transp_def by fastforce next show "sorted_wrt (\\<^sub>t) (Keys_to_list ps)" by (fact Keys_to_list_sorted_wrt) qed lemma length_Keys_to_list: "length (Keys_to_list ps) = card (Keys (set ps))" proof - from distinct_Keys_to_list have "card (set (Keys_to_list ps)) = length (Keys_to_list ps)" by (rule distinct_card) thus ?thesis by (simp only: set_Keys_to_list) qed lemma Keys_to_list_eq_pps_to_list: "Keys_to_list ps = pps_to_list (Keys (set ps))" using _ Keys_to_list_sorted_wrt distinct_Keys_to_list pps_to_list_sorted_wrt distinct_pps_to_list proof (rule sorted_wrt_distinct_set_unique) show "antisymp (\\<^sub>t)" unfolding antisymp_def by fastforce next from finite_set have fin: "finite (Keys (set ps))" by (rule finite_Keys) show "set (Keys_to_list ps) = set (pps_to_list (Keys (set ps)))" by (simp add: set_Keys_to_list set_pps_to_list[OF fin]) qed subsection \Multiplication\ lemma in_keys_mult_scalar_le: assumes "v \ keys (p \ q)" shows "v \\<^sub>t punit.lt p \ lt q" proof - from assms obtain t u where "t \ keys p" and "u \ keys q" and "v = t \ u" by (rule in_keys_mult_scalarE) from \t \ keys p\ have "t \ punit.lt p" by (rule punit.lt_max_keys) from \u \ keys q\ have "u \\<^sub>t lt q" by (rule lt_max_keys) hence "v \\<^sub>t t \ lt q" unfolding \v = t \ u\ by (rule splus_mono) also from \t \ punit.lt p\ have "... \\<^sub>t punit.lt p \ lt q" by (rule splus_mono_left) finally show ?thesis . qed lemma in_keys_mult_scalar_ge: assumes "v \ keys (p \ q)" shows "punit.tt p \ tt q \\<^sub>t v" proof - from assms obtain t u where "t \ keys p" and "u \ keys q" and "v = t \ u" by (rule in_keys_mult_scalarE) from \t \ keys p\ have "punit.tt p \ t" by (rule punit.tt_min_keys) from \u \ keys q\ have "tt q \\<^sub>t u" by (rule tt_min_keys) hence "punit.tt p \ tt q \\<^sub>t punit.tt p \ u" by (rule splus_mono) also from \punit.tt p \ t\ have "... \\<^sub>t v" unfolding \v = t \ u\ by (rule splus_mono_left) finally show ?thesis . qed lemma (in ordered_term) lookup_mult_scalar_lt_lt: "lookup (p \ q) (punit.lt p \ lt q) = punit.lc p * lc q" proof (induct p rule: punit.poly_mapping_tail_induct) case 0 show ?case by simp next case step: (tail p) from step(1) have "punit.lc p \ 0" by (rule punit.lc_not_0) let ?t = "punit.lt p \ lt q" show ?case proof (cases "is_monomial p") case True then obtain c t where "c \ 0" and "punit.lt p = t" and "punit.lc p = c" and p_eq: "p = monomial c t" by (rule punit.is_monomial_monomial_ordered) hence "p \ q = monom_mult (punit.lc p) (punit.lt p) q" by (simp add: mult_scalar_monomial) thus ?thesis by (simp add: lookup_monom_mult_plus lc_def) next case False have "punit.lt (punit.tail p) \ punit.lt p" proof (simp add: punit.tail_def punit.lt_lower_eq_iff, rule) assume "punit.lt p = 0" have "keys p \ {punit.lt p}" proof (rule, simp) fix s assume "s \ keys p" hence "s \ punit.lt p" by (rule punit.lt_max_keys) moreover have "punit.lt p \ s" unfolding \punit.lt p = 0\ by (rule zero_min) ultimately show "s = punit.lt p" by simp qed hence "card (keys p) = 0 \ card (keys p) = 1" using subset_singletonD by fastforce thus False proof assume "card (keys p) = 0" hence "p = 0" by (meson card_0_eq keys_eq_empty finite_keys) with step(1) show False .. next assume "card (keys p) = 1" with False show False unfolding is_monomial_def .. qed qed with punit.lt_lower[of p "punit.lt p"] have "punit.lt (punit.tail p) \ punit.lt p" by (simp add: punit.tail_def) have eq: "lookup ((punit.tail p) \ q) ?t = 0" proof (rule ccontr) assume "lookup ((punit.tail p) \ q) ?t \ 0" hence "?t \\<^sub>t punit.lt (punit.tail p) \ lt q" by (meson in_keys_mult_scalar_le lookup_not_eq_zero_eq_in_keys) hence "punit.lt p \ punit.lt (punit.tail p)" by (rule ord_term_canc_left) also have "... \ punit.lt p" by fact finally show False .. qed from step(2) have "lookup (monom_mult (punit.lc p) (punit.lt p) q) ?t = punit.lc p * lc q" by (simp only: lookup_monom_mult_plus lc_def) thus ?thesis by (simp add: mult_scalar_tail_rec_left[of p q] lookup_add eq) qed qed lemma lookup_mult_scalar_tt_tt: "lookup (p \ q) (punit.tt p \ tt q) = punit.tc p * tc q" proof (induct p rule: punit.poly_mapping_tail_induct) case 0 show ?case by simp next case step: (tail p) from step(1) have "punit.lc p \ 0" by (rule punit.lc_not_0) let ?t = "punit.tt p \ tt q" show ?case proof (cases "is_monomial p") case True then obtain c t where "c \ 0" and "punit.lt p = t" and "punit.lc p = c" and p_eq: "p = monomial c t" by (rule punit.is_monomial_monomial_ordered) from \c \ 0\ have "punit.tt p = t" and "punit.tc p = c" by (simp_all add: p_eq punit.tt_monomial) with p_eq have "p \ q = monom_mult (punit.tc p) (punit.tt p) q" by (simp add: mult_scalar_monomial) thus ?thesis by (simp add: lookup_monom_mult_plus tc_def) next case False from step(1) have "keys p \ {}" by simp with finite_keys have "card (keys p) \ 0" by auto with False have "2 \ card (keys p)" unfolding is_monomial_def by linarith then obtain s t where "s \ keys p" and "t \ keys p" and "s \ t" by (metis (mono_tags, lifting) card.empty card.infinite card_insert_disjoint card_mono empty_iff finite.emptyI insertCI lessI not_less numeral_2_eq_2 ordered_powerprod_lin.infinite_growing ordered_powerprod_lin.neqE preorder_class.less_le_trans subsetI) from this(1) this(3) have "punit.tt p \ t" by (rule punit.tt_less) also from \t \ keys p\ have "t \ punit.lt p" by (rule punit.lt_max_keys) finally have "punit.tt p \ punit.lt p" . hence tt_tail: "punit.tt (punit.tail p) = punit.tt p" and tc_tail: "punit.tc (punit.tail p) = punit.tc p" unfolding punit.tail_def by (rule punit.tt_lower, rule punit.tc_lower) have eq: "lookup (monom_mult (punit.lc p) (punit.lt p) q) ?t = 0" proof (rule ccontr) assume "lookup (monom_mult (punit.lc p) (punit.lt p) q) ?t \ 0" hence "punit.lt p \ tt q \\<^sub>t ?t" by (meson in_keys_iff in_keys_monom_mult_ge) hence "punit.lt p \ punit.tt p" by (rule ord_term_canc_left) also have "... \ punit.lt p" by fact finally show False .. qed from step(2) have "lookup (punit.tail p \ q) ?t = punit.tc p * tc q" by (simp only: tt_tail tc_tail) thus ?thesis by (simp add: mult_scalar_tail_rec_left[of p q] lookup_add eq) qed qed lemma lt_mult_scalar: assumes "p \ 0" and "q \ (0::'t \\<^sub>0 'b::semiring_no_zero_divisors)" shows "lt (p \ q) = punit.lt p \ lt q" proof (rule lt_eqI_keys, simp only: in_keys_iff lookup_mult_scalar_lt_lt) from assms(1) have "punit.lc p \ 0" by (rule punit.lc_not_0) moreover from assms(2) have "lc q \ 0" by (rule lc_not_0) ultimately show "punit.lc p * lc q \ 0" by simp qed (rule in_keys_mult_scalar_le) lemma tt_mult_scalar: assumes "p \ 0" and "q \ (0::'t \\<^sub>0 'b::semiring_no_zero_divisors)" shows "tt (p \ q) = punit.tt p \ tt q" proof (rule tt_eqI, simp only: in_keys_iff lookup_mult_scalar_tt_tt) from assms(1) have "punit.tc p \ 0" by (rule punit.tc_not_0) moreover from assms(2) have "tc q \ 0" by (rule tc_not_0) ultimately show "punit.tc p * tc q \ 0" by simp qed (rule in_keys_mult_scalar_ge) lemma lc_mult_scalar: "lc (p \ q) = punit.lc p * lc (q::'t \\<^sub>0 'b::semiring_no_zero_divisors)" proof (cases "p = 0") case True thus ?thesis by (simp add: lc_def) next case False show ?thesis proof (cases "q = 0") case True thus ?thesis by (simp add: lc_def) next case False with \p \ 0\ show ?thesis by (simp add: lc_def lt_mult_scalar lookup_mult_scalar_lt_lt) qed qed lemma tc_mult_scalar: "tc (p \ q) = punit.tc p * tc (q::'t \\<^sub>0 'b::semiring_no_zero_divisors)" proof (cases "p = 0") case True thus ?thesis by (simp add: tc_def) next case False show ?thesis proof (cases "q = 0") case True thus ?thesis by (simp add: tc_def) next case False with \p \ 0\ show ?thesis by (simp add: tc_def tt_mult_scalar lookup_mult_scalar_tt_tt) qed qed lemma mult_scalar_not_zero: assumes "p \ 0" and "q \ (0::'t \\<^sub>0 'b::semiring_no_zero_divisors)" shows "p \ q \ 0" proof assume "p \ q = 0" hence "0 = lc (p \ q)" by (simp add: lc_def) also have "... = punit.lc p * lc q" by (rule lc_mult_scalar) finally have "punit.lc p * lc q = 0" by simp moreover from assms(1) have "punit.lc p \ 0" by (rule punit.lc_not_0) moreover from assms(2) have "lc q \ 0" by (rule lc_not_0) ultimately show False by simp qed end (* ordered_term_powerprod *) context ordered_powerprod begin lemmas in_keys_times_le = punit.in_keys_mult_scalar_le[simplified] lemmas in_keys_times_ge = punit.in_keys_mult_scalar_ge[simplified] lemmas lookup_times_lp_lp = punit.lookup_mult_scalar_lt_lt[simplified] lemmas lookup_times_tp_tp = punit.lookup_mult_scalar_tt_tt[simplified] lemmas lookup_times_monomial_right_plus = punit.lookup_mult_scalar_monomial_right_plus[simplified] lemmas lookup_times_monomial_right = punit.lookup_mult_scalar_monomial_right[simplified] lemmas lp_times = punit.lt_mult_scalar[simplified] lemmas tp_times = punit.tt_mult_scalar[simplified] lemmas lc_times = punit.lc_mult_scalar[simplified] lemmas tc_times = punit.tc_mult_scalar[simplified] lemmas times_not_zero = punit.mult_scalar_not_zero[simplified] lemmas times_tail_rec_left = punit.mult_scalar_tail_rec_left[simplified] lemmas times_tail_rec_right = punit.mult_scalar_tail_rec_right[simplified] lemmas punit_in_keys_monom_mult_le = punit.in_keys_monom_mult_le[simplified] lemmas punit_in_keys_monom_mult_ge = punit.in_keys_monom_mult_ge[simplified] lemmas lp_monom_mult = punit.lt_monom_mult[simplified] lemmas tp_monom_mult = punit.tt_monom_mult[simplified] end subsection \@{term dgrad_p_set} and @{term dgrad_p_set_le}\ locale gd_term = ordered_term pair_of_term term_of_pair ord ord_strict ord_term ord_term_strict for pair_of_term::"'t \ ('a::graded_dickson_powerprod \ 'k::{the_min,wellorder})" and term_of_pair::"('a \ 'k) \ 't" and ord::"'a \ 'a \ bool" (infixl "\" 50) and ord_strict (infixl "\" 50) and ord_term::"'t \ 't \ bool" (infixl "\\<^sub>t" 50) and ord_term_strict::"'t \ 't \ bool" (infixl "\\<^sub>t" 50) begin sublocale gd_powerprod .. lemma adds_term_antisym: assumes "u adds\<^sub>t v" and "v adds\<^sub>t u" shows "u = v" using assms unfolding adds_term_def using adds_antisym by (metis term_of_pair_pair) definition dgrad_p_set :: "('a \ nat) \ nat \ ('t \\<^sub>0 'b::zero) set" where "dgrad_p_set d m = {p. pp_of_term ` keys p \ dgrad_set d m}" definition dgrad_p_set_le :: "('a \ nat) \ (('t \\<^sub>0 'b) set) \ (('t \\<^sub>0 'b::zero) set) \ bool" where "dgrad_p_set_le d F G \ (dgrad_set_le d (pp_of_term ` Keys F) (pp_of_term ` Keys G))" lemma in_dgrad_p_set_iff: "p \ dgrad_p_set d m \ (\v\keys p. d (pp_of_term v) \ m)" by (auto simp add: dgrad_p_set_def dgrad_set_def) lemma dgrad_p_setI [intro]: assumes "\v. v \ keys p \ d (pp_of_term v) \ m" shows "p \ dgrad_p_set d m" using assms by (auto simp: in_dgrad_p_set_iff) lemma dgrad_p_setD: assumes "p \ dgrad_p_set d m" and "v \ keys p" shows "d (pp_of_term v) \ m" using assms by (simp only: in_dgrad_p_set_iff) lemma zero_in_dgrad_p_set: "0 \ dgrad_p_set d m" by (rule, simp) lemma dgrad_p_set_zero [simp]: "dgrad_p_set (\_. 0) m = UNIV" by auto lemma subset_dgrad_p_set_zero: "F \ dgrad_p_set (\_. 0) m" by simp lemma dgrad_p_set_subset: assumes "m \ n" shows "dgrad_p_set d m \ dgrad_p_set d n" using assms by (auto simp: dgrad_p_set_def dgrad_set_def) lemma dgrad_p_setD_lp: assumes "p \ dgrad_p_set d m" and "p \ 0" shows "d (lp p) \ m" by (rule dgrad_p_setD, fact, rule lt_in_keys, fact) lemma dgrad_p_set_exhaust_expl: assumes "finite F" shows "F \ dgrad_p_set d (Max (d ` pp_of_term ` Keys F))" proof fix f assume "f \ F" from assms have "finite (Keys F)" by (rule finite_Keys) have fin: "finite (d ` pp_of_term ` Keys F)" by (intro finite_imageI, fact) show "f \ dgrad_p_set d (Max (d ` pp_of_term ` Keys F))" proof (rule dgrad_p_setI) fix v assume "v \ keys f" from this \f \ F\ have "v \ Keys F" by (rule in_KeysI) hence "d (pp_of_term v) \ d ` pp_of_term ` Keys F" by simp with fin show "d (pp_of_term v) \ Max (d ` pp_of_term ` Keys F)" by (rule Max_ge) qed qed lemma dgrad_p_set_exhaust: assumes "finite F" obtains m where "F \ dgrad_p_set d m" proof from assms show "F \ dgrad_p_set d (Max (d ` pp_of_term ` Keys F))" by (rule dgrad_p_set_exhaust_expl) qed lemma dgrad_p_set_insert: assumes "F \ dgrad_p_set d m" obtains n where "m \ n" and "f \ dgrad_p_set d n" and "F \ dgrad_p_set d n" proof - have "finite {f}" by simp then obtain m1 where "{f} \ dgrad_p_set d m1" by (rule dgrad_p_set_exhaust) hence "f \ dgrad_p_set d m1" by simp define n where "n = ord_class.max m m1" have "m \ n" and "m1 \ n" by (simp_all add: n_def) from this(1) show ?thesis proof from \m1 \ n\ have "dgrad_p_set d m1 \ dgrad_p_set d n" by (rule dgrad_p_set_subset) with \f \ dgrad_p_set d m1\ show "f \ dgrad_p_set d n" .. next from \m \ n\ have "dgrad_p_set d m \ dgrad_p_set d n" by (rule dgrad_p_set_subset) with assms show "F \ dgrad_p_set d n" by (rule subset_trans) qed qed lemma dgrad_p_set_leI: assumes "\f. f \ F \ dgrad_p_set_le d {f} G" shows "dgrad_p_set_le d F G" unfolding dgrad_p_set_le_def dgrad_set_le_def proof fix s assume "s \ pp_of_term ` Keys F" then obtain v where "v \ Keys F" and "s = pp_of_term v" .. from this(1) obtain f where "f \ F" and "v \ keys f" by (rule in_KeysE) from this(2) have "s \ pp_of_term ` Keys {f}" by (simp add: \s = pp_of_term v\ Keys_insert) from \f \ F\ have "dgrad_p_set_le d {f} G" by (rule assms) from this \s \ pp_of_term ` Keys {f}\ show "\t\pp_of_term ` Keys G. d s \ d t" unfolding dgrad_p_set_le_def dgrad_set_le_def .. qed lemma dgrad_p_set_le_trans [trans]: assumes "dgrad_p_set_le d F G" and "dgrad_p_set_le d G H" shows "dgrad_p_set_le d F H" using assms unfolding dgrad_p_set_le_def by (rule dgrad_set_le_trans) lemma dgrad_p_set_le_subset: assumes "F \ G" shows "dgrad_p_set_le d F G" unfolding dgrad_p_set_le_def by (rule dgrad_set_le_subset, rule image_mono, rule Keys_mono, fact) lemma dgrad_p_set_leI_insert_keys: assumes "dgrad_p_set_le d F G" and "dgrad_set_le d (pp_of_term ` keys f) (pp_of_term ` Keys G)" shows "dgrad_p_set_le d (insert f F) G" using assms by (simp add: dgrad_p_set_le_def Keys_insert dgrad_set_le_Un image_Un) lemma dgrad_p_set_leI_insert: assumes "dgrad_p_set_le d F G" and "dgrad_p_set_le d {f} G" shows "dgrad_p_set_le d (insert f F) G" using assms by (simp add: dgrad_p_set_le_def Keys_insert dgrad_set_le_Un image_Un) lemma dgrad_p_set_leI_Un: assumes "dgrad_p_set_le d F1 G" and "dgrad_p_set_le d F2 G" shows "dgrad_p_set_le d (F1 \ F2) G" using assms by (auto simp: dgrad_p_set_le_def dgrad_set_le_def Keys_Un) lemma dgrad_p_set_le_dgrad_p_set: assumes "dgrad_p_set_le d F G" and "G \ dgrad_p_set d m" shows "F \ dgrad_p_set d m" proof fix f assume "f \ F" show "f \ dgrad_p_set d m" proof (rule dgrad_p_setI) fix v assume "v \ keys f" from this \f \ F\ have "v \ Keys F" by (rule in_KeysI) hence "pp_of_term v \ pp_of_term ` Keys F" by simp with assms(1) obtain s where "s \ pp_of_term ` Keys G" and "d (pp_of_term v) \ d s" unfolding dgrad_p_set_le_def by (rule dgrad_set_leE) from this(1) obtain u where "u \ Keys G" and s: "s = pp_of_term u" .. from this(1) obtain g where "g \ G" and "u \ keys g" by (rule in_KeysE) from this(1) assms(2) have "g \ dgrad_p_set d m" .. from this \u \ keys g\ have "d s \ m" unfolding s by (rule dgrad_p_setD) with \d (pp_of_term v) \ d s\ show "d (pp_of_term v) \ m" by (rule le_trans) qed qed lemma dgrad_p_set_le_except: "dgrad_p_set_le d {except p S} {p}" by (auto simp add: dgrad_p_set_le_def Keys_insert keys_except intro: dgrad_set_le_subset) lemma dgrad_p_set_le_tail: "dgrad_p_set_le d {tail p} {p}" by (simp only: tail_def lower_def, fact dgrad_p_set_le_except) lemma dgrad_p_set_le_plus: "dgrad_p_set_le d {p + q} {p, q}" by (simp add: dgrad_p_set_le_def Keys_insert, rule dgrad_set_le_subset, rule image_mono, fact Poly_Mapping.keys_add) lemma dgrad_p_set_le_uminus: "dgrad_p_set_le d {-p} {p}" by (simp add: dgrad_p_set_le_def Keys_insert keys_uminus, fact dgrad_set_le_refl) lemma dgrad_p_set_le_minus: "dgrad_p_set_le d {p - q} {p, q}" by (simp add: dgrad_p_set_le_def Keys_insert, rule dgrad_set_le_subset, rule image_mono, fact keys_minus) lemma dgrad_set_le_monom_mult: assumes "dickson_grading d" shows "dgrad_set_le d (pp_of_term ` keys (monom_mult c t p)) (insert t (pp_of_term ` keys p))" proof (rule dgrad_set_leI) fix s assume "s \ pp_of_term ` keys (monom_mult c t p)" with keys_monom_mult_subset have "s \ pp_of_term ` ((\) t ` keys p)" by fastforce then obtain v where "v \ keys p" and s: "s = pp_of_term (t \ v)" by fastforce have "d s = ord_class.max (d t) (d (pp_of_term v))" by (simp only: s pp_of_term_splus dickson_gradingD1[OF assms(1)]) hence "d s = d t \ d s = d (pp_of_term v)" by auto thus "\t\insert t (pp_of_term ` keys p). d s \ d t" proof assume "d s = d t" thus ?thesis by simp next assume "d s = d (pp_of_term v)" show ?thesis proof from \d s = d (pp_of_term v)\ show "d s \ d (pp_of_term v)" by simp next from \v \ keys p\ show "pp_of_term v \ insert t (pp_of_term ` keys p)" by simp qed qed qed lemma dgrad_p_set_closed_plus: assumes "p \ dgrad_p_set d m" and "q \ dgrad_p_set d m" shows "p + q \ dgrad_p_set d m" proof - from dgrad_p_set_le_plus have "{p + q} \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set) from assms show "{p, q} \ dgrad_p_set d m" by simp qed thus ?thesis by simp qed lemma dgrad_p_set_closed_uminus: assumes "p \ dgrad_p_set d m" shows "-p \ dgrad_p_set d m" proof - from dgrad_p_set_le_uminus have "{-p} \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set) from assms show "{p} \ dgrad_p_set d m" by simp qed thus ?thesis by simp qed lemma dgrad_p_set_closed_minus: assumes "p \ dgrad_p_set d m" and "q \ dgrad_p_set d m" shows "p - q \ dgrad_p_set d m" proof - from dgrad_p_set_le_minus have "{p - q} \ dgrad_p_set d m" proof (rule dgrad_p_set_le_dgrad_p_set) from assms show "{p, q} \ dgrad_p_set d m" by simp qed thus ?thesis by simp qed lemma dgrad_p_set_closed_monom_mult: assumes "dickson_grading d" and "d t \ m" and "p \ dgrad_p_set d m" shows "monom_mult c t p \ dgrad_p_set d m" proof (rule dgrad_p_setI) fix v assume "v \ keys (monom_mult c t p)" hence "pp_of_term v \ pp_of_term ` keys (monom_mult c t p)" by simp with dgrad_set_le_monom_mult[OF assms(1)] obtain s where "s \ insert t (pp_of_term ` keys p)" and "d (pp_of_term v) \ d s" by (rule dgrad_set_leE) from this(1) have "s = t \ s \ pp_of_term ` keys p" by simp thus "d (pp_of_term v) \ m" proof assume "s = t" with \d (pp_of_term v) \ d s\ assms(2) show ?thesis by simp next assume "s \ pp_of_term ` keys p" then obtain u where "u \ keys p" and "s = pp_of_term u" .. from assms(3) this(1) have "d s \ m" unfolding \s = pp_of_term u\ by (rule dgrad_p_setD) with \d (pp_of_term v) \ d s\ show ?thesis by (rule le_trans) qed qed lemma dgrad_p_set_closed_monom_mult_zero: assumes "p \ dgrad_p_set d m" shows "monom_mult c 0 p \ dgrad_p_set d m" proof (rule dgrad_p_setI) fix v assume "v \ keys (monom_mult c 0 p)" hence "pp_of_term v \ pp_of_term ` keys (monom_mult c 0 p)" by simp then obtain u where "u \ keys (monom_mult c 0 p)" and eq: "pp_of_term v = pp_of_term u" .. from this(1) have "u \ keys p" by (metis keys_monom_multE splus_zero) with assms have "d (pp_of_term u) \ m" by (rule dgrad_p_setD) thus "d (pp_of_term v) \ m" by (simp only: eq) qed lemma dgrad_p_set_closed_except: assumes "p \ dgrad_p_set d m" shows "except p S \ dgrad_p_set d m" by (rule dgrad_p_setI, rule dgrad_p_setD, rule assms, simp add: keys_except) lemma dgrad_p_set_closed_tail: assumes "p \ dgrad_p_set d m" shows "tail p \ dgrad_p_set d m" unfolding tail_def lower_def using assms by (rule dgrad_p_set_closed_except) subsection \Dickson's Lemma for Sequences of Terms\ lemma Dickson_term: assumes "dickson_grading d" and "finite K" shows "almost_full_on (adds\<^sub>t) {t. pp_of_term t \ dgrad_set d m \ component_of_term t \ K}" (is "almost_full_on _ ?A") proof (rule almost_full_onI) fix seq :: "nat \ 't" assume *: "\i. seq i \ ?A" define seq' where "seq' = (\i. (pp_of_term (seq i), component_of_term (seq i)))" have "pp_of_term ` ?A \ {x. d x \ m}" by (auto dest: dgrad_setD) moreover from assms(1) have "almost_full_on (adds) {x. d x \ m}" by (rule dickson_gradingD2) ultimately have "almost_full_on (adds) (pp_of_term ` ?A)" by (rule almost_full_on_subset) moreover have "almost_full_on (=) (component_of_term ` ?A)" proof (rule eq_almost_full_on_finite_set) have "component_of_term ` ?A \ K" by blast thus "finite (component_of_term ` ?A)" using assms(2) by (rule finite_subset) qed ultimately have "almost_full_on (prod_le (adds) (=)) (pp_of_term ` ?A \ component_of_term ` ?A)" by (rule almost_full_on_Sigma) moreover from * have "\i. seq' i \ pp_of_term ` ?A \ component_of_term ` ?A" by (simp add: seq'_def) ultimately obtain i j where "i < j" and "prod_le (adds) (=) (seq' i) (seq' j)" by (rule almost_full_onD) from this(2) have "seq i adds\<^sub>t seq j" by (simp add: seq'_def prod_le_def adds_term_def) with \i < j\ show "good (adds\<^sub>t) seq" by (rule goodI) qed corollary Dickson_termE: assumes "dickson_grading d" and "finite (component_of_term ` range (f::nat \ 't))" and "pp_of_term ` range f \ dgrad_set d m" obtains i j where "i < j" and "f i adds\<^sub>t f j" proof - let ?A = "{t. pp_of_term t \ dgrad_set d m \ component_of_term t \ component_of_term ` range f}" from assms(1, 2) have "almost_full_on (adds\<^sub>t) ?A" by (rule Dickson_term) moreover from assms(3) have "\i. f i \ ?A" by blast ultimately obtain i j where "i < j" and "f i adds\<^sub>t f j" by (rule almost_full_onD) thus ?thesis .. qed lemma ex_finite_adds_term: assumes "dickson_grading d" and "finite (component_of_term ` S)" and "pp_of_term ` S \ dgrad_set d m" obtains T where "finite T" and "T \ S" and "\s. s \ S \ (\t\T. t adds\<^sub>t s)" proof - let ?A = "{t. pp_of_term t \ dgrad_set d m \ component_of_term t \ component_of_term ` S}" have "reflp ((adds\<^sub>t)::'t \ _)" by (simp add: reflp_def adds_term_refl) moreover have "almost_full_on (adds\<^sub>t) S" proof (rule almost_full_on_subset) from assms(3) show "S \ ?A" by blast next from assms(1, 2) show "almost_full_on (adds\<^sub>t) ?A" by (rule Dickson_term) qed ultimately obtain T where "finite T" and "T \ S" and "\s. s \ S \ (\t\T. t adds\<^sub>t s)" by (rule almost_full_on_finite_subsetE, blast) thus ?thesis .. qed subsection \Well-foundedness\ definition dickson_less_v :: "('a \ nat) \ nat \ 't \ 't \ bool" where "dickson_less_v d m v u \ (d (pp_of_term v) \ m \ d (pp_of_term u) \ m \ v \\<^sub>t u)" definition dickson_less_p :: "('a \ nat) \ nat \ ('t \\<^sub>0 'b) \ ('t \\<^sub>0 'b::zero) \ bool" where "dickson_less_p d m p q \ ({p, q} \ dgrad_p_set d m \ p \\<^sub>p q)" lemma dickson_less_vI: assumes "d (pp_of_term v) \ m" and "d (pp_of_term u) \ m" and "v \\<^sub>t u" shows "dickson_less_v d m v u" using assms by (simp add: dickson_less_v_def) lemma dickson_less_vD1: assumes "dickson_less_v d m v u" shows "d (pp_of_term v) \ m" using assms by (simp add: dickson_less_v_def) lemma dickson_less_vD2: assumes "dickson_less_v d m v u" shows "d (pp_of_term u) \ m" using assms by (simp add: dickson_less_v_def) lemma dickson_less_vD3: assumes "dickson_less_v d m v u" shows "v \\<^sub>t u" using assms by (simp add: dickson_less_v_def) lemma dickson_less_v_irrefl: "\ dickson_less_v d m v v" by (simp add: dickson_less_v_def) lemma dickson_less_v_trans: assumes "dickson_less_v d m v u" and "dickson_less_v d m u w" shows "dickson_less_v d m v w" using assms by (auto simp add: dickson_less_v_def) lemma wf_dickson_less_v_aux1: assumes "dickson_grading d" and "\i::nat. dickson_less_v d m (seq (Suc i)) (seq i)" obtains i where "\j. j > i \ component_of_term (seq j) < component_of_term (seq i)" proof - let ?Q = "pp_of_term ` range seq" have "pp_of_term (seq 0) \ ?Q" by simp with wf_dickson_less[OF assms(1)] obtain t where "t \ ?Q" and *: "\s. dickson_less d m s t \ s \ ?Q" by (rule wfE_min[to_pred], blast) from this(1) obtain i where t: "t = pp_of_term (seq i)" by fastforce show ?thesis proof fix j assume "i < j" with _ assms(2) have dlv: "dickson_less_v d m (seq j) (seq i)" proof (rule transp_sequence) from dickson_less_v_trans show "transp (dickson_less_v d m)" by (rule transpI) qed hence "seq j \\<^sub>t seq i" by (rule dickson_less_vD3) define s where "s = pp_of_term (seq j)" have "pp_of_term (seq j) \ ?Q" by simp hence "\ dickson_less d m s t" unfolding s_def using * by blast moreover from dlv have "d s \ m" and "d t \ m" unfolding s_def t by (rule dickson_less_vD1, rule dickson_less_vD2) ultimately have "t \ s" by (simp add: dickson_less_def) show "component_of_term (seq j) < component_of_term (seq i)" proof (rule ccontr, simp only: not_less) assume "component_of_term (seq i) \ component_of_term (seq j)" with \t \ s\ have "seq i \\<^sub>t seq j" unfolding s_def t by (rule ord_termI) moreover from dlv have "seq j \\<^sub>t seq i" by (rule dickson_less_vD3) ultimately show False by simp qed qed qed lemma wf_dickson_less_v_aux2: assumes "dickson_grading d" and "\i::nat. dickson_less_v d m (seq (Suc i)) (seq i)" and "\i::nat. component_of_term (seq i) < k" shows thesis using assms(2, 3) proof (induct k arbitrary: seq thesis rule: less_induct) case (less k) from assms(1) less(2) obtain i where *: "\j. j > i \ component_of_term (seq j) < component_of_term (seq i)" by (rule wf_dickson_less_v_aux1, blast) define seq1 where "seq1 = (\j. seq (Suc (i + j)))" from less(3) show ?case proof (rule less(1)) fix j show "dickson_less_v d m (seq1 (Suc j)) (seq1 j)" by (simp add: seq1_def, fact less(2)) next fix j show "component_of_term (seq1 j) < component_of_term (seq i)" by (simp add: seq1_def *) qed qed lemma wf_dickson_less_v: assumes "dickson_grading d" shows "wfP (dickson_less_v d m)" proof (rule wfP_chain, rule, elim exE) fix seq::"nat \ 't" assume "\i. dickson_less_v d m (seq (Suc i)) (seq i)" hence *: "\i. dickson_less_v d m (seq (Suc i)) (seq i)" .. with assms obtain i where **: "\j. j > i \ component_of_term (seq j) < component_of_term (seq i)" by (rule wf_dickson_less_v_aux1, blast) define seq1 where "seq1 = (\j. seq (Suc (i + j)))" from assms show False proof (rule wf_dickson_less_v_aux2) fix j show "dickson_less_v d m (seq1 (Suc j)) (seq1 j)" by (simp add: seq1_def, fact *) next fix j show "component_of_term (seq1 j) < component_of_term (seq i)" by (simp add: seq1_def **) qed qed lemma dickson_less_v_zero: "dickson_less_v (\_. 0) m = (\\<^sub>t)" by (rule, rule, simp add: dickson_less_v_def) lemma dickson_less_pI: assumes "p \ dgrad_p_set d m" and "q \ dgrad_p_set d m" and "p \\<^sub>p q" shows "dickson_less_p d m p q" using assms by (simp add: dickson_less_p_def) lemma dickson_less_pD1: assumes "dickson_less_p d m p q" shows "p \ dgrad_p_set d m" using assms by (simp add: dickson_less_p_def) lemma dickson_less_pD2: assumes "dickson_less_p d m p q" shows "q \ dgrad_p_set d m" using assms by (simp add: dickson_less_p_def) lemma dickson_less_pD3: assumes "dickson_less_p d m p q" shows "p \\<^sub>p q" using assms by (simp add: dickson_less_p_def) lemma dickson_less_p_irrefl: "\ dickson_less_p d m p p" by (simp add: dickson_less_p_def) lemma dickson_less_p_trans: assumes "dickson_less_p d m p q" and "dickson_less_p d m q r" shows "dickson_less_p d m p r" using assms by (auto simp add: dickson_less_p_def) lemma dickson_less_p_mono: assumes "dickson_less_p d m p q" and "m \ n" shows "dickson_less_p d n p q" proof - from assms(2) have "dgrad_p_set d m \ dgrad_p_set d n" by (rule dgrad_p_set_subset) moreover from assms(1) have "p \ dgrad_p_set d m" and "q \ dgrad_p_set d m" and "p \\<^sub>p q" by (rule dickson_less_pD1, rule dickson_less_pD2, rule dickson_less_pD3) ultimately have "p \ dgrad_p_set d n" and "q \ dgrad_p_set d n" by auto from this \p \\<^sub>p q\ show ?thesis by (rule dickson_less_pI) qed lemma dickson_less_p_zero: "dickson_less_p (\_. 0) m = (\\<^sub>p)" by (rule, rule, simp add: dickson_less_p_def) lemma wf_dickson_less_p_aux: assumes "dickson_grading d" assumes "x \ Q" and "\y\Q. y \ 0 \ (y \ dgrad_p_set d m \ dickson_less_v d m (lt y) u)" shows "\p\Q. (\q\Q. \ dickson_less_p d m q p)" using assms(2) assms(3) proof (induct u arbitrary: x Q rule: wfP_induct[OF wf_dickson_less_v, OF assms(1)]) fix u::'t and x::"'t \\<^sub>0 'b" and Q::"('t \\<^sub>0 'b) set" assume hyp: "\u0. dickson_less_v d m u0 u \ (\x0 Q0::('t \\<^sub>0 'b) set. x0 \ Q0 \ (\y\Q0. y \ 0 \ (y \ dgrad_p_set d m \ dickson_less_v d m (lt y) u0)) \ (\p\Q0. \q\Q0. \ dickson_less_p d m q p))" assume "x \ Q" assume "\y\Q. y \ 0 \ (y \ dgrad_p_set d m \ dickson_less_v d m (lt y) u)" hence bounded: "\y. y \ Q \ y \ 0 \ (y \ dgrad_p_set d m \ dickson_less_v d m (lt y) u)" by auto show "\p\Q. \q\Q. \ dickson_less_p d m q p" proof (cases "0 \ Q") case True show ?thesis proof (rule, rule, rule) fix q::"'t \\<^sub>0 'b" assume "dickson_less_p d m q 0" hence "q \\<^sub>p 0" by (rule dickson_less_pD3) thus False using ord_p_zero_min[of q] by simp next from True show "0 \ Q" . qed next case False define Q1 where "Q1 = {lt p | p. p \ Q}" from \x \ Q\ have "lt x \ Q1" unfolding Q1_def by auto with wf_dickson_less_v[OF assms(1)] obtain v where "v \ Q1" and v_min_1: "\q. dickson_less_v d m q v \ q \ Q1" by (rule wfE_min[to_pred], auto) have v_min: "\q. q \ Q \ \ dickson_less_v d m (lt q) v" proof - fix q assume "q \ Q" hence "lt q \ Q1" unfolding Q1_def by auto thus "\ dickson_less_v d m (lt q) v" using v_min_1 by auto qed from \v \ Q1\ obtain p where "lt p = v" and "p \ Q" unfolding Q1_def by auto hence "p \ 0" using False by auto with \p \ Q\ have "p \ dgrad_p_set d m \ dickson_less_v d m (lt p) u" by (rule bounded) hence "p \ dgrad_p_set d m" and "dickson_less_v d m (lt p) u" by simp_all moreover from this(1) \p \ 0\ have "d (pp_of_term (lt p)) \ m" by (rule dgrad_p_setD_lp) ultimately have "d (pp_of_term v) \ m" by (simp only: \lt p = v\) define Q2 where "Q2 = {tail p | p. p \ Q \ lt p = v}" from \p \ Q\ \lt p = v\ have "tail p \ Q2" unfolding Q2_def by auto have "\q\Q2. q \ 0 \ (q \ dgrad_p_set d m \ dickson_less_v d m (lt q) (lt p))" proof (intro ballI impI) fix q assume "q \ Q2" then obtain q0 where q: "q = tail q0" and "lt q0 = lt p" and "q0 \ Q" using \lt p = v\ by (auto simp add: Q2_def) assume "q \ 0" hence "tail q0 \ 0" using \q = tail q0\ by simp hence "q0 \ 0" by auto with \q0 \ Q\ have "q0 \ dgrad_p_set d m \ dickson_less_v d m (lt q0) u" by (rule bounded) hence "q0 \ dgrad_p_set d m" and "dickson_less_v d m (lt q0) u" by simp_all from this(1) have "q \ dgrad_p_set d m" unfolding q by (rule dgrad_p_set_closed_tail) show "q \ dgrad_p_set d m \ dickson_less_v d m (lt q) (lt p)" proof show "dickson_less_v d m (lt q) (lt p)" proof (rule dickson_less_vI) from \q \ dgrad_p_set d m\ \q \ 0\ show "d (pp_of_term (lt q)) \ m" by (rule dgrad_p_setD_lp) next from \dickson_less_v d m (lt p) u\ show "d (pp_of_term (lt p)) \ m" by (rule dickson_less_vD1) next from lt_tail[OF \tail q0 \ 0\] \q = tail q0\ \lt q0 = lt p\ show "lt q \\<^sub>t lt p" by simp qed qed fact qed with hyp \dickson_less_v d m (lt p) u\ \tail p \ Q2\ have "\p\Q2. \q\Q2. \ dickson_less_p d m q p" by blast then obtain q where "q \ Q2" and q_min: "\r\Q2. \ dickson_less_p d m r q" .. from \q \ Q2\ obtain q0 where "q = tail q0" and "q0 \ Q" and "lt q0 = v" unfolding Q2_def by auto from q_min \q = tail q0\ have q0_tail_min: "\r. r \ Q2 \ \ dickson_less_p d m r (tail q0)" by simp from \q0 \ Q\ show ?thesis proof show "\r\Q. \ dickson_less_p d m r q0" proof (intro ballI notI) fix r assume "dickson_less_p d m r q0" hence "r \ dgrad_p_set d m" and "q0 \ dgrad_p_set d m" and "r \\<^sub>p q0" by (rule dickson_less_pD1, rule dickson_less_pD2, rule dickson_less_pD3) from this(3) have "lt r \\<^sub>t lt q0" by (simp add: ord_p_lt) with \lt q0 = v\ have "lt r \\<^sub>t v" by simp assume "r \ Q" hence "\ dickson_less_v d m (lt r) v" by (rule v_min) from False \r \ Q\ have "r \ 0" using False by blast with \r \ dgrad_p_set d m\ have "d (pp_of_term (lt r)) \ m" by (rule dgrad_p_setD_lp) have "\ lt r \\<^sub>t v" proof assume "lt r \\<^sub>t v" with \d (pp_of_term (lt r)) \ m\ \d (pp_of_term v) \ m\ have "dickson_less_v d m (lt r) v" by (rule dickson_less_vI) with \\ dickson_less_v d m (lt r) v\ show False .. qed with \lt r \\<^sub>t v\ have "lt r = v" by simp with \r \ Q\ have "tail r \ Q2" by (auto simp add: Q2_def) have "dickson_less_p d m (tail r) (tail q0)" proof (rule dickson_less_pI) show "tail r \ dgrad_p_set d m" by (rule dgrad_p_set_closed_tail, fact) next show "tail q0 \ dgrad_p_set d m" by (rule dgrad_p_set_closed_tail, fact) next have "lt r = lt q0" by (simp only: \lt r = v\ \lt q0 = v\) from \r \ 0\ this \r \\<^sub>p q0\ show "tail r \\<^sub>p tail q0" by (rule ord_p_tail) qed with q0_tail_min[OF \tail r \ Q2\] show False .. qed qed qed qed theorem wf_dickson_less_p: assumes "dickson_grading d" shows "wfP (dickson_less_p d m)" proof (rule wfI_min[to_pred]) fix Q::"('t \\<^sub>0 'b) set" and x assume "x \ Q" show "\z\Q. \y. dickson_less_p d m y z \ y \ Q" proof (cases "0 \ Q") case True show ?thesis proof (rule, rule, rule) from True show "0 \ Q" . next fix q::"'t \\<^sub>0 'b" assume "dickson_less_p d m q 0" hence "q \\<^sub>p 0" by (rule dickson_less_pD3) thus "q \ Q" using ord_p_zero_min[of q] by simp qed next case False show ?thesis proof (cases "Q \ dgrad_p_set d m") case True let ?L = "lt ` Q" from \x \ Q\ have "lt x \ ?L" by simp with wf_dickson_less_v[OF assms] obtain v where "v \ ?L" and v_min: "\u. dickson_less_v d m u v \ u \ ?L" by (rule wfE_min[to_pred], blast) from this(1) obtain x1 where "x1 \ Q" and "v = lt x1" .. from this(1) True False have "x1 \ dgrad_p_set d m" and "x1 \ 0" by auto hence "d (pp_of_term v) \ m" unfolding \v = lt x1\ by (rule dgrad_p_setD_lp) define Q1 where "Q1 = {tail p | p. p \ Q \ lt p = v}" from \x1 \ Q\ have "tail x1 \ Q1" by (auto simp add: Q1_def \v = lt x1\) with assms have "\p\Q1. \q\Q1. \ dickson_less_p d m q p" proof (rule wf_dickson_less_p_aux) show "\y\Q1. y \ 0 \ y \ dgrad_p_set d m \ dickson_less_v d m (lt y) v" proof (intro ballI impI) fix y assume "y \ Q1" and "y \ 0" from this(1) obtain y1 where "y1 \ Q" and "v = lt y1" and "y = tail y1" unfolding Q1_def by blast from this(1) True have "y1 \ dgrad_p_set d m" .. hence "y \ dgrad_p_set d m" unfolding \y = tail y1\ by (rule dgrad_p_set_closed_tail) thus "y \ dgrad_p_set d m \ dickson_less_v d m (lt y) v" proof show "dickson_less_v d m (lt y) v" proof (rule dickson_less_vI) from \y \ dgrad_p_set d m\ \y \ 0\ show "d (pp_of_term (lt y)) \ m" by (rule dgrad_p_setD_lp) next from \y \ 0\ show "lt y \\<^sub>t v" unfolding \v = lt y1\ \y = tail y1\ by (rule lt_tail) qed fact qed qed qed then obtain p0 where "p0 \ Q1" and p0_min: "\q. q \ Q1 \ \ dickson_less_p d m q p0" by blast from this(1) obtain p where "p \ Q" and "v = lt p" and "p0 = tail p" unfolding Q1_def by blast from this(1) False have "p \ 0" by blast show ?thesis proof (intro bexI allI impI notI) fix y assume "y \ Q" hence "y \ 0" using False by blast assume "dickson_less_p d m y p" hence "y \ dgrad_p_set d m" and "p \ dgrad_p_set d m" and "y \\<^sub>p p" by (rule dickson_less_pD1, rule dickson_less_pD2, rule dickson_less_pD3) from this(3) have "y \\<^sub>p p" by simp hence "lt y \\<^sub>t lt p" by (rule ord_p_lt) moreover have "\ lt y \\<^sub>t lt p" proof assume "lt y \\<^sub>t lt p" have "dickson_less_v d m (lt y) v" unfolding \v = lt p\ by (rule dickson_less_vI, rule dgrad_p_setD_lp, fact+, rule dgrad_p_setD_lp, fact+) hence "lt y \ ?L" by (rule v_min) hence "y \ Q" by fastforce from this \y \ Q\ show False .. qed ultimately have "lt y = lt p" by simp from \y \ 0\ this \y \\<^sub>p p\ have "tail y \\<^sub>p tail p" by (rule ord_p_tail) from \y \ Q\ have "tail y \ Q1" by (auto simp add: Q1_def \v = lt p\ \lt y = lt p\[symmetric]) hence "\ dickson_less_p d m (tail y) p0" by (rule p0_min) moreover have "dickson_less_p d m (tail y) p0" unfolding \p0 = tail p\ by (rule dickson_less_pI, rule dgrad_p_set_closed_tail, fact, rule dgrad_p_set_closed_tail, fact+) ultimately show False .. qed fact next case False then obtain q where "q \ Q" and "q \ dgrad_p_set d m" by blast from this(1) show ?thesis proof show "\y. dickson_less_p d m y q \ y \ Q" proof (intro allI impI) fix y assume "dickson_less_p d m y q" hence "q \ dgrad_p_set d m" by (rule dickson_less_pD2) with \q \ dgrad_p_set d m\ show "y \ Q" .. qed qed qed qed qed corollary ord_p_minimum_dgrad_p_set: assumes "dickson_grading d" and "x \ Q" and "Q \ dgrad_p_set d m" obtains q where "q \ Q" and "\y. y \\<^sub>p q \ y \ Q" proof - from assms(1) have "wfP (dickson_less_p d m)" by (rule wf_dickson_less_p) from this assms(2) obtain q where "q \ Q" and *: "\y. dickson_less_p d m y q \ y \ Q" by (rule wfE_min[to_pred], auto) from assms(3) \q \ Q\ have "q \ dgrad_p_set d m" .. from \q \ Q\ show ?thesis proof fix y assume "y \\<^sub>p q" show "y \ Q" proof assume "y \ Q" with assms(3) have "y \ dgrad_p_set d m" .. from this \q \ dgrad_p_set d m\ \y \\<^sub>p q\ have "dickson_less_p d m y q" by (rule dickson_less_pI) hence "y \ Q" by (rule *) from this \y \ Q\ show False .. qed qed qed lemma ord_term_minimum_dgrad_set: assumes "dickson_grading d" and "v \ V" and "pp_of_term ` V \ dgrad_set d m" obtains u where "u \ V" and "\w. w \\<^sub>t u \ w \ V" proof - from assms(1) have "wfP (dickson_less_v d m)" by (rule wf_dickson_less_v) then obtain u where "u \ V" and *: "\w. dickson_less_v d m w u \ w \ V" using assms(2) by (rule wfE_min[to_pred]) blast from this(1) have "pp_of_term u \ pp_of_term ` V" by (rule imageI) with assms(3) have "pp_of_term u \ dgrad_set d m" .. hence "d (pp_of_term u) \ m" by (rule dgrad_setD) from \u \ V\ show ?thesis proof fix w assume "w \\<^sub>t u" show "w \ V" proof assume "w \ V" hence "pp_of_term w \ pp_of_term ` V" by (rule imageI) with assms(3) have "pp_of_term w \ dgrad_set d m" .. hence "d (pp_of_term w) \ m" by (rule dgrad_setD) from this \d (pp_of_term u) \ m\ \w \\<^sub>t u\ have "dickson_less_v d m w u" by (rule dickson_less_vI) hence "w \ V" by (rule *) from this \w \ V\ show False .. qed qed qed end (* gd_term *) subsection \More Interpretations\ context gd_powerprod begin sublocale punit: gd_term to_pair_unit fst "(\)" "(\)" "(\)" "(\)" .. end locale od_term = ordered_term pair_of_term term_of_pair ord ord_strict ord_term ord_term_strict for pair_of_term::"'t \ ('a::dickson_powerprod \ 'k::{the_min,wellorder})" and term_of_pair::"('a \ 'k) \ 't" and ord::"'a \ 'a \ bool" (infixl "\" 50) and ord_strict (infixl "\" 50) and ord_term::"'t \ 't \ bool" (infixl "\\<^sub>t" 50) and ord_term_strict::"'t \ 't \ bool" (infixl "\\<^sub>t" 50) begin sublocale gd_term .. lemma ord_p_wf: "wfP (\\<^sub>p)" proof - from dickson_grading_zero have "wfP (dickson_less_p (\_. 0) 0)" by (rule wf_dickson_less_p) thus ?thesis by (simp only: dickson_less_p_zero) qed end (* od_term *) end (* theory *) diff --git a/thys/Refine_Imperative_HOL/IICF/Impl/IICF_Sepl_Binding.thy b/thys/Refine_Imperative_HOL/IICF/Impl/IICF_Sepl_Binding.thy --- a/thys/Refine_Imperative_HOL/IICF/Impl/IICF_Sepl_Binding.thy +++ b/thys/Refine_Imperative_HOL/IICF/Impl/IICF_Sepl_Binding.thy @@ -1,813 +1,813 @@ section \Sepref Bindings for Imp/HOL Collections\ theory IICF_Sepl_Binding imports Separation_Logic_Imperative_HOL.Imp_Map_Spec Separation_Logic_Imperative_HOL.Imp_Set_Spec Separation_Logic_Imperative_HOL.Imp_List_Spec Separation_Logic_Imperative_HOL.Hash_Map_Impl Separation_Logic_Imperative_HOL.Array_Map_Impl Separation_Logic_Imperative_HOL.To_List_GA Separation_Logic_Imperative_HOL.Hash_Set_Impl Separation_Logic_Imperative_HOL.Array_Set_Impl Separation_Logic_Imperative_HOL.Open_List Separation_Logic_Imperative_HOL.Circ_List "../Intf/IICF_Map" "../Intf/IICF_Set" "../Intf/IICF_List" Collections.Locale_Code begin text \This theory binds collection data structures from the basic collection framework established in \AFP/Separation_Logic_Imperative_HOL\ for usage with Sepref. \ (* TODO: Move, addition to Imp_Map_Spec *) locale imp_map_contains_key = imp_map + constrains is_map :: "('k \ 'v) \ 'm \ assn" fixes contains_key :: "'k \ 'm \ bool Heap" assumes contains_key_rule[sep_heap_rules]: " contains_key k p <\r. is_map m p * \(r\k\dom m)>\<^sub>t" (* TODO: Move to Imp_Map_Spec *) locale gen_contains_key_by_lookup = imp_map_lookup begin definition "contains_key k m \ do {r \ lookup k m; return (\is_None r)}" sublocale imp_map_contains_key is_map contains_key apply unfold_locales unfolding contains_key_def apply (sep_auto split: option.splits) done end (* TODO: Move to Imp_List_Spec *) locale imp_list_tail = imp_list + constrains is_list :: "'a list \ 'l \ assn" fixes tail :: "'l \ 'l Heap" assumes tail_rule[sep_heap_rules]: "l\[] \ tail p \<^sub>t" (* TODO: Move to Open_List *) definition os_head :: "'a::heap os_list \ ('a) Heap" where "os_head p \ case p of None \ raise STR ''os_Head: Empty list'' | Some p \ do { m \!p; return (val m) }" primrec os_tl :: "'a::heap os_list \ ('a os_list) Heap" where "os_tl None = raise STR ''os_tl: Empty list''" | "os_tl (Some p) = do { m \!p; return (next m) }" interpretation os: imp_list_head os_list os_head by unfold_locales (sep_auto simp: os_head_def neq_Nil_conv) interpretation os: imp_list_tail os_list os_tl by unfold_locales (sep_auto simp: os_tl_def neq_Nil_conv) (* TODO: Move to Circ_List *) definition cs_is_empty :: "'a::heap cs_list \ bool Heap" where "cs_is_empty p \ return (is_None p)" interpretation cs: imp_list_is_empty cs_list cs_is_empty by unfold_locales (sep_auto simp: cs_is_empty_def split: option.splits) definition cs_head :: "'a::heap cs_list \ 'a Heap" where "cs_head p \ case p of None \ raise STR ''cs_head: Empty list'' | Some p \ do { n \ !p; return (val n)}" interpretation cs: imp_list_head cs_list cs_head by unfold_locales (sep_auto simp: neq_Nil_conv cs_head_def) definition cs_tail :: "'a::heap cs_list \ 'a cs_list Heap" where "cs_tail p \ do { (_,r) \ cs_pop p; return r }" interpretation cs: imp_list_tail cs_list cs_tail by unfold_locales (sep_auto simp: cs_tail_def) (* TODO: Move to hashmap/hashset *) lemma is_hashmap_finite[simp]: "h \ is_hashmap m mi \ finite (dom m)" unfolding is_hashmap_def is_hashmap'_def by auto lemma is_hashset_finite[simp]: "h \ is_hashset s si \ finite s" unfolding is_hashset_def by (auto dest: is_hashmap_finite) (* TODO: Move to array-map/ array-set *) definition "ias_is_it s a si \ \(a',i). \\<^sub>Al. a\\<^sub>al * \(a'=a \ s=ias_of_list l \ (i=length l \ si={} \ i i\s \ si=s \ {x. x\i} )) " context begin private function first_memb where "first_memb lmax a i = do { if i Array.nth a i; if x then return i else first_memb lmax a (Suc i) } else return i }" by pat_completeness auto termination by (relation "measure (\(l,_,i). l-i)") auto declare first_memb.simps[simp del] private lemma first_memb_rl_aux: assumes "lmax \ length l" "i\lmax" shows "< a \\<^sub>a l > first_memb lmax a i <\k. a\\<^sub>a l * \(k\lmax \ (\j. i\j \ j \l!j) \ i\k \ (k=lmax \ l!k)) >" using assms proof (induction lmax a i rule: first_memb.induct) case (1 lmax a i) show ?case apply (subst first_memb.simps) using "1.prems" apply (sep_auto heap: "1.IH"; ((sep_auto;fail) | metis eq_iff not_less_eq_eq)) done qed private lemma first_memb_rl[sep_heap_rules]: assumes "lmax \ length l" "i\lmax" shows "< a \\<^sub>a l > first_memb lmax a i <\k. a\\<^sub>a l * \(ias_of_list l \ {i.. i\k \ (k k\ias_of_list l \ k=lmax) ) >" using assms by (sep_auto simp: ias_of_list_def heap: first_memb_rl_aux) definition "ias_it_init a = do { l \ Array.len a; i \ first_memb l a 0; return (a,i) }" definition "ias_it_has_next \ \(a,i). do { l \ Array.len a; return (i \(a,i). do { l \ Array.len a; i' \ first_memb l a (Suc i); return (i,(a,i')) }" (* TODO: Move *) lemma ias_of_list_bound: "ias_of_list l \ {0.. is_ias S x \ finite S" unfolding is_ias_def by auto (* TODO: Move, replace original rules by this stronger var! *) lemma to_list_ga_rec_rule: assumes "imp_set_iterate is_set is_it it_init it_has_next it_next" assumes "imp_list_prepend is_list l_prepend" assumes FIN: "finite it" assumes DIS: "distinct l" "set l \ it = {}" shows " < is_it s si it iti * is_list l li > to_list_ga_rec it_has_next it_next l_prepend iti li < \r. \\<^sub>Al'. is_set s si * is_list l' r * \(distinct l' \ set l' = set l \ it) >\<^sub>t" proof - interpret imp_set_iterate is_set is_it it_init it_has_next it_next + imp_list_prepend is_list l_prepend by fact+ from FIN DIS show ?thesis proof (induction arbitrary: l li iti rule: finite_psubset_induct) case (psubset it) show ?case apply (subst to_list_ga_rec.simps) using psubset.prems apply (sep_auto heap: psubset.IH) apply (rule ent_frame_fwd[OF quit_iteration]) apply frame_inference apply solve_entails done qed qed lemma to_list_ga_rule: assumes IT: "imp_set_iterate is_set is_it it_init it_has_next it_next" assumes EM: "imp_list_empty is_list l_empty" assumes PREP: "imp_list_prepend is_list l_prepend" assumes FIN: "finite s" shows " to_list_ga it_init it_has_next it_next l_empty l_prepend si <\r. \\<^sub>Al. is_set s si * is_list l r * true * \(distinct l \ set l = s)>" proof - interpret imp_list_empty is_list l_empty + imp_set_iterate is_set is_it it_init it_has_next it_next by fact+ note [sep_heap_rules] = to_list_ga_rec_rule[OF IT PREP] show ?thesis unfolding to_list_ga_def by (sep_auto simp: FIN) qed subsection \Binding Locales\ method solve_sepl_binding = ( unfold_locales; (unfold option_assn_pure_conv)?; sep_auto intro!: hfrefI hn_refineI[THEN hn_refine_preI] simp: invalid_assn_def hn_ctxt_def pure_def ) subsubsection \Map\ locale bind_map = imp_map is_map for is_map :: "('ki \ 'vi) \ 'm \ assn" begin definition "assn K V \ hr_comp is_map (\the_pure K,the_pure V\map_rel)" lemmas [fcomp_norm_unfold] = assn_def[symmetric] lemmas [safe_constraint_rules] = CN_FALSEI[of is_pure "assn K V" for K V] end locale bind_map_empty = imp_map_empty + bind_map begin lemma empty_hnr_aux: "(uncurry0 empty,uncurry0 (RETURN op_map_empty)) \ unit_assn\<^sup>k \\<^sub>a is_map" by solve_sepl_binding sepref_decl_impl (no_register) empty: empty_hnr_aux . end locale bind_map_is_empty = imp_map_is_empty + bind_map begin lemma is_empty_hnr_aux: "(is_empty,RETURN o op_map_is_empty) \ is_map\<^sup>k \\<^sub>a bool_assn" by solve_sepl_binding sepref_decl_impl is_empty: is_empty_hnr_aux . end locale bind_map_update = imp_map_update + bind_map begin lemma update_hnr_aux: "(uncurry2 update,uncurry2 (RETURN ooo op_map_update)) \ id_assn\<^sup>k *\<^sub>a id_assn\<^sup>k *\<^sub>a is_map\<^sup>d \\<^sub>a is_map" by solve_sepl_binding sepref_decl_impl update: update_hnr_aux . end locale bind_map_delete = imp_map_delete + bind_map begin lemma delete_hnr_aux: "(uncurry delete,uncurry (RETURN oo op_map_delete)) \ id_assn\<^sup>k *\<^sub>a is_map\<^sup>d \\<^sub>a is_map" by solve_sepl_binding sepref_decl_impl delete: delete_hnr_aux . end locale bind_map_lookup = imp_map_lookup + bind_map begin lemma lookup_hnr_aux: "(uncurry lookup,uncurry (RETURN oo op_map_lookup)) \ id_assn\<^sup>k *\<^sub>a is_map\<^sup>k \\<^sub>a id_assn" by solve_sepl_binding sepref_decl_impl lookup: lookup_hnr_aux . end locale bind_map_contains_key = imp_map_contains_key + bind_map begin lemma contains_key_hnr_aux: "(uncurry contains_key,uncurry (RETURN oo op_map_contains_key)) \ id_assn\<^sup>k *\<^sub>a is_map\<^sup>k \\<^sub>a bool_assn" by solve_sepl_binding sepref_decl_impl contains_key: contains_key_hnr_aux . end subsubsection \Set\ locale bind_set = imp_set is_set for is_set :: "('ai set) \ 'm \ assn" + fixes A :: "'a \ 'ai \ assn" begin definition "assn \ hr_comp is_set (\the_pure A\set_rel)" lemmas [safe_constraint_rules] = CN_FALSEI[of is_pure "assn"] end locale bind_set_setup = bind_set begin (* TODO: Use sepref_decl_impl (see map) *) lemmas [fcomp_norm_unfold] = assn_def[symmetric] lemma APA: "\PROP Q; CONSTRAINT is_pure A\ \ PROP Q" . lemma APAlu: "\PROP Q; CONSTRAINT (IS_PURE IS_LEFT_UNIQUE) A\ \ PROP Q" . lemma APAru: "\PROP Q; CONSTRAINT (IS_PURE IS_RIGHT_UNIQUE) A\ \ PROP Q" . lemma APAbu: "\PROP Q; CONSTRAINT (IS_PURE IS_LEFT_UNIQUE) A; CONSTRAINT (IS_PURE IS_RIGHT_UNIQUE) A\ \ PROP Q" . end locale bind_set_empty = imp_set_empty + bind_set begin lemma hnr_empty_aux: "(uncurry0 empty,uncurry0 (RETURN op_set_empty))\unit_assn\<^sup>k \\<^sub>a is_set" by solve_sepl_binding interpretation bind_set_setup by standard lemmas hnr_op_empty = hnr_empty_aux[FCOMP op_set_empty.fref[where A="the_pure A"]] lemmas hnr_mop_empty = hnr_op_empty[FCOMP mk_mop_rl0_np[OF mop_set_empty_alt]] end locale bind_set_is_empty = imp_set_is_empty + bind_set begin lemma hnr_is_empty_aux: "(is_empty, RETURN o op_set_is_empty)\is_set\<^sup>k \\<^sub>a bool_assn" by solve_sepl_binding interpretation bind_set_setup by standard lemmas hnr_op_is_empty[sepref_fr_rules] = hnr_is_empty_aux[THEN APA,FCOMP op_set_is_empty.fref[where A="the_pure A"]] lemmas hnr_mop_is_empty[sepref_fr_rules] = hnr_op_is_empty[FCOMP mk_mop_rl1_np[OF mop_set_is_empty_alt]] end locale bind_set_member = imp_set_memb + bind_set begin lemma hnr_member_aux: "(uncurry memb, uncurry (RETURN oo op_set_member))\id_assn\<^sup>k *\<^sub>a is_set\<^sup>k \\<^sub>a bool_assn" by solve_sepl_binding interpretation bind_set_setup by standard lemmas hnr_op_member[sepref_fr_rules] = hnr_member_aux[THEN APAbu,FCOMP op_set_member.fref[where A="the_pure A"]] lemmas hnr_mop_member[sepref_fr_rules] = hnr_op_member[FCOMP mk_mop_rl2_np[OF mop_set_member_alt]] end locale bind_set_insert = imp_set_ins + bind_set begin lemma hnr_insert_aux: "(uncurry ins, uncurry (RETURN oo op_set_insert))\id_assn\<^sup>k *\<^sub>a is_set\<^sup>d \\<^sub>a is_set" by solve_sepl_binding interpretation bind_set_setup by standard lemmas hnr_op_insert[sepref_fr_rules] = hnr_insert_aux[THEN APAru,FCOMP op_set_insert.fref[where A="the_pure A"]] lemmas hnr_mop_insert[sepref_fr_rules] = hnr_op_insert[FCOMP mk_mop_rl2_np[OF mop_set_insert_alt]] end locale bind_set_delete = imp_set_delete + bind_set begin lemma hnr_delete_aux: "(uncurry delete, uncurry (RETURN oo op_set_delete))\id_assn\<^sup>k *\<^sub>a is_set\<^sup>d \\<^sub>a is_set" by solve_sepl_binding interpretation bind_set_setup by standard lemmas hnr_op_delete[sepref_fr_rules] = hnr_delete_aux[THEN APAbu,FCOMP op_set_delete.fref[where A="the_pure A"]] lemmas hnr_mop_delete[sepref_fr_rules] = hnr_op_delete[FCOMP mk_mop_rl2_np[OF mop_set_delete_alt]] end primrec sorted_wrt' where "sorted_wrt' R [] \ True" | "sorted_wrt' R (x#xs) \ list_all (R x) xs \ sorted_wrt' R xs" lemma sorted_wrt'_eq: "sorted_wrt' = sorted_wrt" proof (intro ext iffI) fix R :: "'a \ 'a \ bool" and xs :: "'a list" { assume "sorted_wrt R xs" thus "sorted_wrt' R xs" - by (induction xs)(auto simp: list_all_iff sorted_sorted_wrt[symmetric]) + by (induction xs)(auto simp: list_all_iff) } { assume "sorted_wrt' R xs" thus "sorted_wrt R xs" by (induction xs) (auto simp: list_all_iff) } qed lemma param_sorted_wrt[param]: "(sorted_wrt, sorted_wrt) \ (A \ A \ bool_rel) \ \A\list_rel \ bool_rel" unfolding sorted_wrt'_eq[symmetric] sorted_wrt'_def by parametricity lemma obtain_list_from_setrel: assumes SV: "single_valued A" assumes "(set l,s) \ \A\set_rel" obtains m where "s=set m" "(l,m)\\A\list_rel" using assms(2) proof (induction l arbitrary: s thesis) case Nil show ?case apply (rule Nil(1)[where m="[]"]) using Nil(2) by auto next case (Cons x l) obtain s' y where "s=insert y s'" "(x,y)\A" "(set l,s')\\A\set_rel" proof - from Cons.prems(2) obtain y where X0: "y\s" "(x,y)\A" unfolding set_rel_def by auto from Cons.prems(2) have X1: "\a\set l. \b\s. (a,b)\A" and X2: "\b\s. \a\insert x (set l). (a,b)\A" unfolding set_rel_def by auto show ?thesis proof (cases "\a\set l. (a,y)\A") case True show ?thesis apply (rule that[of y s]) subgoal using X0 by auto subgoal by fact subgoal apply (rule set_relI) subgoal using X1 by blast subgoal by (metis IS_RIGHT_UNIQUED SV True X0(2) X2 insert_iff) done done next case False show ?thesis apply (rule that[of y "s-{y}"]) subgoal using X0 by auto subgoal by fact subgoal apply (rule set_relI) subgoal using False X1 by fastforce subgoal using IS_RIGHT_UNIQUED SV X0(2) X2 by fastforce done done qed qed moreover from Cons.IH[OF _ \(set l,s')\\A\set_rel\] obtain m where "s'=set m" "(l,m)\\A\list_rel" . ultimately show thesis apply - apply (rule Cons.prems(1)[of "y#m"]) by auto qed lemma param_it_to_sorted_list[param]: "\IS_LEFT_UNIQUE A; IS_RIGHT_UNIQUE A\ \ (it_to_sorted_list, it_to_sorted_list) \ (A \ A \ bool_rel) \ \A\set_rel \ \\A\list_rel\nres_rel" unfolding it_to_sorted_list_def[abs_def] apply (auto simp: it_to_sorted_list_def pw_nres_rel_iff refine_pw_simps) apply (rule obtain_list_from_setrel; assumption?; clarsimp) apply (intro exI conjI; assumption?) using param_distinct[param_fo] apply blast apply simp using param_sorted_wrt[param_fo] apply blast done locale bind_set_iterate = imp_set_iterate + bind_set + assumes is_set_finite: "h \ is_set S x \ finite S" begin context begin private lemma is_imp_set_iterate: "imp_set_iterate is_set is_it it_init it_has_next it_next" by unfold_locales private lemma is_imp_list_empty: "imp_list_empty (list_assn id_assn) (return [])" apply unfold_locales apply solve_constraint apply sep_auto done private lemma is_imp_list_prepend: "imp_list_prepend (list_assn id_assn) (return oo List.Cons)" apply unfold_locales apply solve_constraint apply (sep_auto simp: pure_def) done definition "to_list \ to_list_ga it_init it_has_next it_next (return []) (return oo List.Cons)" private lemmas tl_rl = to_list_ga_rule[OF is_imp_set_iterate is_imp_list_empty is_imp_list_prepend, folded to_list_def] private lemma to_list_sorted1: "(to_list,PR_CONST (it_to_sorted_list (\_ _. True))) \ is_set\<^sup>k \\<^sub>a list_assn id_assn" unfolding PR_CONST_def apply (intro hfrefI) apply (rule hn_refine_preI) apply (rule hn_refineI) unfolding it_to_sorted_list_def apply (sep_auto intro: hfrefI hn_refineI intro: is_set_finite heap: tl_rl) done private lemma to_list_sorted2: "\ CONSTRAINT (IS_PURE IS_LEFT_UNIQUE) A; CONSTRAINT (IS_PURE IS_RIGHT_UNIQUE) A\ \ (PR_CONST (it_to_sorted_list (\_ _. True)), PR_CONST (it_to_sorted_list (\_ _. True))) \ \the_pure A\set_rel \ \\the_pure A\list_rel\nres_rel" unfolding PR_CONST_def CONSTRAINT_def IS_PURE_def by clarify parametricity lemmas to_list_hnr = to_list_sorted1[FCOMP to_list_sorted2, folded assn_def] lemmas to_list_is_to_sorted_list = IS_TO_SORTED_LISTI[OF to_list_hnr] lemma to_list_gen[sepref_gen_algo_rules]: "\CONSTRAINT (IS_PURE IS_LEFT_UNIQUE) A; CONSTRAINT (IS_PURE IS_RIGHT_UNIQUE) A\ \ GEN_ALGO to_list (IS_TO_SORTED_LIST (\_ _. True) (bind_set.assn is_set A) A)" by (simp add: GEN_ALGO_def to_list_is_to_sorted_list) end end subsubsection \List\ locale bind_list = imp_list is_list for is_list :: "('ai list) \ 'm \ assn" + fixes A :: "'a \ 'ai \ assn" begin (*abbreviation "Ap \ the_pure A"*) definition "assn \ hr_comp is_list (\the_pure A\list_rel)" lemmas [safe_constraint_rules] = CN_FALSEI[of is_pure "assn"] end locale bind_list_empty = imp_list_empty + bind_list begin lemma hnr_aux: "(uncurry0 empty,uncurry0 (RETURN op_list_empty))\(pure unit_rel)\<^sup>k \\<^sub>a is_list" apply rule apply rule apply (sep_auto simp: pure_def) done lemmas hnr = hnr_aux[FCOMP op_list_empty.fref[of "the_pure A"], folded assn_def] lemmas hnr_mop = hnr[FCOMP mk_mop_rl0_np[OF mop_list_empty_alt]] end locale bind_list_is_empty = imp_list_is_empty + bind_list begin lemma hnr_aux: "(is_empty,RETURN o op_list_is_empty)\(is_list)\<^sup>k \\<^sub>a pure bool_rel" apply rule apply rule apply (sep_auto simp: pure_def) done lemmas hnr[sepref_fr_rules] = hnr_aux[FCOMP op_list_is_empty.fref, of "the_pure A", folded assn_def] lemmas hnr_mop[sepref_fr_rules] = hnr[FCOMP mk_mop_rl1_np[OF mop_list_is_empty_alt]] end locale bind_list_append = imp_list_append + bind_list begin lemma hnr_aux: "(uncurry (swap_args2 append),uncurry (RETURN oo op_list_append)) \(is_list)\<^sup>d *\<^sub>a (pure Id)\<^sup>k \\<^sub>a is_list" by solve_sepl_binding lemmas hnr[sepref_fr_rules] = hnr_aux[FCOMP op_list_append.fref,of A, folded assn_def] lemmas hnr_mop[sepref_fr_rules] = hnr[FCOMP mk_mop_rl2_np[OF mop_list_append_alt]] end locale bind_list_prepend = imp_list_prepend + bind_list begin lemma hnr_aux: "(uncurry prepend,uncurry (RETURN oo op_list_prepend)) \(pure Id)\<^sup>k *\<^sub>a (is_list)\<^sup>d \\<^sub>a is_list" by solve_sepl_binding lemmas hnr[sepref_fr_rules] = hnr_aux[FCOMP op_list_prepend.fref,of A, folded assn_def] lemmas hnr_mop[sepref_fr_rules] = hnr[FCOMP mk_mop_rl2_np[OF mop_list_prepend_alt]] end locale bind_list_hd = imp_list_head + bind_list begin lemma hnr_aux: "(head,RETURN o op_list_hd) \[\l. l\[]]\<^sub>a (is_list)\<^sup>d \ pure Id" by solve_sepl_binding lemmas hnr[sepref_fr_rules] = hnr_aux[FCOMP op_list_hd.fref,of A, folded assn_def] lemmas hnr_mop[sepref_fr_rules] = hnr[FCOMP mk_mop_rl1[OF mop_list_hd_alt]] end locale bind_list_tl = imp_list_tail + bind_list begin lemma hnr_aux: "(tail,RETURN o op_list_tl) \[\l. l\[]]\<^sub>a (is_list)\<^sup>d \ is_list" by solve_sepl_binding lemmas hnr[sepref_fr_rules] = hnr_aux[FCOMP op_list_tl.fref,of "the_pure A", folded assn_def] lemmas hnr_mop[sepref_fr_rules] = hnr[FCOMP mk_mop_rl1[OF mop_list_tl_alt]] end locale bind_list_rotate1 = imp_list_rotate + bind_list begin lemma hnr_aux: "(rotate,RETURN o op_list_rotate1) \(is_list)\<^sup>d \\<^sub>a is_list" by solve_sepl_binding lemmas hnr[sepref_fr_rules] = hnr_aux[FCOMP op_list_rotate1.fref,of "the_pure A", folded assn_def] lemmas hnr_mop[sepref_fr_rules] = hnr[FCOMP mk_mop_rl1_np[OF mop_list_rotate1_alt]] end locale bind_list_rev = imp_list_reverse + bind_list begin lemma hnr_aux: "(reverse,RETURN o op_list_rev) \(is_list)\<^sup>d \\<^sub>a is_list" by solve_sepl_binding lemmas hnr[sepref_fr_rules] = hnr_aux[FCOMP op_list_rev.fref,of "the_pure A", folded assn_def] lemmas hnr_mop[sepref_fr_rules] = hnr[FCOMP mk_mop_rl1_np[OF mop_list_rev_alt]] end subsection \Array Map (iam)\ definition "op_iam_empty \ IICF_Map.op_map_empty" interpretation iam: bind_map_empty is_iam iam_new by unfold_locales interpretation iam: map_custom_empty op_iam_empty by unfold_locales (simp add: op_iam_empty_def) lemmas [sepref_fr_rules] = iam.empty_hnr[folded op_iam_empty_def] definition [simp]: "op_iam_empty_sz (N::nat) \ IICF_Map.op_map_empty" lemma [def_pat_rules]: "op_iam_empty_sz$N \ UNPROTECT (op_iam_empty_sz N)" by simp interpretation iam_sz: map_custom_empty "PR_CONST (op_iam_empty_sz N)" apply unfold_locales apply (simp) done lemma [sepref_fr_rules]: "(uncurry0 iam_new, uncurry0 (RETURN (PR_CONST (op_iam_empty_sz N)))) \ unit_assn\<^sup>k \\<^sub>a iam.assn K V" using iam.empty_hnr[of K V] by simp interpretation iam: bind_map_update is_iam Array_Map_Impl.iam_update by unfold_locales interpretation iam: bind_map_delete is_iam Array_Map_Impl.iam_delete by unfold_locales interpretation iam: bind_map_lookup is_iam Array_Map_Impl.iam_lookup by unfold_locales setup Locale_Code.open_block interpretation iam: gen_contains_key_by_lookup is_iam Array_Map_Impl.iam_lookup by unfold_locales setup Locale_Code.close_block interpretation iam: bind_map_contains_key is_iam iam.contains_key by unfold_locales subsection \Array Set (ias)\ definition [simp]: "op_ias_empty \ op_set_empty" interpretation ias: bind_set_empty is_ias ias_new for A by unfold_locales interpretation ias: set_custom_empty ias_new op_ias_empty by unfold_locales simp lemmas [sepref_fr_rules] = ias.hnr_op_empty[folded op_ias_empty_def] definition [simp]: "op_ias_empty_sz (N::nat) \ op_set_empty" lemma [def_pat_rules]: "op_ias_empty_sz$N \ UNPROTECT (op_ias_empty_sz N)" by simp interpretation ias_sz: bind_set_empty is_ias "ias_new_sz N" for N A by unfold_locales interpretation ias_sz: set_custom_empty "ias_new_sz N" "PR_CONST (op_ias_empty_sz N)" for A by unfold_locales simp lemma [sepref_fr_rules]: "(uncurry0 (ias_new_sz N), uncurry0 (RETURN (PR_CONST (op_ias_empty_sz N)))) \ unit_assn\<^sup>k \\<^sub>a ias.assn A" using ias_sz.hnr_op_empty[of N A] by simp interpretation ias: bind_set_member is_ias Array_Set_Impl.ias_memb for A by unfold_locales interpretation ias: bind_set_insert is_ias Array_Set_Impl.ias_ins for A by unfold_locales interpretation ias: bind_set_delete is_ias Array_Set_Impl.ias_delete for A by unfold_locales setup Locale_Code.open_block interpretation ias: bind_set_iterate is_ias ias_is_it ias_it_init ias_it_has_next ias_it_next for A by unfold_locales auto setup Locale_Code.close_block subsection \Hash Map (hm)\ interpretation hm: bind_map_empty is_hashmap hm_new by unfold_locales definition "op_hm_empty \ IICF_Map.op_map_empty" interpretation hm: map_custom_empty op_hm_empty by unfold_locales (simp add: op_hm_empty_def) lemmas [sepref_fr_rules] = hm.empty_hnr[folded op_hm_empty_def] interpretation hm: bind_map_is_empty is_hashmap Hash_Map.hm_isEmpty by unfold_locales interpretation hm: bind_map_update is_hashmap Hash_Map.hm_update by unfold_locales interpretation hm: bind_map_delete is_hashmap Hash_Map.hm_delete by unfold_locales interpretation hm: bind_map_lookup is_hashmap Hash_Map.hm_lookup by unfold_locales setup Locale_Code.open_block interpretation hm: gen_contains_key_by_lookup is_hashmap Hash_Map.hm_lookup by unfold_locales setup Locale_Code.close_block interpretation hm: bind_map_contains_key is_hashmap hm.contains_key by unfold_locales subsection \Hash Set (hs)\ interpretation hs: bind_set_empty is_hashset hs_new for A by unfold_locales definition "op_hs_empty \ IICF_Set.op_set_empty" interpretation hs: set_custom_empty hs_new op_hs_empty for A by unfold_locales (simp add: op_hs_empty_def) lemmas [sepref_fr_rules] = hs.hnr_op_empty[folded op_hs_empty_def] interpretation hs: bind_set_is_empty is_hashset Hash_Set_Impl.hs_isEmpty for A by unfold_locales interpretation hs: bind_set_member is_hashset Hash_Set_Impl.hs_memb for A by unfold_locales interpretation hs: bind_set_insert is_hashset Hash_Set_Impl.hs_ins for A by unfold_locales interpretation hs: bind_set_delete is_hashset Hash_Set_Impl.hs_delete for A by unfold_locales setup Locale_Code.open_block interpretation hs: bind_set_iterate is_hashset hs_is_it hs_it_init hs_it_has_next hs_it_next for A by unfold_locales simp setup Locale_Code.close_block subsection \Open Singly Linked List (osll)\ interpretation osll: bind_list os_list for A by unfold_locales interpretation osll_empty: bind_list_empty os_list os_empty for A by unfold_locales definition "osll_empty \ op_list_empty" interpretation osll: list_custom_empty "osll.assn A" os_empty osll_empty apply unfold_locales apply (rule osll_empty.hnr) by (simp add: osll_empty_def) interpretation osll_is_empty: bind_list_is_empty os_list os_is_empty for A by unfold_locales interpretation osll_prepend: bind_list_prepend os_list os_prepend for A by unfold_locales interpretation osll_hd: bind_list_hd os_list os_head for A by unfold_locales interpretation osll_tl: bind_list_tl os_list os_tl for A by unfold_locales interpretation osll_rev: bind_list_rev os_list os_reverse for A by unfold_locales subsection \Circular Singly Linked List (csll)\ (* TODO: In-place reversal of circular list! *) interpretation csll: bind_list cs_list for A by unfold_locales interpretation csll_empty: bind_list_empty cs_list cs_empty for A by unfold_locales definition "csll_empty \ op_list_empty" interpretation csll: list_custom_empty "csll.assn A" cs_empty csll_empty apply unfold_locales apply (rule csll_empty.hnr) by (simp add: csll_empty_def) interpretation csll_is_empty: bind_list_is_empty cs_list cs_is_empty for A by unfold_locales interpretation csll_prepend: bind_list_prepend cs_list cs_prepend for A by unfold_locales interpretation csll_append: bind_list_append cs_list cs_append for A by unfold_locales interpretation csll_hd: bind_list_hd cs_list cs_head for A by unfold_locales interpretation csll_tl: bind_list_tl cs_list cs_tail for A by unfold_locales interpretation csll_rotate1: bind_list_rotate1 cs_list cs_rotate for A by unfold_locales schematic_goal "hn_refine (emp) (?c::?'c Heap) ?\' ?R (do { x \ mop_list_empty; RETURN (1 \ dom [1::nat \ True, 2\False], {1,2::nat}, 1#(2::nat)#x) })" apply (subst iam_sz.fold_custom_empty[where N=10]) apply (subst hs.fold_custom_empty) apply (subst osll.fold_custom_empty) by sepref end diff --git a/thys/Smith_Normal_Form/Cauchy_Binet.thy b/thys/Smith_Normal_Form/Cauchy_Binet.thy --- a/thys/Smith_Normal_Form/Cauchy_Binet.thy +++ b/thys/Smith_Normal_Form/Cauchy_Binet.thy @@ -1,1425 +1,1425 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \The Cauchy--Binet formula\ theory Cauchy_Binet imports Diagonal_To_Smith SNF_Missing_Lemmas begin subsection \Previous missing results about @{text "pick"} and @{text "insert"}\ lemma pick_insert: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (insort a (sorted_list_of_set I)) a = a'*) and ia': "i < a'" (*Case 1*) and a'_card: "a' < card I + 1" shows "pick (insert a I) i = pick I i" proof - have finI: "finite I" using i2 using card.infinite by force have "pick (insert a I) i = sorted_list_of_set (insert a I) ! i" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "i < length (sorted_list_of_set (insert a I))" by (metis a_notin_I card_insert_disjoint distinct_card finite_insert i2 less_Suc_eq sorted_list_of_set(1) sorted_list_of_set(3)) qed also have "... = insort a (sorted_list_of_set I) ! i" using sorted_list_of_set.insert by (metis a_notin_I card.infinite i2 not_less0) also have "... = (sorted_list_of_set I) ! i" proof (rule insort_nth[OF]) show "sorted (sorted_list_of_set I)" by auto show "a \ set (sorted_list_of_set I)" using a_notin_I by (metis card.infinite i2 not_less_zero set_sorted_list_of_set) have "index (sorted_list_of_set (insert a I)) a = a'" using pick_index a_def using a'_card a_notin_I finI by auto hence "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) thus "i < index (insort a (sorted_list_of_set I)) a" using ia' by auto show "sorted_list_of_set I \ []" using finI i2 by fastforce qed also have "... = pick I i" proof (rule sorted_list_of_set_eq_pick) have "finite I" using card.infinite i2 by fastforce thus "i < length (sorted_list_of_set I)" by (metis distinct_card distinct_sorted_list_of_set i2 set_sorted_list_of_set) qed finally show ?thesis . qed lemma pick_insert2: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (sorted_list_of_set (insert a I)) a = a'*) and ia': "i \ a'" (*Case 2*) and a'_card: "a' < card I + 1" shows "pick (insert a I) i < pick I i" proof (cases "i = 0") case True then show ?thesis by (auto, metis (mono_tags, lifting) DL_Missing_Sublist.pick.simps(1) Least_le a_def a_notin_I dual_order.order_iff_strict i2 ia' insertCI le_zero_eq not_less_Least pick_in_set_le) next case False hence i0: "i = Suc (i - 1)" using a'_card ia' by auto have finI: "finite I" using i2 card.infinite by force have index_a'1: "index (sorted_list_of_set (insert a I)) a = a'" using pick_index using a'_card a_def a_notin_I finI by auto hence index_a': "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) have i1_length: "i - 1 < length (sorted_list_of_set I)" using i2 by (metis distinct_card distinct_sorted_list_of_set finI less_imp_diff_less set_sorted_list_of_set) have 1: "pick (insert a I) i = sorted_list_of_set (insert a I) ! i" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "i < length (sorted_list_of_set (insert a I))" by (metis a_notin_I card_insert_disjoint distinct_card finite_insert i2 less_Suc_eq sorted_list_of_set(1) sorted_list_of_set(3)) qed also have 2: "... = insort a (sorted_list_of_set I) ! i" using sorted_list_of_set.insert by (metis a_notin_I card.infinite i2 not_less0) also have "... = insort a (sorted_list_of_set I) ! Suc (i-1)" using i0 by auto also have "... < pick I i" proof (cases "i = a'") case True have "(sorted_list_of_set I) ! i > a" by (smt "1" Suc_less_eq True a_def a_notin_I distinct_card distinct_sorted_list_of_set finI i2 ia' index_a' insort_nth2 length_insort lessI list.size(3) nat_less_le not_less_zero pick_in_set_le set_sorted_list_of_set sorted_list_of_set(2) sorted_list_of_set.insert - sorted_list_of_set_eq_pick sorted_sorted_wrt sorted_wrt_nth_less) + sorted_list_of_set_eq_pick sorted_wrt_nth_less) moreover have "a = insort a (sorted_list_of_set I) ! i" using True 1 2 a_def by auto ultimately show ?thesis using 1 2 by (metis distinct_card finI i0 i2 set_sorted_list_of_set sorted_list_of_set(3) sorted_list_of_set_eq_pick) next case False have "insort a (sorted_list_of_set I) ! Suc (i-1) = (sorted_list_of_set I) ! (i-1)" by (rule insort_nth2, insert i1_length False ia' index_a', auto simp add: a_notin_I finI) also have "... = pick I (i-1)" by (rule sorted_list_of_set_eq_pick[OF i1_length]) also have "... < pick I i" using i0 i2 pick_mono_le by auto finally show ?thesis . qed finally show ?thesis . qed lemma pick_insert3: assumes a_notin_I: "a \ I" and i2: "i < card I" and a_def: "pick (insert a I) a' = a" (*Alternative: index (sorted_list_of_set (insert a I)) a = a'.*) and ia': "i \ a'" (*Case 2*) and a'_card: "a' < card I + 1" shows "pick (insert a I) (Suc i) = pick I i" proof (cases "i = 0") case True have a_LEAST: "a < (LEAST aa. aa\I)" using True a_def a_notin_I i2 ia' pick_insert2 by fastforce have Least_rw: "(LEAST aa. aa = a \ aa \ I) = a" by (rule Least_equality, insert a_notin_I, auto, metis a_LEAST le_less_trans nat_le_linear not_less_Least) let ?P = "\aa. (aa = a \ aa \ I) \ (LEAST aa. aa = a \ aa \ I) < aa" let ?Q = "\aa. aa \ I" have "?P = ?Q" unfolding Least_rw fun_eq_iff by (auto, metis a_LEAST le_less_trans not_le not_less_Least) thus ?thesis using True by auto next case False have finI: "finite I" using i2 card.infinite by force have index_a'1: "index (sorted_list_of_set (insert a I)) a = a'" using pick_index using a'_card a_def a_notin_I finI by auto hence index_a': "index (insort a (sorted_list_of_set I)) a = a'" by (simp add: a_notin_I finI) have i1_length: "i < length (sorted_list_of_set I)" using i2 by (metis distinct_card distinct_sorted_list_of_set finI set_sorted_list_of_set) have 1: "pick (insert a I) (Suc i) = sorted_list_of_set (insert a I) ! (Suc i)" proof (rule sorted_list_of_set_eq_pick[symmetric]) have "finite (insert a I)" using card.infinite i2 by force thus "Suc i < length (sorted_list_of_set (insert a I))" by (metis Suc_mono a_notin_I card_insert_disjoint distinct_card distinct_sorted_list_of_set finI i2 set_sorted_list_of_set) qed also have 2: "... = insort a (sorted_list_of_set I) ! Suc i" using sorted_list_of_set.insert by (metis a_notin_I card.infinite i2 not_less0) also have "... = pick I i" proof (cases "i = a'") case True show ?thesis by (metis True a_notin_I finI i1_length index_a' insort_nth2 le_refl list.size(3) not_less0 set_sorted_list_of_set sorted_list_of_set(2) sorted_list_of_set_eq_pick) next case False have "insort a (sorted_list_of_set I) ! Suc i = (sorted_list_of_set I) ! i" by (rule insort_nth2, insert i1_length False ia' index_a', auto simp add: a_notin_I finI) also have "... = pick I i" by (rule sorted_list_of_set_eq_pick[OF i1_length]) finally show ?thesis . qed finally show ?thesis . qed lemma pick_insert_index: assumes Ik: "card I = k" and a_notin_I: "a \ I" and ik: "i < k" and a_def: "pick (insert a I) a' = a" and a'k: "a' < card I + 1" shows "pick (insert a I) (insert_index a' i) = pick I i" proof (cases "i I" by (simp add: Ik ik pick_in_set_le) show "pick (insert a I) i < pick I i" by (rule pick_insert2[OF a_notin_I _ a_def _ a'k], insert False, auto simp add: Ik ik) fix y assume y: "y \ I \ pick (insert a I) i < y" let ?xs = "sorted_list_of_set (insert a I)" have "y \ set ?xs" using y by (metis fin_aI insertI2 set_sorted_list_of_set y) from this obtain j where xs_j_y: "?xs ! j = y" and j: "j < length ?xs" using in_set_conv_nth by metis have ij: "i pick (insert a I) j" by (metis Ik Suc_lessI card.infinite distinct_card distinct_sorted_list_of_set eq_iff finite_insert ij ik j less_imp_le_nat not_less_zero pick_mono_le set_sorted_list_of_set) also have "... = ?xs ! j" by (rule sorted_list_of_set_eq_pick[symmetric, OF j]) also have "... = y" by (rule xs_j_y) finally show "pick I i \ y" . qed finally show ?thesis unfolding insert_index_def using False by auto qed subsection\Start of the proof\ definition "strict_from_inj n f = (\i. if i\{0.. nat" assumes "inj_on f {0.. {0.. {0.. f ` {0.. f ` {0.. strict_from_inj n f ` {0..)|f \. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i)} \ {\. \ permutes {0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0..T \ carrier_mat m n" using A by auto have BT: "B\<^sup>T \ carrier_mat n m" using B by auto let ?f = "(\i. finsum_vec TYPE('a) n (\k. B\<^sup>T $$ (i, k) \\<^sub>v Matrix.row A\<^sup>T k) {0..k\{0..T $$ (i, k) \\<^sub>v row A\<^sup>T k) $ j)" by (rule index_finsum_vec[OF _ j_n], auto simp add: A) also have "... = (\k\{0..\<^sub>v col A k) $ j)" proof (rule sum.cong, auto) fix x assume x: "xT x = col A x" by (rule row_transpose, insert A x, auto) have B_rw: "B\<^sup>T $$ (i,x) = B $$ (x, i)" by (rule index_transpose_mat, insert x i B, auto) have "(B\<^sup>T $$ (i, x) \\<^sub>v Matrix.row A\<^sup>T x) $v j = B\<^sup>T $$ (i, x) * Matrix.row A\<^sup>T x $v j" by (rule index_smult_vec, insert A j_n, auto) also have "... = B $$ (x, i) * col A x $v j" unfolding row_rw B_rw by simp also have "... = (B $$ (x, i) \\<^sub>v col A x) $v j" by (rule index_smult_vec[symmetric], insert A j_n, auto) finally show " (B\<^sup>T $$ (i, x) \\<^sub>v Matrix.row A\<^sup>T x) $v j = (B $$ (x, i) \\<^sub>v col A x) $v j" . qed also have "... = ?g i $v j" by (rule index_finsum_vec[symmetric, OF _ j_n], auto simp add: A) also have "... = ?rhs $$ (i, j)" by (rule index_mat[symmetric], insert i j, auto) finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" . qed have "det (A*B) = det (B\<^sup>T*A\<^sup>T)" using det_transpose by (metis A B Matrix.transpose_mult mult_carrier_mat) also have "... = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a) n (\k. B\<^sup>T $$ (i, k) \\<^sub>v Matrix.row A\<^sup>T k) {0..r n n (\i. finsum_vec TYPE('a) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\f | (\i\{0.. {0.. (\i. i \ {0.. f i = i). (\i = 0..r n n (\i. col A (f i))))" proof - let ?V="{0..r n n (\i. B $$ (f i, i) \\<^sub>v col A (f i))) = (prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i)))" if f: "(\i\{0.. {0.. (\i. i \ {0.. f i = i)" for f by (rule det_rows_mul, insert A col_dim, auto) have "det (A*B) = det (mat\<^sub>r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) ?U))" by (rule det_mul_finsum_alt[OF A B]) also have "... = sum ?g ?F" by (rule det_linear_rows_sum[OF fm], auto simp add: A) also have "... = (\f\?F. prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i))))" using det_rw by auto finally show ?thesis . qed lemma det_cols_mul': assumes A: "A \ carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\f | (\i\{0.. {0.. (\i. i \ {0.. f i = i). (\i = 0..r n n (\i. row B (f i))))" proof - let ?F="{f. (\i\{0.. {0.. (\i. i \ {0.. f i = i)}" have t: "A * B = (B\<^sup>T*A\<^sup>T)\<^sup>T" using transpose_mult[OF A B] transpose_transpose by metis have "det (B\<^sup>T*A\<^sup>T) = (\f\?F. (\i = 0..T $$ (f i, i)) * det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))))" by (rule det_cols_mul, auto simp add: A B) also have "... = (\f \?F. (\i = 0..r n n (\i. row B (f i))))" proof (rule sum.cong, rule refl) fix f assume f: "f \ ?F" have "(\i = 0..T $$ (f i, i)) = (\i = 0.. {0..T $$ (f x, x) = A $$ (x, f x)" by (rule index_transpose_mat(1), insert f A x, auto) qed moreover have "det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))) = det (mat\<^sub>r n n (\i. row B (f i)))" proof - have row_eq_colT: "row B (f i) $v j = col B\<^sup>T (f i) $v j" if i: "i < n" and j: "j < n" for i j proof - have fi_m: "f i < m" using f i by auto have "col B\<^sup>T (f i) $v j = B\<^sup>T $$(j, f i)" by (rule index_col, insert B fi_m j, auto) also have "... = B $$ (f i, j)" using B fi_m j by auto also have "... = row B (f i) $v j" by (rule index_row[symmetric], insert B fi_m j, auto) finally show ?thesis .. qed show ?thesis by (rule arg_cong[of _ _ det], rule eq_matI, insert row_eq_colT, auto) qed ultimately show "(\i = 0..T $$ (f i, i)) * det (mat\<^sub>r n n (\i. col B\<^sup>T (f i))) = (\i = 0..r n n (\i. row B (f i)))" by simp qed finally show ?thesis by (metis (no_types, lifting) A B det_transpose transpose_mult mult_carrier_mat) qed (*We need a more general version of this lemma*) lemma assumes F: "F= {f. f \ {0.. {0.. (\i. i \ {0.. f i = i)}" and p: " \ permutes {0..f\F. (\i = 0.. i))) = (\f\F. (\i = 0.. \ = g \ \" have "f x = g x" for x proof (cases "x \ {0.. F" show "xa \ \ \ F" unfolding o_def F using F PiE p xa by (auto, smt F atLeastLessThan_iff mem_Collect_eq p permutes_def xa) show "\x\F. xa = x \ \" proof (rule bexI[of _ "xa \ Hilbert_Choice.inv \"]) show "xa = xa \ Hilbert_Choice.inv \ \ \" using p by auto show "xa \ Hilbert_Choice.inv \ \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) qed qed have prod_rw: "(\i = 0..i = 0.. i), \ i))" if "f\F" for f using prod.permute[OF p] by auto let ?g = "\f. (\i = 0.. i))" have "(\f\F. (\i = 0..f\F. (\i = 0.. i), \ i)))" using prod_rw by auto also have "... = (\f\(?h`F). \i = 0.. i))" using sum.reindex[OF inj_on_F, of ?g] unfolding hF by auto also have "... = (\f\F. \i = 0.. i))" unfolding hF by auto finally show ?thesis .. qed lemma detAB_Znm_aux: assumes F: "F= {f. f \ {0.. {0.. (\i. i \ {0.. f i = i)}" shows"(\\ | \ permutes {0..f\F. prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))) = (\\ | \ permutes {0..f\F. (\i = 0.. i)) * (signof \ * (\i = 0..\ | \ permutes {0..f\F. prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))) = (\\ | \ permutes {0..f\F. signof \ * (\i = 0.. i, f i)))" by (smt mult.left_commute prod.cong prod.distrib sum.cong) also have "... = (\\ | \ permutes {0..f\F. signof (Hilbert_Choice.inv \) * (\i = 0.. i, f i)))" by (rule sum_permutations_inverse) also have "... = (\\ | \ permutes {0..f\F. signof (Hilbert_Choice.inv \) * (\i = 0.. i), (\ i)) * A $$ (Hilbert_Choice.inv \ (\ i), f (\ i))))" proof (rule sum.cong) fix x assume x: "x \ {\. \ permutes {0..i = 0..i = 0.. F" for f using prod.permute[OF p] by auto then show "(\f\F. signof ?inv_x * (\i = 0..f\F. signof ?inv_x * (\i = 0..\ | \ permutes {0..f\F. signof \ * (\i = 0.. i), (\ i)) * A $$ (i, f (\ i))))" by (rule sum.cong, auto, rule sum.cong, auto) (metis (no_types, lifting) finite_atLeastLessThan signof_inv) also have "... = (\\ | \ permutes {0..f\F. signof \ * (\i = 0.. i)) * A $$ (i, f i)))" proof (rule sum.cong) fix \ assume p: "\ \ {\. \ permutes {0.. permutes {0.. ?inv_pi = g \ ?inv_pi" have "f x = g x" for x proof (cases "x \ {0.. F" show "xa \ ?inv_pi \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) show "\x\F. xa = x \ ?inv_pi" proof (rule bexI[of _ "xa \ \"]) show "xa = xa \ \ \ Hilbert_Choice.inv \ " using p by auto show "xa \ \ \ F" unfolding o_def F using F PiE p xa by (auto, smt atLeastLessThan_iff permutes_def permutes_less(3)) qed qed let ?g = "\f. signof \ * (\i = 0.. i), \ i) * A $$ (i, f (\ i)))" show "(\f\F. signof \ * (\i = 0.. i), \ i) * A $$ (i, f (\ i)))) = (\f\F. signof \ * (\i = 0.. i) * A $$ (i, f i)))" using sum.reindex[OF inj_on_F, of "?g"] p unfolding hF unfolding o_def by auto qed (simp) also have "... = (\\ | \ permutes {0..f\F. (\i = 0.. i)) * (signof \ * (\i = 0.. carrier_mat n m" and B: "B \ carrier_mat m n" shows "det (A*B) = (\(f, \)\Z n m. signof \ * (\i = 0.. i)))" proof - let ?V="{0.. {0.. {0.. (\i. i \ {0.. f i = i)}" by auto have det_rw: "det (mat\<^sub>r n n (\i. B $$ (f i, i) \\<^sub>v col A (f i))) = (prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i)))" if f: "(\i\{0.. {0.. (\i. i \ {0.. f i = i)" for f by (rule det_rows_mul, insert A col_dim, auto) have det_rw2: "det (mat\<^sub>r n n (\i. col A (f i))) = (\\ | \ permutes {0.. * (\i = 0.. i, f i)))" if f: "f \ ?F" for f proof (unfold Determinant.det_def, auto, rule sum.cong, auto) fix x assume x: "x permutes {0..i = 0..i = 0.. {0..i = 0..i = 0..r n n (\i. finsum_vec TYPE('a::comm_ring_1) n (\k. B $$ (k, i) \\<^sub>v Matrix.col A k) {0..f\?F. prod (\i. B $$ (f i, i)) {0..r n n (\i. col A (f i))))" using det_rw by auto also have "... = (\f\?F. prod (\i. B $$ (f i, i)) {0..\ | \ permutes {0.. * (\i = 0.. i, f (i)))))" by (rule sum.cong, auto simp add: det_rw2) also have "... = (\f\?F. \\ | \ permutes {0..i. B $$ (f i, i)) {0.. * (\i = 0.. i, f (i)))))" by (simp add: mult_hom.hom_sum) also have "... = (\\ | \ permutes {0..f\?F.prod (\i. B $$ (f i, i)) {0.. * (\i = 0.. i, f i))))" by (rule VS_Connect.class_semiring.finsum_finsum_swap, insert finite_permutations finite_bounded_functions[OF fin_m fin_n], auto) thm detAB_Znm_aux also have "... = (\\ | \ permutes {0..f\?F. (\i = 0.. i)) * (signof \ * (\i = 0..f\?F.\\ | \ permutes {0..i = 0.. i)) * (signof \ * (\i = 0..f\?F.\\ | \ permutes {0.. * (\i = 0.. i)))" unfolding prod.distrib by (rule sum.cong, auto, rule sum.cong, auto) also have "... = sum (\(f,\). (signof \) * (prod (\i. A$$(i,f i) * B $$ (f i, \ i)) {0.. carrier_mat n m" and B: "B \ carrier_mat m n" begin private definition "Z_inj = ({f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ inj_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ strict_mono_on f {0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ strict_mono_on f {0.. {\. \ permutes {0.. = (signof \) * (prod (\i. A$$(i,f i) * B $$ (f i, \ i)) {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ strict_mono_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ \ inj_on f {0.. {0.. {0.. (\i. i \ {0.. f i = i)}" text\The Cauchy--Binet formula is proven in \url{https://core.ac.uk/download/pdf/82475020.pdf} In that work, they define @{text "\ \ inv \ \ \"}. I had problems following this proof in Isabelle, since I was demanded to show that such permutations commute, which is false. It is a notation problem of the @{text "\"} operator, the author means @{text "\ \ \ \ inv \"} using the Isabelle notation (i.e., @{text "\ x = \ ((inv \) x)"}). \ lemma step_weight: fixes \ \ defines "\ \ \ \ Hilbert_Choice.inv \" assumes f_inj: "f \ F_inj" and gF: "g \ F" and pi: "\ permutes {0.. permutes {0..x \ {0.. x)" shows "weight f \ = (signof \) * (\i = 0.. i))) * (signof \) * (\i = 0.. i))" proof - let ?A = "(\i = 0.. i))) " let ?B = "(\i = 0.. i))" have sigma: "\ permutes {0.._def by (rule permutes_compose[OF permutes_inv[OF phi] pi]) have A_rw: "?A = (\i = 0..i = 0.. i), \ (\ i)))" by (rule prod.permute[unfolded o_def, OF phi]) also have "... = (\i = 0.. i))" using fg_phi unfolding \_def unfolding o_def unfolding permutes_inverses(2)[OF phi] by auto finally have B_rw: "?B = (\i = 0.. i))" . have "(signof \) * ?A * (signof \) * ?B = (signof \) * (signof \) * ?A * ?B" by auto also have "... = signof (\ \ \) * ?A * ?B" unfolding signof_compose[OF phi sigma] by simp also have "... = signof \ * ?A * ?B" by (metis (no_types, lifting) \_def mult.commute o_inv_o_cancel permutes_inj phi sigma signof_compose) also have "... = signof \ * (\i = 0..i = 0.. i))" using A_rw B_rw by auto also have "... = signof \ * (\i = 0.. i))" by auto also have "... = weight f \" unfolding weight_def by simp finally show ?thesis .. qed lemma Z_good_fun_alt_sum: fixes g defines "Z_good_fun \ {f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj" shows "(\f\Z_good_fun. P f)= (\\\{\. \ permutes {0.. \))" proof - let ?f = "\\. g \ \" let ?P = "{\. \ permutes {0.. i < n" hence "xa i = i" unfolding permutes_def by auto thus "g (xa i) = i" using g i_ge_n unfolding F_inj_def by auto next fix xa assume "xa permutes {0.. xa) {0.. xb assume "\ permutes {0.. (\x. g (\ x)) ` {0.. {0.. {0..i. \ i < n \ x i = i" and inj_on_x: "inj_on x {0.. = "\i. if i x i = g j) else i" show "x \ (\) g ` {\. \ permutes {0..], rule conjI) have "?\ i = i" if i: "i \ {0..!j. ?\ j = i" for i proof (cases "i x a = g j) = i" proof (rule theI2) show "i < n \ x a = g i" using xa_gi True by auto fix xa assume "xa < n \ x a = g xa" thus "xa = i" by (metis (mono_tags, lifting) F_inj_def True atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq xa_gi) thus "xa = i" . qed thus ta: "?\ a = i" using a by auto fix j assume tj: "?\ j = i" show "j = a" proof (cases "j x j = g ja) = i" using tj True by auto have "?P (THE ja. ?P ja)" proof (rule theI) show "b < n \ x j = g b" using xj_gb b by auto fix xa assume "xa < n \ x j = g xa" thus "xa = b" by (metis (mono_tags, lifting) F_inj_def b atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq xj_gb) qed hence "x j = g i" unfolding the_ji by auto hence "x j = x a" using xa_gi by auto then show ?thesis using inj_on_x a True unfolding inj_on_def by auto next case False then show ?thesis using tj True by auto qed qed next case False note i_ge_n = False show ?thesis proof (rule ex1I[of _ i]) show "?\ i = i" using False by simp fix j assume tj: "?\ j = i" show "j = i" proof (cases "j x j = g ja) < n" proof (rule theI2) show "a < n \ x j = g a" using xj_ga a by auto fix xa assume a1: "xa < n \ x j = g xa" thus "xa = a" using F_inj_def a atLeast0LessThan g inj_on_eq_iff xj_ga by fastforce show "xa < n" by (simp add: a1) qed then show ?thesis using tj i_ge_n by auto next case False then show ?thesis using tj by auto qed qed qed ultimately show "?\ permutes {0.. ?\" proof - have "x xa = g (THE j. j < n \ x xa = g j)" if xa: "xa < n" for xa proof - obtain c where c: "c < n" and xxa_gc: "x xa = g c" by (metis (mono_tags, hide_lams) atLeast0LessThan imageE imageI lessThan_iff xa xg) show ?thesis proof (rule theI2) show c1: "c < n \ x xa = g c" using c xxa_gc by auto fix xb assume c2: "xb < n \ x xa = g xb" thus "xb = c" by (metis (mono_tags, lifting) F_inj_def c1 atLeast0LessThan g inj_onD lessThan_iff mem_Collect_eq) show "x xa = g xb" using c1 c2 by simp qed qed moreover have "x xa = g xa" if xa: "\ xa < n" for xa using g x1 x2 xa unfolding F_inj_def by simp ultimately show ?thesis unfolding o_def fun_eq_iff by auto qed qed qed have inj: "inj_on ?f ?P" proof (rule inj_onI) fix x y assume x: "x \ ?P" and y: "y \ ?P" and gx_gy: "g \ x = g \ y" have "x i = y i" for i proof (cases "i {0.. {0..f\Z_good_fun. P f) = (\f\?f`?P. P f)" using fP by simp also have "... = sum (P \ (\) g) {\. \ permutes {0..\ | \ permutes {0.. \))" by auto finally show ?thesis . qed lemma F_injI: assumes "f \ {0.. {0..i. i \ {0.. f i = i)" and "inj_on f {0.. F_inj" using assms unfolding F_inj_def by simp lemma F_inj_composition_permutation: assumes phi: "\ permutes {0.. F_inj" shows "g \ \ \ F_inj" proof (rule F_injI) show "g \ \ \ {0.. {0..i. i \ {0.. (g \ \) i = i" using g phi unfolding permutes_def F_inj_def by simp show "inj_on (g \ \) {0.. F_strict" shows "f \ F_inj" using f strict_mono_on_imp_inj_on unfolding F_strict_def F_inj_def by auto lemma one_step: assumes g1: "g \ F_strict" shows "det (submatrix A UNIV (g`{0..(x, y) \ Z_good g. weight x y)" (is "?lhs = ?rhs") proof - define Z_good_fun where "Z_good_fun = {f. f \ {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj" by (rule F_strict_imp_F_inj[OF g1]) have detA: "(\\\{\. \ permutes {0.. * (\i = 0.. i)))) = det (submatrix A UNIV (g`{0.. j \ g ` {0.. g ` {0.. j \ g ` {0.. carrier_mat n n" unfolding submatrix_def card_J using A by auto have "det (submatrix A UNIV (g`{0..p | p permutes {0..i = 0..\\{\. \ permutes {0.. * (\i = 0.. i))))" proof (rule sum.cong) fix x assume x: "x \ {\. \ permutes {0..i = 0..i = 0.. {0.. (g ` {0..i = 0..i = 0..\ \ ?Perm. signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i))) = (\\ \ ?Perm. signof (\) * (\i = 0.. i)))" if phi: "\ permutes {0.. proof - let ?h="\\. \ \ ?inv \" let ?g = "\\. signof (\) * (\i = 0.. i))" have "?h`?Perm = ?Perm" proof - have "\ \ ?inv \ permutes {0.. permutes {0.. using permutes_compose permutes_inv phi that by blast moreover have "x \ (\\. \ \ ?inv \) ` ?Perm" if "x permutes {0.. \ permutes {0.. \ \ ?inv \" using phi by auto ultimately show ?thesis unfolding image_def by auto qed ultimately show ?thesis by auto qed hence "(\\ \ ?Perm. ?g \) = (\\ \ ?h`?Perm. ?g \)" by simp also have "... = sum (?g \ ?h) ?Perm" proof (rule sum.reindex) show "inj_on (\\. \ \ ?inv \) {\. \ permutes {0..\ \ ?Perm. signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i)))" unfolding o_def by auto finally show ?thesis by simp qed have detB: "det (submatrix B (g`{0..\ \ ?Perm. signof \ * (\i = 0.. i)))" proof - have "{i. i < dim_row B \ i \ g ` {0.. g ` {0.. j \ g ` {0.. carrier_mat n n" unfolding submatrix_def using card_I B by auto have "det (submatrix B (g`{0..p \ ?Perm. signof p * (\i=0..\ \ ?Perm. signof \ * (\i = 0.. i)))" proof (rule sum.cong, rule refl) fix x assume x: "x \ {\. \ permutes {0..i=0..i=0.. {0.. (g ` {0..i = 0..i = 0..f\Z_good_fun. \\\?Perm. weight f \)" unfolding Z_good_def sum.cartesian_product Z_good_fun_def by blast also have "... = (\\\{\. \ permutes {0.. \))" unfolding Z_good_fun_def by (rule Z_good_fun_alt_sum[OF g]) also have "... = (\\\{\. \ permutes {0..\\{\. \ permutes {0.. * (\i = 0.. i))) * signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i)))" proof (rule sum.cong, simp, rule sum.cong, simp) fix \ \ assume phi: "\ \ ?Perm" and pi: "\ \ ?Perm" show "weight (g \ \) \ = signof \ * (\i = 0.. i))) * signof (\ \ ?inv \) * (\i = 0.. \ ?inv \) i))" proof (rule step_weight) show "g \ \ \ F_inj" by (rule F_inj_composition_permutation[OF _ g], insert phi, auto) show "g \ F" using g unfolding F_def F_inj_def by simp qed (insert phi pi, auto) qed also have "... = (\\\{\. \ permutes {0.. * (\i = 0.. i))) * (\\ | \ permutes {0.. \ ?inv \) * (\i = 0.. \ ?inv \) i))))" by (metis (mono_tags, lifting) Groups.mult_ac(1) semiring_0_class.sum_distrib_left sum.cong) also have "... = (\\ \ ?Perm. signof \ * (\i = 0.. i))) * (\\ \ ?Perm. signof \ * (\i = 0.. i))))" using detB_rw by auto also have "... = (\\ \ ?Perm. signof \ * (\i = 0.. i)))) * (\\ \ ?Perm. signof \ * (\i = 0.. i)))" by (simp add: semiring_0_class.sum_distrib_right) also have "... = ?lhs" unfolding detA detB .. finally show ?thesis .. qed lemma gather_by_strictness: "sum (\g. sum (\(f,\). weight f \) (Z_good g)) F_strict = sum (\g. det (submatrix A UNIV (g`{0.. F_strict" show "(\(x, y)\Z_good f. weight x y) = det (submatrix A UNIV (f ` {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite ?A" using rev_finite_subset by blast show "finite {\. \ permutes {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite F_inj" unfolding F_inj_def using rev_finite_subset by blast qed lemma finite_F_strict[simp]: "finite F_strict" proof - have finN: "finite {0..i\{0.. {0.. (\i. i \ {0.. f i = i)} = ?B" by auto have "?A\?B" by auto moreover have "finite ?B" using B finite_bounded_functions[OF finM finN] by auto ultimately show "finite F_strict" unfolding F_strict_def using rev_finite_subset by blast qed lemma nth_strict_mono: fixes f::"nat \ nat" assumes strictf: "strict_mono f" and i: "i ?I. a < f i} = i" using i proof (induct i) case 0 then show ?case by (auto simp add: strict_mono_less strictf) next case (Suc i) have i: "i < n" using Suc.prems by auto let ?J'="{a \ f ` {0.. f i" and 2: "f xa < f (Suc i)" show "f xa < f i" using 1 2 not_less_less_Suc_eq strict_mono_less strictf by fastforce next fix xa assume "f xa < f i" thus "f xa < f (Suc i)" using less_SucI strict_mono_less strictf by blast next show "f i \ f ` {0..card (f ` {0.. i) also have "... = pick ?I (card {a \ ?I. a < f i})" unfolding card_eq by simp also have "... = f i" by (rule pick_card_in_set, simp add: i) finally show ?thesis .. qed lemma nth_strict_mono_on: fixes f::"nat \ nat" assumes strictf: "strict_mono_on f {0.. ?I. a < f i} = i" using i proof (induct i) case 0 then show ?case by (auto, metis (no_types, lifting) atLeast0LessThan lessThan_iff less_Suc_eq not_less0 not_less_eq strict_mono_on_def strictf) next case (Suc i) have i: "i < n" using Suc.prems by auto let ?J'="{a \ f ` {0.. f i" and 2: "f xa < f (Suc i)" and 3: "xa < n" show "f xa < f i" by (metis (full_types) 1 2 3 antisym_conv3 atLeast0LessThan i lessThan_iff less_SucE order.asym strict_mono_onD strictf) next fix xa assume "f xa < f i" and "xa < n" thus "f xa < f (Suc i)" using less_SucI strictf by (metis (no_types, lifting) Suc.prems atLeast0LessThan lessI lessThan_iff less_trans strict_mono_onD) next show "f i \ f ` {0..card (f ` {0.. i) also have "... = pick ?I (card {a \ ?I. a < f i})" unfolding card_eq by simp also have "... = f i" by (rule pick_card_in_set, simp add: i) finally show ?thesis .. qed lemma strict_fun_eq: assumes f: "f \ F_strict" and g: "g \ F_strict" and fg: "f`{0.. F_inj" shows "strict_from_inj n f \ F" proof - { fix x assume x: "x < n" have inj_on: "inj_on f {0.. a \ f ` {0.. a \ f ` {0.. F_strict" if xa: "xa \ F_inj" for xa proof - have "strict_mono_on (strict_from_inj n xa) {0.. F_inj" shows "strict_from_inj n f ` {0.. a \ f ` {0.. a \ f ` {0.. f ` {0..card (f ` {0.. xa) finally show "strict_from_inj n f xa \ f ` {0.. strict_from_inj n f ` {0.. F_strict" shows "Z_good g = {x \ F_inj. strict_from_inj n x = g} \ {\. \ permutes {0.. {0.. {0.. (\i. i \ {0.. f i = i) \ inj_on f {0.. (f`{0.. F_inj. strict_from_inj n x = g}" proof (auto) fix f assume f: "f \ Z_good_fun" thus f_inj: "f \ F_inj" unfolding F_inj_def Z_good_fun_def by auto show "strict_from_inj n f = g" proof (rule strict_fun_eq[OF _ g]) show "strict_from_inj n f ` {0.. F_strict" using F_strict_def f_inj strict_from_inj_F_strict by blast qed next fix f assume f_inj: "f \ F_inj" and g_strict_f: "g = strict_from_inj n f" have "f xa \ g ` {0.. f ` {0.. Z_good_fun" using f_inj g_strict_f unfolding Z_good_fun_def F_inj_def by auto qed thus ?thesis unfolding Z_good_fun_def Z_good_def by simp qed lemma weight_0: "(\(f, \) \ Z_not_inj. weight f \) = 0" proof - let ?F="{f. (\i\{0.. {0.. (\i. i \ {0.. f i = i)}" let ?Perm = "{\. \ permutes {0..(f, \)\Z_not_inj. weight f \) = (\f \ F_not_inj. (\i = 0..r n n (\i. row B (f i))))" proof - have dim_row_rw: "dim_row (mat\<^sub>r n n (\i. col A (f i))) = n" for f by auto have dim_row_rw2: "dim_row (mat\<^sub>r n n (\i. Matrix.row B (f i))) = n" for f by auto have prod_rw: "(\i = 0.. i)) = (\i = 0.. i)" if f: "f \ F_not_inj" and pi: "\ \ ?Perm" for f \ proof (rule prod.cong, rule refl) fix x assume x: "x \ {0.. x < dim_col B" using x pi B by auto ultimately show "B $$ (f x, \ x) = Matrix.row B (f x) $v \ x" by (rule index_row[symmetric]) qed have sum_rw: "(\\ | \ permutes {0.. * (\i = 0.. i))) = det (mat\<^sub>r n n (\i. row B (f i)))" if f: "f \ F_not_inj" for f unfolding Determinant.det_def using dim_row_rw2 prod_rw f by auto have "(\(f, \)\Z_not_inj. weight f \) = (\f\F_not_inj.\\ \ ?Perm. weight f \)" unfolding Z_not_inj_def unfolding sum.cartesian_product unfolding F_not_inj_def by simp also have "... = (\f\F_not_inj. \\ | \ permutes {0.. * (\i = 0.. i)))" unfolding weight_def by simp also have "... = (\f\F_not_inj. (\i = 0..\ | \ permutes {0.. * (\i = 0.. i))))" by (rule sum.cong, rule refl, auto) (metis (no_types, lifting) mult.left_commute mult_hom.hom_sum sum.cong) also have "... = (\f \ F_not_inj. (\i = 0..r n n (\i. row B (f i))))" using sum_rw by auto finally show ?thesis by auto qed also have "... = 0" by (rule sum.neutral, insert det_not_inj_on[of _ n B], auto simp add: F_not_inj_def) finally show ?thesis . qed subsection \Final theorem\ lemma Cauchy_Binet1: shows "det (A*B) = sum (\f. det (submatrix A UNIV (f`{0..(f, \) \ Z_not_inj. weight f \) = 0" by (rule weight_0) let ?f = "strict_from_inj n" have sum_rw: "sum g F_inj = (\y \ F_strict. sum g {x \ F_inj. ?f x = y})" for g by (rule sum.group[symmetric], insert strict_from_inj_F_strict, auto) have Z_Union: "Z_inj \ Z_not_inj = Z n m" unfolding Z_def Z_not_inj_def Z_inj_def by auto have Z_Inter: "Z_inj \ Z_not_inj = {}" unfolding Z_def Z_not_inj_def Z_inj_def by auto have "det (A*B) = (\(f, \)\Z n m. weight f \)" using detAB_Znm[OF A B] unfolding weight_def by auto also have "... = (\(f, \)\Z_inj. weight f \) + (\(f, \)\Z_not_inj. weight f \)" by (metis Z_Inter Z_Union finite_Un finite_Znm sum.union_disjoint) also have "... = (\(f, \)\Z_inj. weight f \)" using sum0 by force also have "... = (\f \ F_inj. \\\{\. \ permutes {0..)" unfolding Z_inj_def unfolding F_inj_def sum.cartesian_product .. also have "... = (\y\F_strict. \f\{x \ F_inj. strict_from_inj n x = y}. sum (weight f) {\. \ permutes {0..y\F_strict. \(f,\)\({x \ F_inj. strict_from_inj n x = y} \ {\. \ permutes {0..)" unfolding F_inj_def sum.cartesian_product .. also have "... = sum (\g. sum (\(f,\). weight f \) (Z_good g)) F_strict" using Z_good_alt by auto also have "... = ?rhs" unfolding gather_by_strictness by simp finally show ?thesis . qed lemma Cauchy_Binet: "det (A*B) = (\I\{I. I\{0.. card I=n}. det (submatrix A UNIV I) * det (submatrix B I UNIV))" proof - let ?f="(\I. (\i. if i ?setI" and J: "J \ ?setI" and fI_fJ: "?f I = ?f J" have "x \ J" if x: "x \ I" for x by (metis (mono_tags) fI_fJ I J distinct_card in_set_conv_nth mem_Collect_eq sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite x) moreover have "x \ I" if x: "x \ J" for x by (metis (mono_tags) fI_fJ I J distinct_card in_set_conv_nth mem_Collect_eq sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite x) ultimately show "I = J" by auto qed have rw: "?f I ` {0.. ?setI" for I proof - have "sorted_list_of_set I ! xa \ I" if "xa < n" for xa by (metis (mono_tags, lifting) I distinct_card distinct_sorted_list_of_set mem_Collect_eq nth_mem set_sorted_list_of_set subset_eq_atLeast0_lessThan_finite that) moreover have "\xa\{0..I" for x by (metis (full_types) x I atLeast0LessThan distinct_card in_set_conv_nth mem_Collect_eq lessThan_iff sorted_list_of_set(1) sorted_list_of_set(3) subset_eq_atLeast0_lessThan_finite) ultimately show ?thesis unfolding image_def by auto qed have f_setI: "?f` ?setI = F_strict" proof - have "sorted_list_of_set I ! xa < m" if I: "I \ {0..xa < card I\ atLeast0LessThan distinct_card finite_atLeastLessThan lessThan_iff pick_in_set_le rev_finite_subset sorted_list_of_set(1) sorted_list_of_set(3) sorted_list_of_set_eq_pick subsetCE) moreover have "strict_mono_on (\i. if i < card I then sorted_list_of_set I ! i else i) {0.. {0..I \ {0.. atLeastLessThan_iff distinct_card finite_atLeastLessThan pick_mono_le rev_finite_subset sorted_list_of_set(1) sorted_list_of_set(3) sorted_list_of_set_eq_pick strict_mono_on_def) moreover have "x \ ?f ` {I. I \ {0.. card I = n}" if x1: "x \ {0.. {0..i. \ i < n \ x i = i" and s: "strict_mono_on x {0..i. if i < n then sorted_list_of_set (x ` {0..f. det (submatrix A UNIV (f ` {0.. ?f) {I. I \ {0.. card I = n}" unfolding Cauchy_Binet1 f_setI[symmetric] by (rule sum.reindex[OF inj_on]) also have "... = (\I\{I. I\{0.. card I=n}.det(submatrix A UNIV I)*det(submatrix B I UNIV))" by (rule sum.cong, insert rw, auto) finally show ?thesis . qed end end diff --git a/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy b/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy --- a/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy +++ b/thys/Smith_Normal_Form/SNF_Missing_Lemmas.thy @@ -1,1168 +1,1168 @@ (* Author: Jose Divasón Email: jose.divason@unirioja.es *) section \Missing results\ theory SNF_Missing_Lemmas imports Hermite.Hermite Mod_Type_Connect Jordan_Normal_Form.DL_Rank_Submatrix "List-Index.List_Index" begin text \This theory presents some missing lemmas that are required for the Smith normal form development. Some of them could be added to different AFP entries, such as the Jordan Normal Form AFP entry by Ren\'e Thiemann and Akihisa Yamada. However, not all the lemmas can be added directly, since some imports are required.\ hide_const (open) C hide_const (open) measure subsection \Miscellaneous lemmas\ lemma sum_two_rw: "(\i = 0..<2. f i) = (\i \ {0,1::nat}. f i)" by (rule sum.cong, auto) lemma sum_common_left: fixes f::"'a \ 'b::comm_ring_1" assumes "finite A" shows "sum (\i. c * f i) A = c * sum f A" by (simp add: mult_hom.hom_sum) lemma prod3_intro: assumes "fst A = a" and "fst (snd A) = b" and "snd (snd A) = c" shows "A = (a,b,c)" using assms by auto subsection \Transfer rules for the HMA\_Connect file of the Perron-Frobenius development\ hide_const (open) HMA_M HMA_I to_hma\<^sub>m from_hma\<^sub>m hide_fact (open) from_hma\<^sub>m_def from_hma_to_hma\<^sub>m HMA_M_def HMA_I_def dim_row_transfer_rule dim_col_transfer_rule context includes lifting_syntax begin lemma HMA_invertible_matrix[transfer_rule]: "((HMA_Connect.HMA_M :: _ \ 'a :: comm_ring_1 ^ 'n ^ 'n \ _) ===> (=)) invertible_mat invertible" proof (intro rel_funI, goal_cases) case (1 x y) note rel_xy[transfer_rule] = "1" have eq_dim: "dim_col x = dim_row x" using HMA_Connect.dim_col_transfer_rule HMA_Connect.dim_row_transfer_rule rel_xy by fastforce moreover have "\A'. y ** A' = Finite_Cartesian_Product.mat 1 \ A' ** y = Finite_Cartesian_Product.mat 1" if xB: "x * B = 1\<^sub>m (dim_row x)" and Bx: "B * x = 1\<^sub>m (dim_row B)" for B proof - let ?A' = "HMA_Connect.to_hma\<^sub>m B:: 'a :: comm_ring_1 ^ 'n ^ 'n" have rel_BA[transfer_rule]: "HMA_M B ?A'" by (metis (no_types, lifting) Bx HMA_M_def eq_dim carrier_mat_triv dim_col_mat(1) from_hma\<^sub>m_def from_hma_to_hma\<^sub>m index_mult_mat(3) index_one_mat(3) rel_xy xB) have [simp]: "dim_row B = CARD('n)" using dim_row_transfer_rule rel_BA by blast have [simp]: "dim_row x = CARD('n)" using dim_row_transfer_rule rel_xy by blast have "y ** ?A' = Finite_Cartesian_Product.mat 1" using xB by (transfer, simp) moreover have "?A' ** y = Finite_Cartesian_Product.mat 1" using Bx by (transfer, simp) ultimately show ?thesis by blast qed moreover have "\B. x * B = 1\<^sub>m (dim_row x) \ B * x = 1\<^sub>m (dim_row B)" if yA: "y ** A' = Finite_Cartesian_Product.mat 1" and Ay: "A' ** y = Finite_Cartesian_Product.mat 1" for A' proof - let ?B = "(from_hma\<^sub>m A')" have [simp]: "dim_row x = CARD('n)" using dim_row_transfer_rule rel_xy by blast have [transfer_rule]: "HMA_M ?B A'" by (simp add: HMA_M_def) hence [simp]: "dim_row ?B = CARD('n)" using dim_row_transfer_rule by auto have "x * ?B = 1\<^sub>m (dim_row x)" using yA by (transfer', auto) moreover have "?B * x = 1\<^sub>m (dim_row ?B)" using Ay by (transfer', auto) ultimately show ?thesis by auto qed ultimately show ?case unfolding invertible_mat_def invertible_def inverts_mat_def by auto qed end subsection \Lemmas obtained from HOL Analysis using local type definitions\ thm Cartesian_Space.invertible_mult (*In HOL Analysis*) thm invertible_iff_is_unit (*In HOL Analysis*) thm det_non_zero_imp_unit (*In JNF, but only for fields*) thm mat_mult_left_right_inverse (*In JNF, but only for fields*) lemma invertible_mat_zero: assumes A: "A \ carrier_mat 0 0" shows "invertible_mat A" using A unfolding invertible_mat_def inverts_mat_def one_mat_def times_mat_def scalar_prod_def Matrix.row_def col_def carrier_mat_def by (auto, metis (no_types, lifting) cong_mat not_less_zero) lemma invertible_mult_JNF: fixes A::"'a::comm_ring_1 mat" assumes A: "A\carrier_mat n n" and B: "B\carrier_mat n n" and inv_A: "invertible_mat A" and inv_B: "invertible_mat B" shows "invertible_mat (A*B)" proof (cases "n = 0") case True then show ?thesis using assms by (simp add: invertible_mat_zero) next case False then show ?thesis using invertible_mult[where ?'a="'a::comm_ring_1", where ?'b="'n::finite", where ?'c="'n::finite", where ?'d="'n::finite", untransferred, cancel_card_constraint, OF assms] by auto qed lemma invertible_iff_is_unit_JNF: assumes A: "A \ carrier_mat n n" shows "invertible_mat A \ (Determinant.det A) dvd 1" proof (cases "n=0") case True then show ?thesis using det_dim_zero invertible_mat_zero A by auto next case False then show ?thesis using invertible_iff_is_unit[untransferred, cancel_card_constraint] A by auto qed subsection \Lemmas about matrices, submatrices and determinants\ (*This is a generalization of thm mat_mult_left_right_inverse*) thm mat_mult_left_right_inverse lemma mat_mult_left_right_inverse: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and AB: "A * B = 1\<^sub>m n" shows "B * A = 1\<^sub>m n" proof - have "Determinant.det (A * B) = Determinant.det (1\<^sub>m n)" using AB by auto hence "Determinant.det A * Determinant.det B = 1" using Determinant.det_mult[OF A B] det_one by auto hence det_A: "(Determinant.det A) dvd 1" and det_B: "(Determinant.det B) dvd 1" using dvd_triv_left dvd_triv_right by metis+ hence inv_A: "invertible_mat A" and inv_B: "invertible_mat B" using A B invertible_iff_is_unit_JNF by blast+ obtain B' where inv_BB': "inverts_mat B B'" and inv_B'B: "inverts_mat B' B" using inv_B unfolding invertible_mat_def by auto have B'_carrier: "B' \ carrier_mat n n" by (metis B inv_B'B inv_BB' carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(3) index_one_mat(3) inverts_mat_def) have "B * A * B = B" using A AB B by auto hence "B * A * (B * B') = B * B'" by (smt A AB B B'_carrier assoc_mult_mat carrier_matD(1) inv_BB' inverts_mat_def one_carrier_mat) thus ?thesis by (metis A B carrier_matD(1) carrier_matD(2) index_mult_mat(3) inv_BB' inverts_mat_def right_mult_one_mat') qed context comm_ring_1 begin lemma col_submatrix_UNIV: assumes "j < card {i. i < dim_col A \ i \ J}" shows "col (submatrix A UNIV J) j = col A (pick J j)" proof (rule eq_vecI) show dim_eq:"dim_vec (col (submatrix A UNIV J) j) = dim_vec (col A (pick J j))" by (simp add: dim_submatrix(1)) fix i assume "i < dim_vec (col A (pick J j))" show "col (submatrix A UNIV J) j $v i = col A (pick J j) $v i" by (smt Collect_cong assms col_def dim_col dim_eq dim_submatrix(1) eq_vecI index_vec pick_UNIV submatrix_index) qed lemma submatrix_split2: "submatrix A I J = submatrix (submatrix A I UNIV) UNIV J" (is "?lhs = ?rhs") proof (rule eq_matI) show dr: "dim_row ?lhs = dim_row ?rhs" by (simp add: dim_submatrix(1)) show dc: "dim_col ?lhs = dim_col ?rhs" by (simp add: dim_submatrix(2)) fix i j assume i: "i < dim_row ?rhs" and j: "j < dim_col ?rhs" have "?rhs $$ (i, j) = (submatrix A I UNIV) $$ (pick UNIV i, pick J j)" proof (rule submatrix_index) show "i < card {i. i < dim_row (submatrix A I UNIV) \ i \ UNIV}" by (metis (full_types) dim_submatrix(1) i) show "j < card {j. j < dim_col (submatrix A I UNIV) \ j \ J}" by (metis (full_types) dim_submatrix(2) j) qed also have "... = A $$ (pick I (pick UNIV i), pick UNIV (pick J j))" proof (rule submatrix_index) show "pick UNIV i < card {i. i < dim_row A \ i \ I}" by (metis (full_types) dr dim_submatrix(1) i pick_UNIV) show "pick J j < card {j. j < dim_col A \ j \ UNIV}" by (metis (full_types) dim_submatrix(2) j pick_le) qed also have "... = ?lhs $$ (i,j)" proof (unfold pick_UNIV, rule submatrix_index[symmetric]) show "i < card {i. i < dim_row A \ i \ I}" by (metis (full_types) dim_submatrix(1) dr i) show "j < card {j. j < dim_col A \ j \ J}" by (metis (full_types) dim_submatrix(2) dc j) qed finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" .. qed lemma submatrix_mult: "submatrix (A*B) I J = submatrix A I UNIV * submatrix B UNIV J" (is "?lhs = ?rhs") proof (rule eq_matI) show "dim_row ?lhs = dim_row ?rhs" unfolding submatrix_def by auto show "dim_col ?lhs = dim_col ?rhs" unfolding submatrix_def by auto fix i j assume i: "i < dim_row ?rhs" and j: "j < dim_col ?rhs" have i1: "i < card {i. i < dim_row (A * B) \ i \ I}" by (metis (full_types) dim_submatrix(1) i index_mult_mat(2)) have j1: "j < card {j. j < dim_col (A * B) \ j \ J}" by (metis dim_submatrix(2) index_mult_mat(3) j) have pi: "pick I i < dim_row A" using i1 pick_le by auto have pj: "pick J j < dim_col B" using j1 pick_le by auto have row_rw: "Matrix.row (submatrix A I UNIV) i = Matrix.row A (pick I i)" using i1 row_submatrix_UNIV by auto have col_rw: "col (submatrix B UNIV J) j = col B (pick J j)" using j1 col_submatrix_UNIV by auto have "?lhs $$ (i,j) = (A*B) $$ (pick I i, pick J j)" by (rule submatrix_index[OF i1 j1]) also have "... = Matrix.row A (pick I i) \ col B (pick J j)" by (rule index_mult_mat(1)[OF pi pj]) also have "... = Matrix.row (submatrix A I UNIV) i \ col (submatrix B UNIV J) j" using row_rw col_rw by simp also have "... = (?rhs) $$ (i,j)" by (rule index_mult_mat[symmetric], insert i j, auto) finally show "?lhs $$ (i, j) = ?rhs $$ (i, j)" . qed lemma det_singleton: assumes A: "A \ carrier_mat 1 1" shows "det A = A $$ (0,0)" using A unfolding carrier_mat_def Determinant.det_def by auto lemma submatrix_singleton_index: assumes A: "A \ carrier_mat n m" and an: "a < n" and bm: "b < m" shows "submatrix A {a} {b} $$ (0,0) = A $$ (a,b)" proof - have a: "{i. i = a \ i < dim_row A} = {a}" using an A unfolding carrier_mat_def by auto have b: "{i. i = b \ i < dim_col A} = {b}" using bm A unfolding carrier_mat_def by auto have "submatrix A {a} {b} $$ (0,0) = A $$ (pick {a} 0,pick {b} 0)" by (rule submatrix_index, insert a b, auto) moreover have "pick {a} 0 = a" by (auto, metis (full_types) LeastI) moreover have "pick {b} 0 = b" by (auto, metis (full_types) LeastI) ultimately show ?thesis by simp qed end lemma det_not_inj_on: assumes not_inj_on: "\ inj_on f {0..r n n (\i. Matrix.row B (f i))) = 0" proof - obtain i j where i: "ij" using not_inj_on unfolding inj_on_def by auto show ?thesis proof (rule det_identical_rows[OF _ ij i j]) let ?B="(mat\<^sub>r n n (\i. row B (f i)))" show "row ?B i = row ?B j" proof (rule eq_vecI, auto) fix ia assume ia: "ia < n" have "row ?B i $ ia = ?B $$ (i, ia)" by (rule index_row(1), insert i ia, auto) also have "... = ?B $$ (j, ia)" by (simp add: fi_fj i ia j) also have "... = row ?B j $ ia" by (rule index_row(1)[symmetric], insert j ia, auto) finally show "row ?B i $ ia = row (mat\<^sub>r n n (\i. row B (f i))) j $ ia" by simp qed show "mat\<^sub>r n n (\i. Matrix.row B (f i)) \ carrier_mat n n" by auto qed qed lemma mat_row_transpose: "(mat\<^sub>r nr nc f)\<^sup>T = mat nc nr (\(i,j). vec_index (f j) i)" by (rule eq_matI, auto) lemma obtain_inverse_matrix: assumes A: "A \ carrier_mat n n" and i: "invertible_mat A" obtains B where "inverts_mat A B" and "inverts_mat B A" and "B \ carrier_mat n n" proof - have "(\B. inverts_mat A B \ inverts_mat B A)" using i unfolding invertible_mat_def by auto from this obtain B where AB: "inverts_mat A B" and BA: "inverts_mat B A" by auto moreover have "B \ carrier_mat n n" using A AB BA unfolding carrier_mat_def inverts_mat_def by (auto, metis index_mult_mat(3) index_one_mat(3))+ ultimately show ?thesis using that by blast qed lemma invertible_mat_smult_mat: fixes A :: "'a::comm_ring_1 mat" assumes inv_A: "invertible_mat A" and k: "k dvd 1" shows "invertible_mat (k \\<^sub>m A)" proof - obtain n where A: "A \ carrier_mat n n" using inv_A unfolding invertible_mat_def by auto have det_dvd_1: "Determinant.det A dvd 1" using inv_A invertible_iff_is_unit_JNF[OF A] by auto have "Determinant.det (k \\<^sub>m A) = k ^ dim_col A * Determinant.det A" by simp also have "... dvd 1" by (rule unit_prod, insert k det_dvd_1 dvd_power_same, force+) finally show ?thesis using invertible_iff_is_unit_JNF by (metis A smult_carrier_mat) qed lemma invertible_mat_one[simp]: "invertible_mat (1\<^sub>m n)" unfolding invertible_mat_def using inverts_mat_def by fastforce lemma four_block_mat_dim0: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n 0" and C: "C \ carrier_mat 0 n" and D: "D \ carrier_mat 0 0" shows "four_block_mat A B C D = A" unfolding four_block_mat_def using assms by auto lemma det_four_block_mat_lower_right_id: assumes A: "A \ carrier_mat m m" and B: "B = 0\<^sub>m m (n-m)" and C: "C = 0\<^sub>m (n-m) m" and D: "D = 1\<^sub>m (n-m)" and "n>m" shows "Determinant.det (four_block_mat A B C D) = Determinant.det A" using assms proof (induct n arbitrary: A B C D) case 0 then show ?case by auto next case (Suc n) let ?block = "(four_block_mat A B C D)" let ?B = "Matrix.mat m (n-m) (\(i,j). 0)" let ?C = "Matrix.mat (n-m) m (\(i,j). 0)" let ?D = "1\<^sub>m (n-m)" have mat_eq: "(mat_delete ?block n n) = four_block_mat A ?B ?C ?D" (is "?lhs = ?rhs") proof (rule eq_matI) fix i j assume i: "i < dim_row (four_block_mat A ?B ?C ?D)" and j: "j < dim_col (four_block_mat A ?B ?C ?D)" let ?f = " (if i < dim_row A then if j < dim_col A then A $$ (i, j) else B $$ (i, j - dim_col A) else if j < dim_col A then C $$ (i - dim_row A, j) else D $$ (i - dim_row A, j - dim_col A))" let ?g = "(if i < dim_row A then if j < dim_col A then A $$ (i, j) else ?B $$ (i, j - dim_col A) else if j < dim_col A then ?C $$ (i - dim_row A, j) else ?D $$ (i - dim_row A, j - dim_col A))" have "(mat_delete ?block n n) $$ (i,j) = ?block $$ (i,j)" using i j Suc.prems unfolding mat_delete_def by auto also have "... = ?f" by (rule index_mat_four_block, insert Suc.prems i j, auto) also have "... = ?g" using i j Suc.prems by auto also have "... = four_block_mat A ?B ?C ?D $$ (i,j)" by (rule index_mat_four_block[symmetric], insert Suc.prems i j, auto) finally show "?lhs $$ (i,j) = ?rhs $$ (i,j)" . qed (insert Suc.prems, auto) have nn_1: "?block $$ (n, n) = 1" using Suc.prems by auto have rw0: "(\i {..ii carrier_mat 1 n" and B: "B \ carrier_mat m n" and m0: "m \ 0" and r: "Matrix.row A 0 = Matrix.row B 0" shows "Matrix.row (A * V) 0 = Matrix.row (B * V) 0" proof (rule eq_vecI) show "dim_vec (Matrix.row (A * V) 0) = dim_vec (Matrix.row (B * V) 0)" using A B r by auto fix i assume i: "i < dim_vec (Matrix.row (B * V) 0)" have "Matrix.row (A * V) 0 $v i = (A * V) $$ (0,i)" by (rule index_row, insert i A, auto) also have "... = Matrix.row A 0 \ col V i" by (rule index_mult_mat, insert A i, auto) also have "... = Matrix.row B 0 \ col V i" using r by auto also have "... = (B * V) $$ (0,i)" by (rule index_mult_mat[symmetric], insert m0 B i, auto) also have "... = Matrix.row (B * V) 0 $v i" by (rule index_row[symmetric], insert i B m0, auto) finally show "Matrix.row (A * V) 0 $v i = Matrix.row (B * V) 0 $v i" . qed lemma smult_mat_mat_one_element: assumes A: "A \ carrier_mat 1 1" and B: "B \ carrier_mat 1 n" shows "A * B = A $$ (0,0) \\<^sub>m B" proof (rule eq_matI) fix i j assume i: "i < dim_row (A $$ (0, 0) \\<^sub>m B)" and j: "j < dim_col (A $$ (0, 0) \\<^sub>m B)" have i0: "i = 0" using A B i by auto have "(A * B) $$ (i, j) = Matrix.row A i \ col B j" by (rule index_mult_mat, insert i j A B, auto) also have "... = Matrix.row A i $v 0 * col B j $v 0" unfolding scalar_prod_def using B by auto also have "... = A$$(i,i) * B$$(i,j)" using A i i0 j by auto also have "... = (A $$ (i, i) \\<^sub>m B) $$ (i, j)" unfolding i by (rule index_smult_mat[symmetric], insert i j B, auto) finally show "(A * B) $$ (i, j) = (A $$ (0, 0) \\<^sub>m B) $$ (i, j)" using i0 by simp qed (insert A B, auto) lemma determinant_one_element: assumes A: "A \ carrier_mat 1 1" shows "Determinant.det A = A $$ (0,0)" proof - have "Determinant.det A = prod_list (diag_mat A)" by (rule det_upper_triangular[OF _ A], insert A, unfold upper_triangular_def, auto) also have "... = A $$ (0,0)" using A unfolding diag_mat_def by auto finally show ?thesis . qed lemma invertible_mat_transpose: assumes inv_A: "invertible_mat (A::'a::comm_ring_1 mat)" shows "invertible_mat A\<^sup>T" proof - obtain n where A: "A \ carrier_mat n n" using inv_A unfolding invertible_mat_def square_mat.simps by auto hence At: "A\<^sup>T \ carrier_mat n n" by simp have "Determinant.det A\<^sup>T = Determinant.det A" by (metis Determinant.det_def Determinant.det_transpose carrier_matI index_transpose_mat(2) index_transpose_mat(3)) also have "... dvd 1" using invertible_iff_is_unit_JNF[OF A] inv_A by simp finally show ?thesis using invertible_iff_is_unit_JNF[OF At] by auto qed lemma dvd_elements_mult_matrix_left: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat m n" and P: "P \ carrier_mat m m" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (P*A)$$(i,j))" proof - have "x dvd (P * A) $$ (i, j)" if i: "i < m" and j: "j < n" for i j proof - have "(P * A) $$ (i, j) = (\ia = 0..ia = 0.. carrier_mat m n" and Q: "Q \ carrier_mat n n" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (A*Q)$$(i,j))" proof - have "x dvd (A*Q) $$ (i, j)" if i: "i < m" and j: "j < n" for i j proof - have "(A*Q) $$ (i, j) = (\ia = 0..ia = 0.. carrier_mat m n" and P: "P \ carrier_mat m m" and Q: "Q \ carrier_mat n n" and x: "(\i j. i j x dvd A$$(i,j))" shows "(\i j. i j x dvd (P*A*Q)$$(i,j))" using dvd_elements_mult_matrix_left[OF A P x] by (meson P A Q dvd_elements_mult_matrix_right mult_carrier_mat) definition append_cols :: "'a :: zero mat \ 'a mat \ 'a mat" (infixr "@\<^sub>c" 65)where "A @\<^sub>c B = four_block_mat A B (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col B))" lemma append_cols_carrier[simp,intro]: "A \ carrier_mat n a \ B \ carrier_mat n b \ (A @\<^sub>c B) \ carrier_mat n (a+b)" unfolding append_cols_def by auto lemma append_cols_mult_left: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" and P: "P \ carrier_mat n n" shows "P * (A @\<^sub>c B) = (P*A) @\<^sub>c (P*B)" proof - let ?P = "four_block_mat P (0\<^sub>m n 0) (0\<^sub>m 0 n) (0\<^sub>m 0 0)" have "P = ?P" by (rule eq_matI, auto) hence "P * (A @\<^sub>c B) = ?P * (A @\<^sub>c B)" by simp also have "?P * (A @\<^sub>c B) = four_block_mat (P * A + 0\<^sub>m n 0 * 0\<^sub>m 0 (dim_col A)) (P * B + 0\<^sub>m n 0 * 0\<^sub>m 0 (dim_col B)) (0\<^sub>m 0 n * A + 0\<^sub>m 0 0 * 0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 n * B + 0\<^sub>m 0 0 * 0\<^sub>m 0 (dim_col B))" unfolding append_cols_def by (rule mult_four_block_mat, insert A B P, auto) also have "... = four_block_mat (P * A) (P * B) (0\<^sub>m 0 (dim_col (P*A))) (0\<^sub>m 0 (dim_col (P*B)))" by (rule cong_four_block_mat, insert P, auto) also have "... = (P*A) @\<^sub>c (P*B)" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_mult_right_id: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n 1" and B: "B \ carrier_mat n (m-1)" and C: "C = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" and D: "D \ carrier_mat (m-1) (m-1)" shows "(A @\<^sub>c B) * C = A @\<^sub>c (B * D)" proof - let ?C = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" have "(A @\<^sub>c B) * C = (A @\<^sub>c B) * ?C" unfolding C by auto also have "... = four_block_mat A B (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col B)) * ?C" unfolding append_cols_def by auto also have "... = four_block_mat (A * 1\<^sub>m 1 + B * 0\<^sub>m (m - 1) 1) (A * 0\<^sub>m 1 (m - 1) + B * D) (0\<^sub>m 0 (dim_col A) * 1\<^sub>m 1 + 0\<^sub>m 0 (dim_col B) * 0\<^sub>m (m - 1) 1) (0\<^sub>m 0 (dim_col A) * 0\<^sub>m 1 (m - 1) + 0\<^sub>m 0 (dim_col B) * D)" by (rule mult_four_block_mat, insert assms, auto) also have "... = four_block_mat A (B * D) (0\<^sub>m 0 (dim_col A)) (0\<^sub>m 0 (dim_col (B*D)))" by (rule cong_four_block_mat, insert assms, auto) also have "... = A @\<^sub>c (B * D)" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_mult_right_id2: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n a" and B: "B \ carrier_mat n b" and C: "C = four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" and D: "D \ carrier_mat a a" shows "(A @\<^sub>c B) * C = (A * D) @\<^sub>c B" proof - let ?C = "four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" have "(A @\<^sub>c B) * C = (A @\<^sub>c B) * ?C" unfolding C by auto also have "... = four_block_mat A B (0\<^sub>m 0 a) (0\<^sub>m 0 b) * ?C" unfolding append_cols_def using A B by auto also have "... = four_block_mat (A * D + B * 0\<^sub>m b a) (A * 0\<^sub>m a b + B * 1\<^sub>m b) (0\<^sub>m 0 a * D + 0\<^sub>m 0 b * 0\<^sub>m b a) (0\<^sub>m 0 a * 0\<^sub>m a b + 0\<^sub>m 0 b * 1\<^sub>m b)" by (rule mult_four_block_mat, insert A B C D, auto) also have "... = four_block_mat (A * D) B (0\<^sub>m 0 (dim_col (A*D))) (0\<^sub>m 0 (dim_col B))" by (rule cong_four_block_mat, insert assms, auto) also have "... = (A * D) @\<^sub>c B" unfolding append_cols_def by auto finally show ?thesis . qed lemma append_cols_nth: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" and i: "ic B) $$ (i, j) = (if j < dim_col A then A $$(i,j) else B$$(i,j-a))" (is "?lhs = ?rhs") proof - let ?C = "(0\<^sub>m 0 (dim_col A))" let ?D = "(0\<^sub>m 0 (dim_col B))" have i2: "i < dim_row A + dim_row ?D" using i A by auto have j2: "j < dim_col A + dim_col (0\<^sub>m 0 (dim_col B))" using j B A by auto have "(A @\<^sub>c B) $$ (i, j) = four_block_mat A B ?C ?D $$ (i, j)" unfolding append_cols_def by auto also have "... = (if i < dim_row A then if j < dim_col A then A $$ (i, j) else B $$ (i, j - dim_col A) else if j < dim_col A then ?C $$ (i - dim_row A, j) else 0\<^sub>m 0 (dim_col B) $$ (i - dim_row A, j - dim_col A))" by (rule index_mat_four_block(1)[OF i2 j2]) also have "... = ?rhs" using i A by auto finally show ?thesis . qed lemma append_cols_split: assumes d: "dim_col A > 0" shows "A = mat_of_cols (dim_row A) [col A 0] @\<^sub>c mat_of_cols (dim_row A) (map (col A) [1..c ?A2") proof (rule eq_matI) fix i j assume i: "i < dim_row (?A1 @\<^sub>c ?A2)" and j: "j < dim_col (?A1 @\<^sub>c ?A2)" have "(?A1 @\<^sub>c ?A2) $$ (i, j) = (if j < dim_col ?A1 then ?A1 $$(i,j) else ?A2$$(i,j-(dim_col ?A1)))" by (rule append_cols_nth, insert i j, auto simp add: append_cols_def) also have "... = A $$ (i,j)" proof (cases "j< dim_col ?A1") case True then show ?thesis by (metis One_nat_def Suc_eq_plus1 add.right_neutral append_cols_def col_def i index_mat_four_block(2) index_vec index_zero_mat(2) less_one list.size(3) list.size(4) mat_of_cols_Cons_index_0 mat_of_cols_carrier(2) mat_of_cols_carrier(3)) next case False then show ?thesis by (metis (no_types, lifting) Suc_eq_plus1 Suc_less_eq Suc_pred add_diff_cancel_right' append_cols_def diff_zero i index_col index_mat_four_block(2) index_mat_four_block(3) index_zero_mat(2) index_zero_mat(3) j length_map length_upt linordered_semidom_class.add_diff_inverse list.size(3) list.size(4) mat_of_cols_carrier(2) mat_of_cols_carrier(3) mat_of_cols_index nth_map_upt plus_1_eq_Suc upt_0) qed finally show "A $$ (i, j) = (?A1 @\<^sub>c ?A2) $$ (i, j)" .. qed (auto simp add: append_cols_def d) lemma append_rows_nth: assumes A: "A \ carrier_mat a n" and B: "B \ carrier_mat b n" and i: "ir B) $$ (i, j) = (if i < dim_row A then A $$(i,j) else B$$(i-a,j))" (is "?lhs = ?rhs") proof - let ?C = "(0\<^sub>m (dim_row A) 0)" let ?D = "(0\<^sub>m (dim_row B) 0)" have i2: "i < dim_row A + dim_row ?D" using i j A B by auto have j2: "j < dim_col A + dim_col ?D" using i j A B by auto have "(A @\<^sub>r B) $$ (i, j) = four_block_mat A ?C B ?D $$ (i, j)" unfolding append_rows_def by auto also have "... = (if i < dim_row A then if j < dim_col A then A $$ (i, j) else ?C $$ (i, j - dim_col A) else if j < dim_col A then B $$ (i - dim_row A, j) else ?D $$ (i - dim_row A, j - dim_col A))" by (rule index_mat_four_block(1)[OF i2 j2]) also have "... = ?rhs" using i A j B by auto finally show ?thesis . qed lemma append_rows_split: assumes k: "k\dim_row A" shows "A = (mat_of_rows (dim_col A) [Matrix.row A i. i \ [0..r (mat_of_rows (dim_col A) [Matrix.row A i. i \ [k..r ?A2") proof (rule eq_matI) have "(?A1 @\<^sub>r ?A2) \ carrier_mat (k + (dim_row A-k)) (dim_col A)" by (rule carrier_append_rows, insert k, auto) hence A1_A2: "(?A1 @\<^sub>r ?A2) \ carrier_mat (dim_row A) (dim_col A)" using k by simp thus "dim_row A = dim_row (?A1 @\<^sub>r ?A2)" and "dim_col A = dim_col (?A1 @\<^sub>r ?A2)" by auto fix i j assume i: "i < dim_row (?A1 @\<^sub>r ?A2)" and j: "j < dim_col (?A1 @\<^sub>r ?A2)" have "(?A1 @\<^sub>r ?A2) $$ (i, j) = (if i < dim_row ?A1 then ?A1 $$(i,j) else ?A2$$(i-(dim_row ?A1),j))" by (rule append_rows_nth, insert k i j, auto simp add: append_rows_def) also have "... = A $$ (i,j)" proof (cases "ir ?A2) $$ (i,j)" by simp qed lemma transpose_mat_append_rows: assumes A: "A \ carrier_mat a n" and B: "B \ carrier_mat b n" shows "(A @\<^sub>r B)\<^sup>T = A\<^sup>T @\<^sub>c B\<^sup>T" by (smt append_cols_def append_rows_def A B carrier_matD(1) index_transpose_mat(3) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) lemma transpose_mat_append_cols: assumes A: "A \ carrier_mat n a" and B: "B \ carrier_mat n b" shows "(A @\<^sub>c B)\<^sup>T = A\<^sup>T @\<^sub>r B\<^sup>T" by (metis Matrix.transpose_transpose A B carrier_matD(1) carrier_mat_triv index_transpose_mat(3) transpose_mat_append_rows) lemma append_rows_mult_right: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat a n" and B: "B \ carrier_mat b n" and Q: "Q\ carrier_mat n n" shows "(A @\<^sub>r B) * Q = (A * Q) @\<^sub>r (B*Q)" proof - have "transpose_mat ((A @\<^sub>r B) * Q) = Q\<^sup>T * (A @\<^sub>r B)\<^sup>T" by (rule transpose_mult, insert A B Q, auto) also have "... = Q\<^sup>T * (A\<^sup>T @\<^sub>c B\<^sup>T)" using transpose_mat_append_rows assms by metis also have "... = Q\<^sup>T * A\<^sup>T @\<^sub>c Q\<^sup>T * B\<^sup>T" using append_cols_mult_left assms by (metis transpose_carrier_mat) also have "transpose_mat ... = (A * Q) @\<^sub>r (B*Q)" by (smt A B Matrix.transpose_mult Matrix.transpose_transpose append_cols_def append_rows_def Q carrier_mat_triv index_mult_mat(2) index_transpose_mat(2) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) finally show ?thesis by simp qed lemma append_rows_mult_left_id: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat 1 n" and B: "B \ carrier_mat (m-1) n" and C: "C = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) D" and D: "D \ carrier_mat (m-1) (m-1)" shows "C * (A @\<^sub>r B) = A @\<^sub>r (D * B)" proof - have "transpose_mat (C * (A @\<^sub>r B)) = (A @\<^sub>r B)\<^sup>T * C\<^sup>T" by (metis (no_types, lifting) B C D Matrix.transpose_mult append_rows_def A carrier_matD carrier_mat_triv index_mat_four_block(2,3) index_zero_mat(2) one_carrier_mat) also have "... = (A\<^sup>T @\<^sub>c B\<^sup>T) * C\<^sup>T" using transpose_mat_append_rows[OF A B] by auto also have "... = A\<^sup>T @\<^sub>c (B\<^sup>T * D\<^sup>T)" by (rule append_cols_mult_right_id, insert A B C D, auto) also have "transpose_mat ... = A @\<^sub>r (D * B)" by (smt B D Matrix.transpose_mult Matrix.transpose_transpose append_cols_def append_rows_def A carrier_matD(2) carrier_mat_triv index_mult_mat(3) index_transpose_mat(3) transpose_four_block_mat zero_carrier_mat zero_transpose_mat) finally show ?thesis by auto qed lemma append_rows_mult_left_id2: assumes A: "(A::'a::comm_semiring_1 mat) \ carrier_mat a n" and B: "B \ carrier_mat b n" and C: "C = four_block_mat D (0\<^sub>m a b) (0\<^sub>m b a) (1\<^sub>m b)" and D: "D \ carrier_mat a a" shows "C * (A @\<^sub>r B) = (D * A) @\<^sub>r B" proof - have "(C * (A @\<^sub>r B))\<^sup>T = (A @\<^sub>r B)\<^sup>T * C\<^sup>T" by (rule transpose_mult, insert assms, auto) also have "... = (A\<^sup>T @\<^sub>c B\<^sup>T) * C\<^sup>T" by (metis A B transpose_mat_append_rows) also have "... = (A\<^sup>T * D\<^sup>T @\<^sub>c B\<^sup>T)" by (rule append_cols_mult_right_id2, insert assms, auto) also have "...\<^sup>T = (D * A) @\<^sub>r B" by (metis A B D transpose_mult transpose_transpose mult_carrier_mat transpose_mat_append_rows) finally show ?thesis by simp qed lemma four_block_mat_preserves_column: assumes A: "(A::'a::semiring_1 mat) \ carrier_mat n m" and B: "B = four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 (m - 1)) (0\<^sub>m (m - 1) 1) C" and C: "C \ carrier_mat (m-1) (m-1)" and i: "ic ?A2" by (rule append_cols_split[of A, unfolded n2], insert m A, auto) hence "A * B = (?A1 @\<^sub>c ?A2) * B" by simp also have "... = ?A1 @\<^sub>c (?A2 * C)" by (rule append_cols_mult_right_id[OF _ _ B C], insert A, auto) also have "... $$ (i,0) = ?A1 $$ (i,0)" using append_cols_nth by (simp add: append_cols_def i) also have "... = A $$ (i,0)" by (metis A i carrier_matD(1) col_def index_vec mat_of_cols_Cons_index_0) finally show ?thesis . qed definition "lower_triangular A = (\i j. i < j \ i < dim_row A \ j < dim_col A \ A $$ (i,j) = 0)" lemma lower_triangular_index: assumes "lower_triangular A" "i carrier_mat n n" shows "A * (k \\<^sub>m (1\<^sub>m n)) = (k \\<^sub>m (1\<^sub>m n)) * A" proof - have "(\ia = 0..ia = 0..ia \ ({0..ia \ ({0..ia \ ({0..ia \ ({0.. carrier_mat 2 2" shows "Determinant.det A = A$$(0,0) * A $$ (1,1) - A$$(0,1)*A$$(1,0)" proof - let ?A = "(Mod_Type_Connect.to_hma\<^sub>m A)::'a^2^2" have [transfer_rule]: "Mod_Type_Connect.HMA_M A ?A" unfolding Mod_Type_Connect.HMA_M_def using from_hma_to_hma\<^sub>m A by auto have [transfer_rule]: "Mod_Type_Connect.HMA_I 0 0" unfolding Mod_Type_Connect.HMA_I_def by (simp add: to_nat_0) have [transfer_rule]: "Mod_Type_Connect.HMA_I 1 1" unfolding Mod_Type_Connect.HMA_I_def by (simp add: to_nat_1) have "Determinant.det A = Determinants.det ?A" by (transfer, simp) also have "... = ?A $h 1 $h 1 * ?A $h 2 $h 2 - ?A $h 1 $h 2 * ?A $h 2 $h 1" unfolding det_2 by simp also have "... = ?A $h 0 $h 0 * ?A $h 1 $h 1 - ?A $h 0 $h 1 * ?A $h 1 $h 0" by (smt Groups.mult_ac(2) exhaust_2 semiring_norm(160)) also have "... = A$$(0,0) * A $$ (1,1) - A$$(0,1)*A$$(1,0)" unfolding index_hma_def[symmetric] by (transfer, auto) finally show ?thesis . qed lemma mat_diag_smult: "mat_diag n (\ x. (k::'a::comm_ring_1)) = (k \\<^sub>m 1\<^sub>m n)" proof - have "mat_diag n (\ x. k) = mat_diag n (\ x. k * 1)" by auto also have "... = mat_diag n (\ x. k) * mat_diag n (\ x. 1)" using mat_diag_diag by (simp add: mat_diag_def) also have "... = mat_diag n (\ x. k) * (1\<^sub>m n)" by auto thm mat_diag_mult_left also have "... = Matrix.mat n n (\(i, j). k * (1\<^sub>m n) $$ (i, j))" by (rule mat_diag_mult_left, auto) also have "... = (k \\<^sub>m 1\<^sub>m n)" unfolding smult_mat_def by auto finally show ?thesis . qed lemma invertible_mat_four_block_mat_lower_right: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat n n" and inv_A: "invertible_mat A" shows "invertible_mat (four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 n) (0\<^sub>m n 1) A)" proof - let ?I = "(four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 n) (0\<^sub>m n 1) A)" have "Determinant.det ?I = Determinant.det (1\<^sub>m 1) * Determinant.det A" by (rule det_four_block_mat_lower_left_zero_col, insert assms, auto) also have "... = Determinant.det A" by auto finally have "Determinant.det ?I = Determinant.det A" . thus ?thesis by (metis (no_types, lifting) assms carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mat_four_block(2) index_mat_four_block(3) index_one_mat(2) index_one_mat(3) invertible_iff_is_unit_JNF) qed lemma invertible_mat_four_block_mat_lower_right_id: assumes A: "(A::'a::comm_ring_1 mat) \ carrier_mat m m" and B: "B = 0\<^sub>m m (n-m)" and C: "C = 0\<^sub>m (n-m) m" and D: "D = 1\<^sub>m (n-m)" and "n>m" and inv_A: "invertible_mat A" shows "invertible_mat (four_block_mat A B C D)" proof - have "Determinant.det (four_block_mat A B C D) = Determinant.det A" by (rule det_four_block_mat_lower_right_id, insert assms, auto) thus ?thesis using inv_A by (metis (no_types, lifting) assms(1) assms(4) carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mat_four_block(2) index_mat_four_block(3) index_one_mat(2) index_one_mat(3) invertible_iff_is_unit_JNF) qed lemma split_block4_decreases_dim_row: assumes E: "(A,B,C,D) = split_block E 1 1" and E1: "dim_row E > 1" and E2: "dim_col E > 1" shows "dim_row D < dim_row E" proof - have "D \ carrier_mat (1 + (dim_row E - 2)) (1 + (dim_col E - 2))" by (rule split_block(4)[OF E[symmetric]], insert E1 E2, auto) hence "D \ carrier_mat (dim_row E - 1) (dim_col E - 1)" using E1 E2 by auto thus ?thesis using E1 by auto qed lemma inv_P'PAQQ': assumes A: "A \ carrier_mat n n" and P: "P \ carrier_mat n n" and inv_P: "inverts_mat P' P" and inv_Q: "inverts_mat Q Q'" and Q: "Q \ carrier_mat n n" and P': "P' \ carrier_mat n n" and Q': "Q' \ carrier_mat n n" shows "(P'*(P*A*Q)*Q') = A" proof - have "(P'*(P*A*Q)*Q') = (P'*(P*A*Q*Q'))" by (smt P P' Q Q' assoc_mult_mat carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(2) index_mult_mat(3)) also have "... = ((P'*P)*A*(Q*Q'))" by (smt A P P' Q Q' assoc_mult_mat carrier_matD(1) carrier_matD(2) carrier_mat_triv index_mult_mat(3) inv_Q inverts_mat_def right_mult_one_mat') finally show ?thesis by (metis P' Q A inv_P inv_Q carrier_matD(1) inverts_mat_def left_mult_one_mat right_mult_one_mat) qed lemma assumes "U \ carrier_mat 2 2" and "V \ carrier_mat 2 2" and "A = U * V" shows mat_mult2_00: "A $$ (0,0) = U $$ (0,0)*V $$ (0,0) + U $$ (0,1)*V $$ (1,0)" and mat_mult2_01: "A $$ (0,1) = U $$ (0,0)*V $$ (0,1) + U $$ (0,1)*V $$ (1,1)" and mat_mult2_10: "A $$ (1,0) = U $$ (1,0)*V $$ (0,0) + U $$ (1,1)*V $$ (1,0)" and mat_mult2_11: "A $$ (1,1) = U $$ (1,0)*V $$ (0,1) + U $$ (1,1)*V $$ (1,1)" using assms unfolding times_mat_def Matrix.row_def col_def scalar_prod_def using sum_two_rw by auto subsection\Lemmas about @{text "sorted lists"}, @{text "insort"} and @{text "pick"}\ lemma sorted_distinct_imp_sorted_wrt: assumes "sorted xs" and "distinct xs" shows "sorted_wrt (<) xs" using assms by (induct xs, insert le_neq_trans, auto) lemma sorted_map_strict: assumes "strict_mono_on g {0.. g ` {0..x\set (map g [0.. g n" using sg unfolding strict_mono_on_def by (simp add: less_imp_le) qed finally show ?case . qed lemma sorted_nth_strict_mono: "sorted xs \ distinct xs \i < j \ j < length xs \ xs!i < xs!j" by (simp add: less_le nth_eq_iff_index_eq sorted_iff_nth_mono_less) lemma sorted_list_of_set_0_LEAST: assumes finI: "finite I" and I: "I \ {}" shows "sorted_list_of_set I ! 0 = (LEAST n. n\I)" proof (rule Least_equality[symmetric]) show "sorted_list_of_set I ! 0 \ I" by (metis I Max_in finI gr_zeroI in_set_conv_nth not_less_zero set_sorted_list_of_set) fix y assume "y \ I" thus "sorted_list_of_set I ! 0 \ y" by (metis eq_iff finI in_set_conv_nth neq0_conv sorted_iff_nth_mono_less sorted_list_of_set(1) sorted_sorted_list_of_set) qed lemma sorted_list_of_set_eq_pick: assumes i: "i < length (sorted_list_of_set I)" shows "sorted_list_of_set I ! i = pick I i" proof - have finI: "finite I" proof (rule ccontr) assume "infinite I" hence "length (sorted_list_of_set I) = 0" using sorted_list_of_set.infinite by auto thus False using i by simp qed show ?thesis using i proof (induct i) case 0 have I: "I \ {}" using "0.prems" sorted_list_of_set_empty by blast show ?case unfolding pick.simps by (rule sorted_list_of_set_0_LEAST[OF finI I]) next case (Suc i) note x_less = Suc.prems show ?case proof (unfold pick.simps, rule Least_equality[symmetric], rule conjI) show 1: "pick I i < sorted_list_of_set I ! Suc i" by (metis Suc.hyps Suc.prems Suc_lessD distinct_sorted_list_of_set find_first_unique lessI - nat_less_le sorted_sorted_list_of_set sorted_sorted_wrt sorted_wrt_nth_less) + nat_less_le sorted_sorted_list_of_set sorted_wrt_nth_less) show "sorted_list_of_set I ! Suc i \ I" using Suc.prems finI nth_mem set_sorted_list_of_set by blast have rw: "sorted_list_of_set I ! i = pick I i" using Suc.hyps Suc_lessD x_less by blast have sorted_less: "sorted_list_of_set I ! i < sorted_list_of_set I ! Suc i" by (simp add: 1 rw) fix y assume y: "y \ I \ pick I i < y" show "sorted_list_of_set I ! Suc i \ y" by (smt antisym_conv finI in_set_conv_nth less_Suc_eq less_Suc_eq_le nat_neq_iff rw sorted_iff_nth_mono_less sorted_list_of_set(1) sorted_sorted_list_of_set x_less y) qed qed qed text\$b$ is the position where we add, $a$ the element to be added and $i$ the position that is checked\ lemma insort_nth': assumes "\j set xs" and "i < length xs + 1" and "i < b" and "xs \ []" and "b < length xs" shows "insort a xs ! i = xs ! i" using assms proof (induct xs arbitrary: a b i) case Nil then show ?case by auto next case (Cons x xs) note less = Cons.prems(1) note sorted = Cons.prems(2) note a_notin = Cons.prems(3) note i_length = Cons.prems(4) note i_b = Cons.prems(5) note b_length = Cons.prems(7) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! i = (a # x # xs) ! i" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(5) True by force finally show ?thesis . next case False note x_less_a = False have "insort a (x # xs) ! i = (x # insort a xs) ! i" using False by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis by auto next case False have "(x # insort a xs) ! i = (insort a xs) ! (i-1)" by (simp add: False nth_Cons') also have "... = xs ! (i-1)" proof (rule Cons.hyps) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp show "i - 1 < length xs + 1" using i_length False by auto show "xs \ []" using i_b b_length by force show "i - 1 < b - 1" by (simp add: False diff_less_mono i_b leI) show "b - 1 < length xs" using b_length i_b by auto show "\j set xs" and "i < index (insort a xs) a" and "xs \ []" shows "insort a xs ! i = xs ! i" using assms proof (induct xs arbitrary: a i) case Nil then show ?case by auto next case (Cons x xs) note sorted = Cons.prems(1) note a_notin = Cons.prems(2) note i_index = Cons.prems(3) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! i = (a # x # xs) ! i" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(3) True by force finally show ?thesis . next case False note x_less_a = False show ?thesis proof (cases "xs = []") case True have "x \ a" using False by auto then show ?thesis using True i_index False by auto next case False note xs_not_empty = False have "insort a (x # xs) ! i = (x # insort a xs) ! i" using x_less_a by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis by auto next case False note i0 = False have "(x # insort a xs) ! i = (insort a xs) ! (i-1)" by (simp add: False nth_Cons') also have "... = xs ! (i-1)" proof (rule Cons.hyps[OF _ _ _ xs_not_empty]) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp have "index (insort a (x # xs)) a = index ((x # insort a xs)) a" using x_less_a by auto also have "... = index (insort a xs) a + 1" unfolding index_Cons using x_less_a by simp finally show "i - 1 < index (insort a xs) a" using False i_index by linarith qed also have "... = (x # xs) ! i" by (simp add: False nth_Cons') finally show ?thesis . qed finally show ?thesis . qed qed qed lemma insort_nth2: assumes "sorted xs" and "a \ set xs" and "i < length xs" and "i \ index (insort a xs) a" and "xs \ []" shows "insort a xs ! (Suc i) = xs ! i" using assms proof (induct xs arbitrary: a i) case Nil then show ?case by auto next case (Cons x xs) note sorted = Cons.prems(1) note a_notin = Cons.prems(2) note i_length = Cons.prems(3) note index_i = Cons.prems(4) show ?case proof (cases "a \ x") case True have "insort a (x # xs) ! (Suc i) = (a # x # xs) ! (Suc i)" using True by simp also have "... = (x # xs) ! i" using Cons.prems(1) Cons.prems(5) True by force finally show ?thesis . next case False note x_less_a = False have "insort a (x # xs) ! (Suc i) = (x # insort a xs) ! (Suc i)" using False by simp also have "... = (x # xs) ! i" proof (cases "i = 0") case True then show ?thesis using index_i linear x_less_a by fastforce next case False note i0 = False show ?thesis proof - have Suc_i: "Suc (i - 1) = i" using i0 by auto have "(x # insort a xs) ! (Suc i) = (insort a xs) ! i" by (simp add: nth_Cons') also have "... = (insort a xs) ! Suc (i - 1)" using Suc_i by simp also have "... = xs ! (i - 1)" proof (rule Cons.hyps) show "sorted xs" using sorted by simp show "a \ set xs" using a_notin by simp show "i - 1 < length xs" using i_length using Suc_i by auto thus "xs \ []" by auto have "index (insort a (x # xs)) a = index ((x # insort a xs)) a" using x_less_a by simp also have "... = index (insort a xs) a + 1" unfolding index_Cons using x_less_a by simp finally show "index (insort a xs) a \ i - 1" using index_i i0 by auto qed also have "... = (x # xs) ! i" using Suc_i by auto finally show ?thesis . qed qed finally show ?thesis . qed qed lemma pick_index: assumes a: "a \ I" and a'_card: "a' < card I" shows "(pick I a' = a) = (index (sorted_list_of_set I) a = a')" proof - have finI: "finite I" using a'_card card.infinite by force have length_I: "length (sorted_list_of_set I) = card I" by (metis a'_card card.infinite distinct_card distinct_sorted_list_of_set not_less_zero set_sorted_list_of_set) let ?i = "index (sorted_list_of_set I) a" have "(sorted_list_of_set I) ! a' = pick I a'" by (rule sorted_list_of_set_eq_pick, auto simp add: finI a'_card length_I) moreover have "(sorted_list_of_set I) ! ?i = a" by (rule nth_index, simp add: a finI) ultimately show ?thesis by (metis a'_card distinct_sorted_list_of_set index_nth_id length_I) qed end