# HG changeset patch # User haftmann # Date 1324934230 -3600 # Node ID b7b905b23b2ab4f2aa2d6ad4814d250850baf852 # Parent b39256df5f8a57c81cbdb28d7eabf2122b8a9577 incorporated More_Set and More_List into the Main body -- to be consolidated later diff -r b39256df5f8a -r b7b905b23b2a src/HOL/IMP/Abs_Int1_ivl.thy --- a/src/HOL/IMP/Abs_Int1_ivl.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/IMP/Abs_Int1_ivl.thy Mon Dec 26 22:17:10 2011 +0100 @@ -1,7 +1,7 @@ (* Author: Tobias Nipkow *) theory Abs_Int1_ivl -imports Abs_Int1 Abs_Int_Tests "~~/src/HOL/Library/More_Set" +imports Abs_Int1 Abs_Int_Tests begin subsection "Interval Analysis" diff -r b39256df5f8a -r b7b905b23b2a src/HOL/IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy --- a/src/HOL/IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy Mon Dec 26 22:17:10 2011 +0100 @@ -1,7 +1,7 @@ (* Author: Tobias Nipkow *) theory Abs_Int_den1_ivl -imports Abs_Int_den1 "~~/src/HOL/Library/More_Set" +imports Abs_Int_den1 begin subsection "Interval Analysis" diff -r b39256df5f8a -r b7b905b23b2a src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/IsaMakefile Mon Dec 26 22:17:10 2011 +0100 @@ -286,6 +286,8 @@ List.thy \ Main.thy \ Map.thy \ + More_List.thy \ + More_Set.thy \ Nat_Numeral.thy \ Nat_Transfer.thy \ New_DSequence.thy \ @@ -458,7 +460,7 @@ Library/LaTeXsugar.thy Library/Lattice_Algebras.thy \ Library/Lattice_Syntax.thy Library/Library.thy Library/List_Cset.thy \ Library/List_Prefix.thy Library/List_lexord.thy Library/Mapping.thy \ - Library/Monad_Syntax.thy Library/More_List.thy Library/More_Set.thy \ + Library/Monad_Syntax.thy \ Library/Multiset.thy Library/Nat_Bijection.thy \ Library/Numeral_Type.thy Library/Old_Recdef.thy \ Library/OptionalSugar.thy Library/Order_Relation.thy \ diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/AList_Impl.thy --- a/src/HOL/Library/AList_Impl.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Library/AList_Impl.thy Mon Dec 26 22:17:10 2011 +0100 @@ -5,7 +5,7 @@ header {* Implementation of Association Lists *} theory AList_Impl -imports Main More_List +imports Main begin text {* diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/Cset.thy --- a/src/HOL/Library/Cset.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Library/Cset.thy Mon Dec 26 22:17:10 2011 +0100 @@ -4,7 +4,7 @@ header {* A dedicated set type which is executable on its finite part *} theory Cset -imports More_Set More_List +imports Main begin subsection {* Lifting *} diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/Dlist.thy --- a/src/HOL/Library/Dlist.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Library/Dlist.thy Mon Dec 26 22:17:10 2011 +0100 @@ -3,7 +3,7 @@ header {* Lists with elements distinct as canonical example for datatype invariants *} theory Dlist -imports Main More_List +imports Main begin subsection {* The type of distinct lists *} diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Library/Library.thy Mon Dec 26 22:17:10 2011 +0100 @@ -33,7 +33,6 @@ Kleene_Algebra Mapping Monad_Syntax - More_List Multiset Numeral_Type Old_Recdef diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/Monad_Syntax.thy --- a/src/HOL/Library/Monad_Syntax.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Library/Monad_Syntax.thy Mon Dec 26 22:17:10 2011 +0100 @@ -5,7 +5,7 @@ header {* Monad notation for arbitrary types *} theory Monad_Syntax -imports Main "~~/src/Tools/Adhoc_Overloading" "~~/src/HOL/Library/More_List" +imports Main "~~/src/Tools/Adhoc_Overloading" begin text {* diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/More_List.thy --- a/src/HOL/Library/More_List.thy Mon Dec 26 22:17:10 2011 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,312 +0,0 @@ -(* Author: Florian Haftmann, TU Muenchen *) - -header {* Operations on lists beyond the standard List theory *} - -theory More_List -imports Main Multiset -begin - -hide_const (open) Finite_Set.fold - -text {* Repairing code generator setup *} - -declare (in lattice) Inf_fin_set_fold [code_unfold del] -declare (in lattice) Sup_fin_set_fold [code_unfold del] -declare (in linorder) Min_fin_set_fold [code_unfold del] -declare (in linorder) Max_fin_set_fold [code_unfold del] -declare (in complete_lattice) Inf_set_fold [code_unfold del] -declare (in complete_lattice) Sup_set_fold [code_unfold del] - - -text {* Fold combinator with canonical argument order *} - -primrec fold :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where - "fold f [] = id" - | "fold f (x # xs) = fold f xs \ f x" - -lemma foldl_fold: - "foldl f s xs = fold (\x s. f s x) xs s" - by (induct xs arbitrary: s) simp_all - -lemma foldr_fold_rev: - "foldr f xs = fold f (rev xs)" - by (simp add: foldr_foldl foldl_fold fun_eq_iff) - -lemma fold_rev_conv [code_unfold]: - "fold f (rev xs) = foldr f xs" - by (simp add: foldr_fold_rev) - -lemma fold_cong [fundef_cong]: - "a = b \ xs = ys \ (\x. x \ set xs \ f x = g x) - \ fold f xs a = fold g ys b" - by (induct ys arbitrary: a b xs) simp_all - -lemma fold_id: - assumes "\x. x \ set xs \ f x = id" - shows "fold f xs = id" - using assms by (induct xs) simp_all - -lemma fold_commute: - assumes "\x. x \ set xs \ h \ g x = f x \ h" - shows "h \ fold g xs = fold f xs \ h" - using assms by (induct xs) (simp_all add: fun_eq_iff) - -lemma fold_commute_apply: - assumes "\x. x \ set xs \ h \ g x = f x \ h" - shows "h (fold g xs s) = fold f xs (h s)" -proof - - from assms have "h \ fold g xs = fold f xs \ h" by (rule fold_commute) - then show ?thesis by (simp add: fun_eq_iff) -qed - -lemma fold_invariant: - assumes "\x. x \ set xs \ Q x" and "P s" - and "\x s. Q x \ P s \ P (f x s)" - shows "P (fold f xs s)" - using assms by (induct xs arbitrary: s) simp_all - -lemma fold_weak_invariant: - assumes "P s" - and "\s x. x \ set xs \ P s \ P (f x s)" - shows "P (fold f xs s)" - using assms by (induct xs arbitrary: s) simp_all - -lemma fold_append [simp]: - "fold f (xs @ ys) = fold f ys \ fold f xs" - by (induct xs) simp_all - -lemma fold_map [code_unfold]: - "fold g (map f xs) = fold (g o f) xs" - by (induct xs) simp_all - -lemma fold_remove1_split: - assumes f: "\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x" - and x: "x \ set xs" - shows "fold f xs = fold f (remove1 x xs) \ f x" - using assms by (induct xs) (auto simp add: o_assoc [symmetric]) - -lemma fold_multiset_equiv: - assumes f: "\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x" - and equiv: "multiset_of xs = multiset_of ys" - shows "fold f xs = fold f ys" -using f equiv [symmetric] proof (induct xs arbitrary: ys) - case Nil then show ?case by simp -next - case (Cons x xs) - then have *: "set ys = set (x # xs)" by (blast dest: multiset_of_eq_setD) - have "\x y. x \ set ys \ y \ set ys \ f x \ f y = f y \ f x" - by (rule Cons.prems(1)) (simp_all add: *) - moreover from * have "x \ set ys" by simp - ultimately have "fold f ys = fold f (remove1 x ys) \ f x" by (fact fold_remove1_split) - moreover from Cons.prems have "fold f xs = fold f (remove1 x ys)" by (auto intro: Cons.hyps) - ultimately show ?case by simp -qed - -lemma fold_rev: - assumes "\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y" - shows "fold f (rev xs) = fold f xs" - by (rule fold_multiset_equiv, rule assms) (simp_all add: in_multiset_in_set) - -lemma foldr_fold: - assumes "\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y" - shows "foldr f xs = fold f xs" - using assms unfolding foldr_fold_rev by (rule fold_rev) - -lemma fold_Cons_rev: - "fold Cons xs = append (rev xs)" - by (induct xs) simp_all - -lemma rev_conv_fold [code]: - "rev xs = fold Cons xs []" - by (simp add: fold_Cons_rev) - -lemma fold_append_concat_rev: - "fold append xss = append (concat (rev xss))" - by (induct xss) simp_all - -lemma concat_conv_foldr [code]: - "concat xss = foldr append xss []" - by (simp add: fold_append_concat_rev foldr_fold_rev) - -lemma fold_plus_listsum_rev: - "fold plus xs = plus (listsum (rev xs))" - by (induct xs) (simp_all add: add.assoc) - -lemma (in monoid_add) listsum_conv_fold [code]: - "listsum xs = fold (\x y. y + x) xs 0" - by (auto simp add: listsum_foldl foldl_fold fun_eq_iff) - -lemma (in linorder) sort_key_conv_fold: - assumes "inj_on f (set xs)" - shows "sort_key f xs = fold (insort_key f) xs []" -proof - - have "fold (insort_key f) (rev xs) = fold (insort_key f) xs" - proof (rule fold_rev, rule ext) - fix zs - fix x y - assume "x \ set xs" "y \ set xs" - with assms have *: "f y = f x \ y = x" by (auto dest: inj_onD) - have **: "x = y \ y = x" by auto - show "(insort_key f y \ insort_key f x) zs = (insort_key f x \ insort_key f y) zs" - by (induct zs) (auto intro: * simp add: **) - qed - then show ?thesis by (simp add: sort_key_def foldr_fold_rev) -qed - -lemma (in linorder) sort_conv_fold: - "sort xs = fold insort xs []" - by (rule sort_key_conv_fold) simp - - -text {* @{const Finite_Set.fold} and @{const fold} *} - -lemma (in comp_fun_commute) fold_set_remdups: - "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" - by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm insert_absorb) - -lemma (in comp_fun_idem) fold_set: - "Finite_Set.fold f y (set xs) = fold f xs y" - by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm) - -lemma (in ab_semigroup_idem_mult) fold1_set: - assumes "xs \ []" - shows "Finite_Set.fold1 times (set xs) = fold times (tl xs) (hd xs)" -proof - - interpret comp_fun_idem times by (fact comp_fun_idem) - from assms obtain y ys where xs: "xs = y # ys" - by (cases xs) auto - show ?thesis - proof (cases "set ys = {}") - case True with xs show ?thesis by simp - next - case False - then have "fold1 times (insert y (set ys)) = Finite_Set.fold times y (set ys)" - by (simp only: finite_set fold1_eq_fold_idem) - with xs show ?thesis by (simp add: fold_set mult_commute) - qed -qed - -lemma (in lattice) Inf_fin_set_fold: - "Inf_fin (set (x # xs)) = fold inf xs x" -proof - - interpret ab_semigroup_idem_mult "inf :: 'a \ 'a \ 'a" - by (fact ab_semigroup_idem_mult_inf) - show ?thesis - by (simp add: Inf_fin_def fold1_set del: set.simps) -qed - -lemma (in lattice) Inf_fin_set_foldr [code_unfold]: - "Inf_fin (set (x # xs)) = foldr inf xs x" - by (simp add: Inf_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) - -lemma (in lattice) Sup_fin_set_fold: - "Sup_fin (set (x # xs)) = fold sup xs x" -proof - - interpret ab_semigroup_idem_mult "sup :: 'a \ 'a \ 'a" - by (fact ab_semigroup_idem_mult_sup) - show ?thesis - by (simp add: Sup_fin_def fold1_set del: set.simps) -qed - -lemma (in lattice) Sup_fin_set_foldr [code_unfold]: - "Sup_fin (set (x # xs)) = foldr sup xs x" - by (simp add: Sup_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) - -lemma (in linorder) Min_fin_set_fold: - "Min (set (x # xs)) = fold min xs x" -proof - - interpret ab_semigroup_idem_mult "min :: 'a \ 'a \ 'a" - by (fact ab_semigroup_idem_mult_min) - show ?thesis - by (simp add: Min_def fold1_set del: set.simps) -qed - -lemma (in linorder) Min_fin_set_foldr [code_unfold]: - "Min (set (x # xs)) = foldr min xs x" - by (simp add: Min_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) - -lemma (in linorder) Max_fin_set_fold: - "Max (set (x # xs)) = fold max xs x" -proof - - interpret ab_semigroup_idem_mult "max :: 'a \ 'a \ 'a" - by (fact ab_semigroup_idem_mult_max) - show ?thesis - by (simp add: Max_def fold1_set del: set.simps) -qed - -lemma (in linorder) Max_fin_set_foldr [code_unfold]: - "Max (set (x # xs)) = foldr max xs x" - by (simp add: Max_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) - -lemma (in complete_lattice) Inf_set_fold: - "Inf (set xs) = fold inf xs top" -proof - - interpret comp_fun_idem "inf :: 'a \ 'a \ 'a" - by (fact comp_fun_idem_inf) - show ?thesis by (simp add: Inf_fold_inf fold_set inf_commute) -qed - -lemma (in complete_lattice) Inf_set_foldr [code_unfold]: - "Inf (set xs) = foldr inf xs top" - by (simp add: Inf_set_fold ac_simps foldr_fold fun_eq_iff) - -lemma (in complete_lattice) Sup_set_fold: - "Sup (set xs) = fold sup xs bot" -proof - - interpret comp_fun_idem "sup :: 'a \ 'a \ 'a" - by (fact comp_fun_idem_sup) - show ?thesis by (simp add: Sup_fold_sup fold_set sup_commute) -qed - -lemma (in complete_lattice) Sup_set_foldr [code_unfold]: - "Sup (set xs) = foldr sup xs bot" - by (simp add: Sup_set_fold ac_simps foldr_fold fun_eq_iff) - -lemma (in complete_lattice) INFI_set_fold: - "INFI (set xs) f = fold (inf \ f) xs top" - unfolding INF_def set_map [symmetric] Inf_set_fold fold_map .. - -lemma (in complete_lattice) SUPR_set_fold: - "SUPR (set xs) f = fold (sup \ f) xs bot" - unfolding SUP_def set_map [symmetric] Sup_set_fold fold_map .. - - -text {* @{text nth_map} *} - -definition nth_map :: "nat \ ('a \ 'a) \ 'a list \ 'a list" where - "nth_map n f xs = (if n < length xs then - take n xs @ [f (xs ! n)] @ drop (Suc n) xs - else xs)" - -lemma nth_map_id: - "n \ length xs \ nth_map n f xs = xs" - by (simp add: nth_map_def) - -lemma nth_map_unfold: - "n < length xs \ nth_map n f xs = take n xs @ [f (xs ! n)] @ drop (Suc n) xs" - by (simp add: nth_map_def) - -lemma nth_map_Nil [simp]: - "nth_map n f [] = []" - by (simp add: nth_map_def) - -lemma nth_map_zero [simp]: - "nth_map 0 f (x # xs) = f x # xs" - by (simp add: nth_map_def) - -lemma nth_map_Suc [simp]: - "nth_map (Suc n) f (x # xs) = x # nth_map n f xs" - by (simp add: nth_map_def) - - -text {* monad operation *} - -definition bind :: "'a list \ ('a \ 'b list) \ 'b list" where - "bind xs f = concat (map f xs)" - -lemma bind_simps [simp]: - "bind [] f = []" - "bind (x # xs) f = f x @ bind xs f" - by (simp_all add: bind_def) - -end diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/More_Set.thy --- a/src/HOL/Library/More_Set.thy Mon Dec 26 22:17:10 2011 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,284 +0,0 @@ - -(* Author: Florian Haftmann, TU Muenchen *) - -header {* Relating (finite) sets and lists *} - -theory More_Set -imports Main More_List -begin - -lemma comp_fun_idem_remove: - "comp_fun_idem Set.remove" -proof - - have rem: "Set.remove = (\x A. A - {x})" by (simp add: fun_eq_iff remove_def) - show ?thesis by (simp only: comp_fun_idem_remove rem) -qed - -lemma minus_fold_remove: - assumes "finite A" - shows "B - A = Finite_Set.fold Set.remove B A" -proof - - have rem: "Set.remove = (\x A. A - {x})" by (simp add: fun_eq_iff remove_def) - show ?thesis by (simp only: rem assms minus_fold_remove) -qed - -lemma bounded_Collect_code [code_unfold_post]: - "{x \ A. P x} = Set.project P A" - by (simp add: project_def) - - -subsection {* Basic set operations *} - -lemma is_empty_set: - "Set.is_empty (set xs) \ List.null xs" - by (simp add: Set.is_empty_def null_def) - -lemma empty_set: - "{} = set []" - by simp - -lemma insert_set_compl: - "insert x (- set xs) = - set (removeAll x xs)" - by auto - -lemma remove_set_compl: - "Set.remove x (- set xs) = - set (List.insert x xs)" - by (auto simp add: remove_def List.insert_def) - -lemma image_set: - "image f (set xs) = set (map f xs)" - by simp - -lemma project_set: - "Set.project P (set xs) = set (filter P xs)" - by (auto simp add: project_def) - - -subsection {* Functorial set operations *} - -lemma union_set: - "set xs \ A = fold Set.insert xs A" -proof - - interpret comp_fun_idem Set.insert - by (fact comp_fun_idem_insert) - show ?thesis by (simp add: union_fold_insert fold_set) -qed - -lemma union_set_foldr: - "set xs \ A = foldr Set.insert xs A" -proof - - have "\x y :: 'a. insert y \ insert x = insert x \ insert y" - by auto - then show ?thesis by (simp add: union_set foldr_fold) -qed - -lemma minus_set: - "A - set xs = fold Set.remove xs A" -proof - - interpret comp_fun_idem Set.remove - by (fact comp_fun_idem_remove) - show ?thesis - by (simp add: minus_fold_remove [of _ A] fold_set) -qed - -lemma minus_set_foldr: - "A - set xs = foldr Set.remove xs A" -proof - - have "\x y :: 'a. Set.remove y \ Set.remove x = Set.remove x \ Set.remove y" - by (auto simp add: remove_def) - then show ?thesis by (simp add: minus_set foldr_fold) -qed - - -subsection {* Derived set operations *} - -lemma member: - "a \ A \ (\x\A. a = x)" - by simp - -lemma subset_eq: - "A \ B \ (\x\A. x \ B)" - by (fact subset_eq) - -lemma subset: - "A \ B \ A \ B \ \ B \ A" - by (fact less_le_not_le) - -lemma set_eq: - "A = B \ A \ B \ B \ A" - by (fact eq_iff) - -lemma inter: - "A \ B = Set.project (\x. x \ A) B" - by (auto simp add: project_def) - - -subsection {* Theorems on relations *} - -lemma product_code: - "Product_Type.product (set xs) (set ys) = set [(x, y). x \ xs, y \ ys]" - by (auto simp add: Product_Type.product_def) - -lemma Id_on_set: - "Id_on (set xs) = set [(x, x). x \ xs]" - by (auto simp add: Id_on_def) - -lemma trancl_set_ntrancl: "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)" - by (simp add: finite_trancl_ntranl) - -lemma set_rel_comp: - "set xys O set yzs = set ([(fst xy, snd yz). xy \ xys, yz \ yzs, snd xy = fst yz])" - by (auto simp add: Bex_def) - -lemma wf_set: - "wf (set xs) = acyclic (set xs)" - by (simp add: wf_iff_acyclic_if_finite) - - -subsection {* Code generator setup *} - -definition coset :: "'a list \ 'a set" where - [simp]: "coset xs = - set xs" - -code_datatype set coset - - -subsection {* Basic operations *} - -lemma [code]: - "x \ set xs \ List.member xs x" - "x \ coset xs \ \ List.member xs x" - by (simp_all add: member_def) - -lemma [code_unfold]: - "A = {} \ Set.is_empty A" - by (simp add: Set.is_empty_def) - -declare empty_set [code] - -declare is_empty_set [code] - -lemma UNIV_coset [code]: - "UNIV = coset []" - by simp - -lemma insert_code [code]: - "insert x (set xs) = set (List.insert x xs)" - "insert x (coset xs) = coset (removeAll x xs)" - by simp_all - -lemma remove_code [code]: - "Set.remove x (set xs) = set (removeAll x xs)" - "Set.remove x (coset xs) = coset (List.insert x xs)" - by (simp_all add: remove_def Compl_insert) - -declare image_set [code] - -declare project_set [code] - -lemma Ball_set [code]: - "Ball (set xs) P \ list_all P xs" - by (simp add: list_all_iff) - -lemma Bex_set [code]: - "Bex (set xs) P \ list_ex P xs" - by (simp add: list_ex_iff) - -lemma card_set [code]: - "card (set xs) = length (remdups xs)" -proof - - have "card (set (remdups xs)) = length (remdups xs)" - by (rule distinct_card) simp - then show ?thesis by simp -qed - - -subsection {* Derived operations *} - -declare subset_eq [code] - -declare subset [code] - - -subsection {* Functorial operations *} - -lemma inter_code [code]: - "A \ set xs = set (List.filter (\x. x \ A) xs)" - "A \ coset xs = foldr Set.remove xs A" - by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr) - -lemma subtract_code [code]: - "A - set xs = foldr Set.remove xs A" - "A - coset xs = set (List.filter (\x. x \ A) xs)" - by (auto simp add: minus_set_foldr) - -lemma union_code [code]: - "set xs \ A = foldr insert xs A" - "coset xs \ A = coset (List.filter (\x. x \ A) xs)" - by (auto simp add: union_set_foldr) - -definition Inf :: "'a::complete_lattice set \ 'a" where - [simp]: "Inf = Complete_Lattices.Inf" - -hide_const (open) Inf - -lemma [code_unfold_post]: - "Inf = More_Set.Inf" - by simp - -definition Sup :: "'a::complete_lattice set \ 'a" where - [simp]: "Sup = Complete_Lattices.Sup" - -hide_const (open) Sup - -lemma [code_unfold_post]: - "Sup = More_Set.Sup" - by simp - -lemma Inf_code [code]: - "More_Set.Inf (set xs) = foldr inf xs top" - "More_Set.Inf (coset []) = bot" - by (simp_all add: Inf_set_foldr) - -lemma Sup_sup [code]: - "More_Set.Sup (set xs) = foldr sup xs bot" - "More_Set.Sup (coset []) = top" - by (simp_all add: Sup_set_foldr) - -lemma INF_code [code]: - "INFI (set xs) f = foldr (inf \ f) xs top" - by (simp add: INF_def Inf_set_foldr image_set foldr_map del: set_map) - -lemma SUP_sup [code]: - "SUPR (set xs) f = foldr (sup \ f) xs bot" - by (simp add: SUP_def Sup_set_foldr image_set foldr_map del: set_map) - -hide_const (open) coset - - -subsection {* Operations on relations *} - -text {* Initially contributed by Tjark Weber. *} - -declare Domain_fst [code] - -declare Range_snd [code] - -declare trans_join [code] - -declare irrefl_distinct [code] - -declare trancl_set_ntrancl [code] - -declare acyclic_irrefl [code] - -declare product_code [code] - -declare Id_on_set [code] - -declare set_rel_comp [code] - -declare wf_set [code] - -end - diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Library/RBT_Impl.thy --- a/src/HOL/Library/RBT_Impl.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Library/RBT_Impl.thy Mon Dec 26 22:17:10 2011 +0100 @@ -6,7 +6,7 @@ header {* Implementation of Red-Black Trees *} theory RBT_Impl -imports Main More_List +imports Main begin text {* diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Main.thy --- a/src/HOL/Main.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Main.thy Mon Dec 26 22:17:10 2011 +0100 @@ -1,7 +1,7 @@ header {* Main HOL *} theory Main -imports Plain Predicate_Compile Nitpick +imports Plain Predicate_Compile Nitpick More_Set begin text {* diff -r b39256df5f8a -r b7b905b23b2a src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/MicroJava/BV/BVExample.thy Mon Dec 26 22:17:10 2011 +0100 @@ -9,7 +9,6 @@ "../JVM/JVMListExample" BVSpecTypeSafe JVM - "~~/src/HOL/Library/More_Set" begin text {* diff -r b39256df5f8a -r b7b905b23b2a src/HOL/More_List.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/More_List.thy Mon Dec 26 22:17:10 2011 +0100 @@ -0,0 +1,295 @@ +(* Author: Florian Haftmann, TU Muenchen *) + +header {* Operations on lists beyond the standard List theory *} + +theory More_List +imports List +begin + +hide_const (open) Finite_Set.fold + +text {* Repairing code generator setup *} + +declare (in lattice) Inf_fin_set_fold [code_unfold del] +declare (in lattice) Sup_fin_set_fold [code_unfold del] +declare (in linorder) Min_fin_set_fold [code_unfold del] +declare (in linorder) Max_fin_set_fold [code_unfold del] +declare (in complete_lattice) Inf_set_fold [code_unfold del] +declare (in complete_lattice) Sup_set_fold [code_unfold del] + + +text {* Fold combinator with canonical argument order *} + +primrec fold :: "('a \ 'b \ 'b) \ 'a list \ 'b \ 'b" where + "fold f [] = id" + | "fold f (x # xs) = fold f xs \ f x" + +lemma foldl_fold: + "foldl f s xs = fold (\x s. f s x) xs s" + by (induct xs arbitrary: s) simp_all + +lemma foldr_fold_rev: + "foldr f xs = fold f (rev xs)" + by (simp add: foldr_foldl foldl_fold fun_eq_iff) + +lemma fold_rev_conv [code_unfold]: + "fold f (rev xs) = foldr f xs" + by (simp add: foldr_fold_rev) + +lemma fold_cong [fundef_cong]: + "a = b \ xs = ys \ (\x. x \ set xs \ f x = g x) + \ fold f xs a = fold g ys b" + by (induct ys arbitrary: a b xs) simp_all + +lemma fold_id: + assumes "\x. x \ set xs \ f x = id" + shows "fold f xs = id" + using assms by (induct xs) simp_all + +lemma fold_commute: + assumes "\x. x \ set xs \ h \ g x = f x \ h" + shows "h \ fold g xs = fold f xs \ h" + using assms by (induct xs) (simp_all add: fun_eq_iff) + +lemma fold_commute_apply: + assumes "\x. x \ set xs \ h \ g x = f x \ h" + shows "h (fold g xs s) = fold f xs (h s)" +proof - + from assms have "h \ fold g xs = fold f xs \ h" by (rule fold_commute) + then show ?thesis by (simp add: fun_eq_iff) +qed + +lemma fold_invariant: + assumes "\x. x \ set xs \ Q x" and "P s" + and "\x s. Q x \ P s \ P (f x s)" + shows "P (fold f xs s)" + using assms by (induct xs arbitrary: s) simp_all + +lemma fold_weak_invariant: + assumes "P s" + and "\s x. x \ set xs \ P s \ P (f x s)" + shows "P (fold f xs s)" + using assms by (induct xs arbitrary: s) simp_all + +lemma fold_append [simp]: + "fold f (xs @ ys) = fold f ys \ fold f xs" + by (induct xs) simp_all + +lemma fold_map [code_unfold]: + "fold g (map f xs) = fold (g o f) xs" + by (induct xs) simp_all + +lemma fold_remove1_split: + assumes f: "\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x" + and x: "x \ set xs" + shows "fold f xs = fold f (remove1 x xs) \ f x" + using assms by (induct xs) (auto simp add: o_assoc [symmetric]) + +lemma fold_rev: + assumes "\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y" + shows "fold f (rev xs) = fold f xs" +using assms by (induct xs) (simp_all add: fold_commute_apply fun_eq_iff) + +lemma foldr_fold: + assumes "\x y. x \ set xs \ y \ set xs \ f y \ f x = f x \ f y" + shows "foldr f xs = fold f xs" + using assms unfolding foldr_fold_rev by (rule fold_rev) + +lemma fold_Cons_rev: + "fold Cons xs = append (rev xs)" + by (induct xs) simp_all + +lemma rev_conv_fold [code]: + "rev xs = fold Cons xs []" + by (simp add: fold_Cons_rev) + +lemma fold_append_concat_rev: + "fold append xss = append (concat (rev xss))" + by (induct xss) simp_all + +lemma concat_conv_foldr [code]: + "concat xss = foldr append xss []" + by (simp add: fold_append_concat_rev foldr_fold_rev) + +lemma fold_plus_listsum_rev: + "fold plus xs = plus (listsum (rev xs))" + by (induct xs) (simp_all add: add.assoc) + +lemma (in monoid_add) listsum_conv_fold [code]: + "listsum xs = fold (\x y. y + x) xs 0" + by (auto simp add: listsum_foldl foldl_fold fun_eq_iff) + +lemma (in linorder) sort_key_conv_fold: + assumes "inj_on f (set xs)" + shows "sort_key f xs = fold (insort_key f) xs []" +proof - + have "fold (insort_key f) (rev xs) = fold (insort_key f) xs" + proof (rule fold_rev, rule ext) + fix zs + fix x y + assume "x \ set xs" "y \ set xs" + with assms have *: "f y = f x \ y = x" by (auto dest: inj_onD) + have **: "x = y \ y = x" by auto + show "(insort_key f y \ insort_key f x) zs = (insort_key f x \ insort_key f y) zs" + by (induct zs) (auto intro: * simp add: **) + qed + then show ?thesis by (simp add: sort_key_def foldr_fold_rev) +qed + +lemma (in linorder) sort_conv_fold: + "sort xs = fold insort xs []" + by (rule sort_key_conv_fold) simp + + +text {* @{const Finite_Set.fold} and @{const fold} *} + +lemma (in comp_fun_commute) fold_set_fold_remdups: + "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" + by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm insert_absorb) + +lemma (in comp_fun_idem) fold_set_fold: + "Finite_Set.fold f y (set xs) = fold f xs y" + by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm) + +lemma (in ab_semigroup_idem_mult) fold1_set_fold: + assumes "xs \ []" + shows "Finite_Set.fold1 times (set xs) = fold times (tl xs) (hd xs)" +proof - + interpret comp_fun_idem times by (fact comp_fun_idem) + from assms obtain y ys where xs: "xs = y # ys" + by (cases xs) auto + show ?thesis + proof (cases "set ys = {}") + case True with xs show ?thesis by simp + next + case False + then have "fold1 times (insert y (set ys)) = Finite_Set.fold times y (set ys)" + by (simp only: finite_set fold1_eq_fold_idem) + with xs show ?thesis by (simp add: fold_set_fold mult_commute) + qed +qed + +lemma (in lattice) Inf_fin_set_fold: + "Inf_fin (set (x # xs)) = fold inf xs x" +proof - + interpret ab_semigroup_idem_mult "inf :: 'a \ 'a \ 'a" + by (fact ab_semigroup_idem_mult_inf) + show ?thesis + by (simp add: Inf_fin_def fold1_set_fold del: set.simps) +qed + +lemma (in lattice) Inf_fin_set_foldr [code_unfold]: + "Inf_fin (set (x # xs)) = foldr inf xs x" + by (simp add: Inf_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) + +lemma (in lattice) Sup_fin_set_fold: + "Sup_fin (set (x # xs)) = fold sup xs x" +proof - + interpret ab_semigroup_idem_mult "sup :: 'a \ 'a \ 'a" + by (fact ab_semigroup_idem_mult_sup) + show ?thesis + by (simp add: Sup_fin_def fold1_set_fold del: set.simps) +qed + +lemma (in lattice) Sup_fin_set_foldr [code_unfold]: + "Sup_fin (set (x # xs)) = foldr sup xs x" + by (simp add: Sup_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) + +lemma (in linorder) Min_fin_set_fold: + "Min (set (x # xs)) = fold min xs x" +proof - + interpret ab_semigroup_idem_mult "min :: 'a \ 'a \ 'a" + by (fact ab_semigroup_idem_mult_min) + show ?thesis + by (simp add: Min_def fold1_set_fold del: set.simps) +qed + +lemma (in linorder) Min_fin_set_foldr [code_unfold]: + "Min (set (x # xs)) = foldr min xs x" + by (simp add: Min_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) + +lemma (in linorder) Max_fin_set_fold: + "Max (set (x # xs)) = fold max xs x" +proof - + interpret ab_semigroup_idem_mult "max :: 'a \ 'a \ 'a" + by (fact ab_semigroup_idem_mult_max) + show ?thesis + by (simp add: Max_def fold1_set_fold del: set.simps) +qed + +lemma (in linorder) Max_fin_set_foldr [code_unfold]: + "Max (set (x # xs)) = foldr max xs x" + by (simp add: Max_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) + +lemma (in complete_lattice) Inf_set_fold: + "Inf (set xs) = fold inf xs top" +proof - + interpret comp_fun_idem "inf :: 'a \ 'a \ 'a" + by (fact comp_fun_idem_inf) + show ?thesis by (simp add: Inf_fold_inf fold_set_fold inf_commute) +qed + +lemma (in complete_lattice) Inf_set_foldr [code_unfold]: + "Inf (set xs) = foldr inf xs top" + by (simp add: Inf_set_fold ac_simps foldr_fold fun_eq_iff) + +lemma (in complete_lattice) Sup_set_fold: + "Sup (set xs) = fold sup xs bot" +proof - + interpret comp_fun_idem "sup :: 'a \ 'a \ 'a" + by (fact comp_fun_idem_sup) + show ?thesis by (simp add: Sup_fold_sup fold_set sup_commute) +qed + +lemma (in complete_lattice) Sup_set_foldr [code_unfold]: + "Sup (set xs) = foldr sup xs bot" + by (simp add: Sup_set_fold ac_simps foldr_fold fun_eq_iff) + +lemma (in complete_lattice) INFI_set_fold: + "INFI (set xs) f = fold (inf \ f) xs top" + unfolding INF_def set_map [symmetric] Inf_set_fold fold_map .. + +lemma (in complete_lattice) SUPR_set_fold: + "SUPR (set xs) f = fold (sup \ f) xs bot" + unfolding SUP_def set_map [symmetric] Sup_set_fold fold_map .. + + +text {* @{text nth_map} *} + +definition nth_map :: "nat \ ('a \ 'a) \ 'a list \ 'a list" where + "nth_map n f xs = (if n < length xs then + take n xs @ [f (xs ! n)] @ drop (Suc n) xs + else xs)" + +lemma nth_map_id: + "n \ length xs \ nth_map n f xs = xs" + by (simp add: nth_map_def) + +lemma nth_map_unfold: + "n < length xs \ nth_map n f xs = take n xs @ [f (xs ! n)] @ drop (Suc n) xs" + by (simp add: nth_map_def) + +lemma nth_map_Nil [simp]: + "nth_map n f [] = []" + by (simp add: nth_map_def) + +lemma nth_map_zero [simp]: + "nth_map 0 f (x # xs) = f x # xs" + by (simp add: nth_map_def) + +lemma nth_map_Suc [simp]: + "nth_map (Suc n) f (x # xs) = x # nth_map n f xs" + by (simp add: nth_map_def) + + +text {* monad operation *} + +definition bind :: "'a list \ ('a \ 'b list) \ 'b list" where + "bind xs f = concat (map f xs)" + +lemma bind_simps [simp]: + "bind [] f = []" + "bind (x # xs) f = f x @ bind xs f" + by (simp_all add: bind_def) + +end diff -r b39256df5f8a -r b7b905b23b2a src/HOL/More_Set.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/More_Set.thy Mon Dec 26 22:17:10 2011 +0100 @@ -0,0 +1,284 @@ + +(* Author: Florian Haftmann, TU Muenchen *) + +header {* Relating (finite) sets and lists *} + +theory More_Set +imports More_List +begin + +lemma comp_fun_idem_remove: + "comp_fun_idem Set.remove" +proof - + have rem: "Set.remove = (\x A. A - {x})" by (simp add: fun_eq_iff remove_def) + show ?thesis by (simp only: comp_fun_idem_remove rem) +qed + +lemma minus_fold_remove: + assumes "finite A" + shows "B - A = Finite_Set.fold Set.remove B A" +proof - + have rem: "Set.remove = (\x A. A - {x})" by (simp add: fun_eq_iff remove_def) + show ?thesis by (simp only: rem assms minus_fold_remove) +qed + +lemma bounded_Collect_code [code_unfold_post]: + "{x \ A. P x} = Set.project P A" + by (simp add: project_def) + + +subsection {* Basic set operations *} + +lemma is_empty_set: + "Set.is_empty (set xs) \ List.null xs" + by (simp add: Set.is_empty_def null_def) + +lemma empty_set: + "{} = set []" + by simp + +lemma insert_set_compl: + "insert x (- set xs) = - set (removeAll x xs)" + by auto + +lemma remove_set_compl: + "Set.remove x (- set xs) = - set (List.insert x xs)" + by (auto simp add: remove_def List.insert_def) + +lemma image_set: + "image f (set xs) = set (map f xs)" + by simp + +lemma project_set: + "Set.project P (set xs) = set (filter P xs)" + by (auto simp add: project_def) + + +subsection {* Functorial set operations *} + +lemma union_set: + "set xs \ A = fold Set.insert xs A" +proof - + interpret comp_fun_idem Set.insert + by (fact comp_fun_idem_insert) + show ?thesis by (simp add: union_fold_insert fold_set_fold) +qed + +lemma union_set_foldr: + "set xs \ A = foldr Set.insert xs A" +proof - + have "\x y :: 'a. insert y \ insert x = insert x \ insert y" + by auto + then show ?thesis by (simp add: union_set foldr_fold) +qed + +lemma minus_set: + "A - set xs = fold Set.remove xs A" +proof - + interpret comp_fun_idem Set.remove + by (fact comp_fun_idem_remove) + show ?thesis + by (simp add: minus_fold_remove [of _ A] fold_set_fold) +qed + +lemma minus_set_foldr: + "A - set xs = foldr Set.remove xs A" +proof - + have "\x y :: 'a. Set.remove y \ Set.remove x = Set.remove x \ Set.remove y" + by (auto simp add: remove_def) + then show ?thesis by (simp add: minus_set foldr_fold) +qed + + +subsection {* Derived set operations *} + +lemma member: + "a \ A \ (\x\A. a = x)" + by simp + +lemma subset_eq: + "A \ B \ (\x\A. x \ B)" + by (fact subset_eq) + +lemma subset: + "A \ B \ A \ B \ \ B \ A" + by (fact less_le_not_le) + +lemma set_eq: + "A = B \ A \ B \ B \ A" + by (fact eq_iff) + +lemma inter: + "A \ B = Set.project (\x. x \ A) B" + by (auto simp add: project_def) + + +subsection {* Theorems on relations *} + +lemma product_code: + "Product_Type.product (set xs) (set ys) = set [(x, y). x \ xs, y \ ys]" + by (auto simp add: Product_Type.product_def) + +lemma Id_on_set: + "Id_on (set xs) = set [(x, x). x \ xs]" + by (auto simp add: Id_on_def) + +lemma trancl_set_ntrancl: "trancl (set xs) = ntrancl (card (set xs) - 1) (set xs)" + by (simp add: finite_trancl_ntranl) + +lemma set_rel_comp: + "set xys O set yzs = set ([(fst xy, snd yz). xy \ xys, yz \ yzs, snd xy = fst yz])" + by (auto simp add: Bex_def) + +lemma wf_set: + "wf (set xs) = acyclic (set xs)" + by (simp add: wf_iff_acyclic_if_finite) + + +subsection {* Code generator setup *} + +definition coset :: "'a list \ 'a set" where + [simp]: "coset xs = - set xs" + +code_datatype set coset + + +subsection {* Basic operations *} + +lemma [code]: + "x \ set xs \ List.member xs x" + "x \ coset xs \ \ List.member xs x" + by (simp_all add: member_def) + +lemma [code_unfold]: + "A = {} \ Set.is_empty A" + by (simp add: Set.is_empty_def) + +declare empty_set [code] + +declare is_empty_set [code] + +lemma UNIV_coset [code]: + "UNIV = coset []" + by simp + +lemma insert_code [code]: + "insert x (set xs) = set (List.insert x xs)" + "insert x (coset xs) = coset (removeAll x xs)" + by simp_all + +lemma remove_code [code]: + "Set.remove x (set xs) = set (removeAll x xs)" + "Set.remove x (coset xs) = coset (List.insert x xs)" + by (simp_all add: remove_def Compl_insert) + +declare image_set [code] + +declare project_set [code] + +lemma Ball_set [code]: + "Ball (set xs) P \ list_all P xs" + by (simp add: list_all_iff) + +lemma Bex_set [code]: + "Bex (set xs) P \ list_ex P xs" + by (simp add: list_ex_iff) + +lemma card_set [code]: + "card (set xs) = length (remdups xs)" +proof - + have "card (set (remdups xs)) = length (remdups xs)" + by (rule distinct_card) simp + then show ?thesis by simp +qed + + +subsection {* Derived operations *} + +declare subset_eq [code] + +declare subset [code] + + +subsection {* Functorial operations *} + +lemma inter_code [code]: + "A \ set xs = set (List.filter (\x. x \ A) xs)" + "A \ coset xs = foldr Set.remove xs A" + by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr) + +lemma subtract_code [code]: + "A - set xs = foldr Set.remove xs A" + "A - coset xs = set (List.filter (\x. x \ A) xs)" + by (auto simp add: minus_set_foldr) + +lemma union_code [code]: + "set xs \ A = foldr insert xs A" + "coset xs \ A = coset (List.filter (\x. x \ A) xs)" + by (auto simp add: union_set_foldr) + +definition Inf :: "'a::complete_lattice set \ 'a" where + [simp]: "Inf = Complete_Lattices.Inf" + +hide_const (open) Inf + +lemma [code_unfold_post]: + "Inf = More_Set.Inf" + by simp + +definition Sup :: "'a::complete_lattice set \ 'a" where + [simp]: "Sup = Complete_Lattices.Sup" + +hide_const (open) Sup + +lemma [code_unfold_post]: + "Sup = More_Set.Sup" + by simp + +lemma Inf_code [code]: + "More_Set.Inf (set xs) = foldr inf xs top" + "More_Set.Inf (coset []) = bot" + by (simp_all add: Inf_set_foldr) + +lemma Sup_sup [code]: + "More_Set.Sup (set xs) = foldr sup xs bot" + "More_Set.Sup (coset []) = top" + by (simp_all add: Sup_set_foldr) + +lemma INF_code [code]: + "INFI (set xs) f = foldr (inf \ f) xs top" + by (simp add: INF_def Inf_set_foldr image_set foldr_map del: set_map) + +lemma SUP_sup [code]: + "SUPR (set xs) f = foldr (sup \ f) xs bot" + by (simp add: SUP_def Sup_set_foldr image_set foldr_map del: set_map) + +hide_const (open) coset + + +subsection {* Operations on relations *} + +text {* Initially contributed by Tjark Weber. *} + +declare Domain_fst [code] + +declare Range_snd [code] + +declare trans_join [code] + +declare irrefl_distinct [code] + +declare trancl_set_ntrancl [code] + +declare acyclic_irrefl [code] + +declare product_code [code] + +declare Id_on_set [code] + +declare set_rel_comp [code] + +declare wf_set [code] + +end + diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Quotient_Examples/DList.thy --- a/src/HOL/Quotient_Examples/DList.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Quotient_Examples/DList.thy Mon Dec 26 22:17:10 2011 +0100 @@ -8,7 +8,7 @@ header {* Lists with distinct elements *} theory DList -imports "~~/src/HOL/Library/Quotient_List" "~~/src/HOL/Library/More_List" +imports "~~/src/HOL/Library/Quotient_List" begin text {* Some facts about lists *} diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Quotient_Examples/FSet.thy --- a/src/HOL/Quotient_Examples/FSet.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Quotient_Examples/FSet.thy Mon Dec 26 22:17:10 2011 +0100 @@ -6,7 +6,7 @@ *) theory FSet -imports "~~/src/HOL/Library/Quotient_List" "~~/src/HOL/Library/More_List" +imports "~~/src/HOL/Library/Quotient_List" begin text {* diff -r b39256df5f8a -r b7b905b23b2a src/HOL/Quotient_Examples/Quotient_Cset.thy --- a/src/HOL/Quotient_Examples/Quotient_Cset.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/Quotient_Examples/Quotient_Cset.thy Mon Dec 26 22:17:10 2011 +0100 @@ -5,7 +5,7 @@ header {* A variant of theory Cset from Library, defined as a quotient *} theory Quotient_Cset -imports "~~/src/HOL/Library/More_Set" "~~/src/HOL/Library/More_List" "~~/src/HOL/Library/Quotient_Syntax" +imports Main "~~/src/HOL/Library/Quotient_Syntax" begin subsection {* Lifting *} diff -r b39256df5f8a -r b7b905b23b2a src/HOL/ex/Quickcheck_Examples.thy --- a/src/HOL/ex/Quickcheck_Examples.thy Mon Dec 26 22:17:10 2011 +0100 +++ b/src/HOL/ex/Quickcheck_Examples.thy Mon Dec 26 22:17:10 2011 +0100 @@ -6,7 +6,7 @@ header {* Examples for the 'quickcheck' command *} theory Quickcheck_Examples -imports Complex_Main "~~/src/HOL/Library/Dlist" "~~/src/HOL/Library/More_Set" +imports Complex_Main "~~/src/HOL/Library/Dlist" begin text {*