| author | obua | 
| Sat, 30 Jun 2007 17:30:10 +0200 | |
| changeset 23521 | 195fe3fe2831 | 
| parent 23394 | 474ff28210c0 | 
| permissions | -rw-r--r-- | 
| 17632 | 1 | (* Title: HOL/Library/ExecutableSet.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Stefan Berghofer, TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Implementation of finite sets by lists *}
 | |
| 7 | ||
| 8 | theory ExecutableSet | |
| 9 | imports Main | |
| 10 | begin | |
| 11 | ||
| 22665 | 12 | subsection {* Definitional rewrites *}
 | 
| 20597 | 13 | |
| 14 | instance set :: (eq) eq .. | |
| 19791 | 15 | |
| 21153 | 16 | lemma [code target: Set]: | 
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 17 | "A = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A" | 
| 17632 | 18 | by blast | 
| 19 | ||
| 20597 | 20 | lemma [code func]: | 
| 21572 | 21 | "(A\<Colon>'a\<Colon>eq set) = B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A" | 
| 22 | by blast | |
| 23 | ||
| 24 | lemma [code func]: | |
| 22177 | 25 | "(A\<Colon>'a\<Colon>eq set) \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)" | 
| 26 | unfolding subset_def .. | |
| 20597 | 27 | |
| 21572 | 28 | lemma [code func]: | 
| 22177 | 29 | "(A\<Colon>'a\<Colon>eq set) \<subset> B \<longleftrightarrow> A \<subseteq> B \<and> A \<noteq> B" | 
| 30 | by blast | |
| 21572 | 31 | |
| 21323 | 32 | lemma [code]: | 
| 33 | "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. x = a)" | |
| 34 | unfolding bex_triv_one_point1 .. | |
| 17632 | 35 | |
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 36 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 37 |   filter_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
 | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 38 |   "filter_set P xs = {x\<in>xs. P x}"
 | 
| 20597 | 39 | |
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 40 | lemmas [symmetric, code inline] = filter_set_def | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 41 | |
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 42 | |
| 22665 | 43 | subsection {* Operations on lists *}
 | 
| 19791 | 44 | |
| 22665 | 45 | subsubsection {* Basic definitions *}
 | 
| 19791 | 46 | |
| 47 | definition | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 48 |   flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
 | 
| 19791 | 49 | "flip f a b = f b a" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 50 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 51 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 52 | member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 22350 | 53 | "member xs x \<longleftrightarrow> x \<in> set xs" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 54 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 55 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 56 | insertl :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where | 
| 19791 | 57 | "insertl x xs = (if member xs x then xs else x#xs)" | 
| 58 | ||
| 23394 | 59 | lemma [code target: List]: "member [] y \<longleftrightarrow> False" | 
| 22350 | 60 | and [code target: List]: "member (x#xs) y \<longleftrightarrow> y = x \<or> member xs y" | 
| 23394 | 61 | unfolding member_def by (induct xs) simp_all | 
| 19791 | 62 | |
| 22177 | 63 | fun | 
| 64 |   drop_first :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
 | |
| 19791 | 65 | "drop_first f [] = []" | 
| 22492 | 66 | | "drop_first f (x#xs) = (if f x then xs else x # drop_first f xs)" | 
| 19791 | 67 | declare drop_first.simps [code del] | 
| 68 | declare drop_first.simps [code target: List] | |
| 69 | ||
| 70 | declare remove1.simps [code del] | |
| 71 | lemma [code target: List]: | |
| 72 | "remove1 x xs = (if member xs x then drop_first (\<lambda>y. y = x) xs else xs)" | |
| 73 | proof (cases "member xs x") | |
| 74 | case False thus ?thesis unfolding member_def by (induct xs) auto | |
| 75 | next | |
| 76 | case True | |
| 77 | have "remove1 x xs = drop_first (\<lambda>y. y = x) xs" by (induct xs) simp_all | |
| 78 | with True show ?thesis by simp | |
| 79 | qed | |
| 80 | ||
| 81 | lemma member_nil [simp]: | |
| 82 | "member [] = (\<lambda>x. False)" | |
| 83 | proof | |
| 84 | fix x | |
| 85 | show "member [] x = False" unfolding member_def by simp | |
| 86 | qed | |
| 87 | ||
| 88 | lemma member_insertl [simp]: | |
| 89 | "x \<in> set (insertl x xs)" | |
| 90 | unfolding insertl_def member_def mem_iff by simp | |
| 91 | ||
| 92 | lemma insertl_member [simp]: | |
| 93 | fixes xs x | |
| 94 | assumes member: "member xs x" | |
| 95 | shows "insertl x xs = xs" | |
| 96 | using member unfolding insertl_def by simp | |
| 97 | ||
| 98 | lemma insertl_not_member [simp]: | |
| 99 | fixes xs x | |
| 100 | assumes member: "\<not> (member xs x)" | |
| 101 | shows "insertl x xs = x # xs" | |
| 102 | using member unfolding insertl_def by simp | |
| 103 | ||
| 104 | lemma foldr_remove1_empty [simp]: | |
| 105 | "foldr remove1 xs [] = []" | |
| 106 | by (induct xs) simp_all | |
| 107 | ||
| 108 | ||
| 22665 | 109 | subsubsection {* Derived definitions *}
 | 
| 19791 | 110 | |
| 20934 | 111 | function unionl :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 112 | where | 
| 19791 | 113 | "unionl [] ys = ys" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 114 | | "unionl xs ys = foldr insertl xs ys" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 115 | by pat_completeness auto | 
| 21321 | 116 | termination by lexicographic_order | 
| 117 | ||
| 19791 | 118 | lemmas unionl_def = unionl.simps(2) | 
| 119 | ||
| 20934 | 120 | function intersect :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 121 | where | 
| 19791 | 122 | "intersect [] ys = []" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 123 | | "intersect xs [] = []" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 124 | | "intersect xs ys = filter (member xs) ys" | 
| 21321 | 125 | by pat_completeness auto | 
| 126 | termination by lexicographic_order | |
| 127 | ||
| 19791 | 128 | lemmas intersect_def = intersect.simps(3) | 
| 129 | ||
| 20934 | 130 | function subtract :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 131 | where | 
| 19791 | 132 | "subtract [] ys = ys" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 133 | | "subtract xs [] = []" | 
| 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 134 | | "subtract xs ys = foldr remove1 xs ys" | 
| 21321 | 135 | by pat_completeness auto | 
| 136 | termination by lexicographic_order | |
| 137 | ||
| 19791 | 138 | lemmas subtract_def = subtract.simps(3) | 
| 139 | ||
| 20934 | 140 | function map_distinct :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"
 | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 141 | where | 
| 19791 | 142 | "map_distinct f [] = []" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 143 | | "map_distinct f xs = foldr (insertl o f) xs []" | 
| 21321 | 144 | by pat_completeness auto | 
| 145 | termination by lexicographic_order | |
| 146 | ||
| 19791 | 147 | lemmas map_distinct_def = map_distinct.simps(2) | 
| 148 | ||
| 20934 | 149 | function unions :: "'a list list \<Rightarrow> 'a list" | 
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 150 | where | 
| 19791 | 151 | "unions [] = []" | 
| 22492 | 152 | | "unions xs = foldr unionl xs []" | 
| 21321 | 153 | by pat_completeness auto | 
| 154 | termination by lexicographic_order | |
| 155 | ||
| 19791 | 156 | lemmas unions_def = unions.simps(2) | 
| 157 | ||
| 20934 | 158 | consts intersects :: "'a list list \<Rightarrow> 'a list" | 
| 19791 | 159 | primrec | 
| 160 | "intersects (x#xs) = foldr intersect xs x" | |
| 161 | ||
| 162 | definition | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 163 |   map_union :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
 | 
| 19791 | 164 | "map_union xs f = unions (map f xs)" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 165 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 166 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 167 |   map_inter :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
 | 
| 19791 | 168 | "map_inter xs f = intersects (map f xs)" | 
| 169 | ||
| 170 | ||
| 22665 | 171 | subsection {* Isomorphism proofs *}
 | 
| 19791 | 172 | |
| 173 | lemma iso_member: | |
| 22350 | 174 | "member xs x \<longleftrightarrow> x \<in> set xs" | 
| 19791 | 175 | unfolding member_def mem_iff .. | 
| 176 | ||
| 177 | lemma iso_insert: | |
| 178 | "set (insertl x xs) = insert x (set xs)" | |
| 179 | unfolding insertl_def iso_member by (simp add: Set.insert_absorb) | |
| 180 | ||
| 181 | lemma iso_remove1: | |
| 182 | assumes distnct: "distinct xs" | |
| 183 |   shows "set (remove1 x xs) = set xs - {x}"
 | |
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 184 | using distnct set_remove1_eq by auto | 
| 19791 | 185 | |
| 186 | lemma iso_union: | |
| 187 | "set (unionl xs ys) = set xs \<union> set ys" | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
20503diff
changeset | 188 | unfolding unionl_def | 
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 189 | by (induct xs arbitrary: ys) (simp_all add: iso_insert) | 
| 19791 | 190 | |
| 191 | lemma iso_intersect: | |
| 192 | "set (intersect xs ys) = set xs \<inter> set ys" | |
| 193 | unfolding intersect_def Int_def by (simp add: Int_def iso_member) auto | |
| 194 | ||
| 22177 | 195 | definition | 
| 196 | subtract' :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where | |
| 197 | "subtract' = flip subtract" | |
| 198 | ||
| 19791 | 199 | lemma iso_subtract: | 
| 200 | fixes ys | |
| 201 | assumes distnct: "distinct ys" | |
| 22177 | 202 | shows "set (subtract' ys xs) = set ys - set xs" | 
| 23394 | 203 | and "distinct (subtract' ys xs)" | 
| 22177 | 204 | unfolding subtract'_def flip_def subtract_def | 
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 205 | using distnct by (induct xs arbitrary: ys) auto | 
| 19791 | 206 | |
| 207 | lemma iso_map_distinct: | |
| 208 | "set (map_distinct f xs) = image f (set xs)" | |
| 209 | unfolding map_distinct_def by (induct xs) (simp_all add: iso_insert) | |
| 210 | ||
| 211 | lemma iso_unions: | |
| 212 | "set (unions xss) = \<Union> set (map set xss)" | |
| 23394 | 213 | unfolding unions_def | 
| 214 | proof (induct xss) | |
| 19791 | 215 | case Nil show ?case by simp | 
| 216 | next | |
| 217 | case (Cons xs xss) thus ?case by (induct xs) (simp_all add: iso_insert) | |
| 218 | qed | |
| 219 | ||
| 220 | lemma iso_intersects: | |
| 221 | "set (intersects (xs#xss)) = \<Inter> set (map set (xs#xss))" | |
| 222 | by (induct xss) (simp_all add: Int_def iso_member, auto) | |
| 223 | ||
| 224 | lemma iso_UNION: | |
| 225 | "set (map_union xs f) = UNION (set xs) (set o f)" | |
| 226 | unfolding map_union_def iso_unions by simp | |
| 227 | ||
| 228 | lemma iso_INTER: | |
| 229 | "set (map_inter (x#xs) f) = INTER (set (x#xs)) (set o f)" | |
| 230 | unfolding map_inter_def iso_intersects by (induct xs) (simp_all add: iso_member, auto) | |
| 231 | ||
| 232 | definition | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 233 |   Blall :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 19791 | 234 | "Blall = flip list_all" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 235 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21385diff
changeset | 236 |   Blex :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 19791 | 237 | "Blex = flip list_ex" | 
| 238 | ||
| 239 | lemma iso_Ball: | |
| 240 | "Blall xs f = Ball (set xs) f" | |
| 241 | unfolding Blall_def flip_def by (induct xs) simp_all | |
| 242 | ||
| 243 | lemma iso_Bex: | |
| 244 | "Blex xs f = Bex (set xs) f" | |
| 245 | unfolding Blex_def flip_def by (induct xs) simp_all | |
| 246 | ||
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 247 | lemma iso_filter: | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 248 | "set (filter P xs) = filter_set P (set xs)" | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 249 | unfolding filter_set_def by (induct xs) auto | 
| 19791 | 250 | |
| 22665 | 251 | subsection {* code generator setup *}
 | 
| 19791 | 252 | |
| 21008 | 253 | ML {*
 | 
| 254 | nonfix inter; | |
| 255 | nonfix union; | |
| 21875 | 256 | nonfix subset; | 
| 21008 | 257 | *} | 
| 258 | ||
| 21191 | 259 | code_modulename SML | 
| 260 | ExecutableSet List | |
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 261 | Set List | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 262 | |
| 21911 
e29bcab0c81c
added OCaml code generation (without dictionaries)
 haftmann parents: 
21875diff
changeset | 263 | code_modulename OCaml | 
| 
e29bcab0c81c
added OCaml code generation (without dictionaries)
 haftmann parents: 
21875diff
changeset | 264 | ExecutableSet List | 
| 
e29bcab0c81c
added OCaml code generation (without dictionaries)
 haftmann parents: 
21875diff
changeset | 265 | Set List | 
| 
e29bcab0c81c
added OCaml code generation (without dictionaries)
 haftmann parents: 
21875diff
changeset | 266 | |
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 267 | code_modulename Haskell | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 268 | ExecutableSet List | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 269 | Set List | 
| 20934 | 270 | |
| 21063 | 271 | definition [code inline]: | 
| 20934 | 272 | "empty_list = []" | 
| 273 | ||
| 274 | lemma [code func]: | |
| 275 | "insert (x \<Colon> 'a\<Colon>eq) = insert x" .. | |
| 276 | ||
| 277 | lemma [code func]: | |
| 278 | "(xs \<Colon> 'a\<Colon>eq set) \<union> ys = xs \<union> ys" .. | |
| 279 | ||
| 280 | lemma [code func]: | |
| 281 | "(xs \<Colon> 'a\<Colon>eq set) \<inter> ys = xs \<inter> ys" .. | |
| 282 | ||
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 283 | lemma [code func]: | 
| 22177 | 284 | "(op -) (xs \<Colon> 'a\<Colon>eq set) = (op -) (xs \<Colon> 'a\<Colon>eq set)" .. | 
| 20934 | 285 | |
| 286 | lemma [code func]: | |
| 287 | "image (f \<Colon> 'a \<Rightarrow> 'b\<Colon>eq) = image f" .. | |
| 288 | ||
| 289 | lemma [code func]: | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22665diff
changeset | 290 | "Union (xss \<Colon> 'a\<Colon>eq set set) = Union xss" .. | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22665diff
changeset | 291 | |
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22665diff
changeset | 292 | lemma [code func]: | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22665diff
changeset | 293 | "Inter (xss \<Colon> 'a\<Colon>eq set set) = Inter xss" .. | 
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22665diff
changeset | 294 | |
| 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22665diff
changeset | 295 | lemma [code func]: | 
| 20934 | 296 | "UNION xs (f \<Colon> 'a \<Rightarrow> 'b\<Colon>eq set) = UNION xs f" .. | 
| 297 | ||
| 298 | lemma [code func]: | |
| 299 | "INTER xs (f \<Colon> 'a \<Rightarrow> 'b\<Colon>eq set) = INTER xs f" .. | |
| 300 | ||
| 301 | lemma [code func]: | |
| 302 | "Ball (xs \<Colon> 'a\<Colon>type set) = Ball xs" .. | |
| 303 | ||
| 304 | lemma [code func]: | |
| 305 | "Bex (xs \<Colon> 'a\<Colon>type set) = Bex xs" .. | |
| 306 | ||
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 307 | lemma [code func]: | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 308 | "filter_set P (xs \<Colon> 'a\<Colon>type set) = filter_set P xs" .. | 
| 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 309 | |
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22665diff
changeset | 310 | |
| 20934 | 311 | code_abstype "'a set" "'a list" where | 
| 21115 | 312 |   "{}" \<equiv> empty_list
 | 
| 20934 | 313 | insert \<equiv> insertl | 
| 314 | "op \<union>" \<equiv> unionl | |
| 315 | "op \<inter>" \<equiv> intersect | |
| 22177 | 316 | "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" \<equiv> subtract' | 
| 20934 | 317 | image \<equiv> map_distinct | 
| 318 | Union \<equiv> unions | |
| 319 | Inter \<equiv> intersects | |
| 320 | UNION \<equiv> map_union | |
| 321 | INTER \<equiv> map_inter | |
| 322 | Ball \<equiv> Blall | |
| 323 | Bex \<equiv> Blex | |
| 21385 
cf398bb8a8be
added filter_set; adaptions to more strict type discipline for code lemmas
 haftmann parents: 
21323diff
changeset | 324 | filter_set \<equiv> filter | 
| 20934 | 325 | |
| 326 | ||
| 22665 | 327 | subsubsection {* type serializations *}
 | 
| 19791 | 328 | |
| 17632 | 329 | types_code | 
| 330 |   set ("_ list")
 | |
| 331 | attach (term_of) {*
 | |
| 332 | fun term_of_set f T [] = Const ("{}", Type ("set", [T]))
 | |
| 333 |   | term_of_set f T (x :: xs) = Const ("insert",
 | |
| 334 |       T --> Type ("set", [T]) --> Type ("set", [T])) $ f x $ term_of_set f T xs;
 | |
| 335 | *} | |
| 336 | attach (test) {*
 | |
| 337 | fun gen_set' aG i j = frequency | |
| 338 | [(i, fn () => aG j :: gen_set' aG (i-1) j), (1, fn () => [])] () | |
| 339 | and gen_set aG i = gen_set' aG i i; | |
| 340 | *} | |
| 341 | ||
| 19791 | 342 | |
| 22665 | 343 | subsubsection {* const serializations *}
 | 
| 18702 | 344 | |
| 17632 | 345 | consts_code | 
| 22921 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 346 |   "{}" ("{*[]*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 347 |   insert ("{*insertl*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 348 |   "op \<union>" ("{*unionl*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 349 |   "op \<inter>" ("{*intersect*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 350 |   "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" ("{* flip subtract *}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 351 |   image ("{*map_distinct*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 352 |   Union ("{*unions*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 353 |   Inter ("{*intersects*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 354 |   UNION ("{*map_union*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 355 |   INTER ("{*map_inter*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 356 |   Ball ("{*Blall*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 357 |   Bex ("{*Blex*}")
 | 
| 
475ff421a6a3
consts in consts_code Isar commands are now referred to by usual term syntax
 haftmann parents: 
22838diff
changeset | 358 |   filter_set ("{*filter*}")
 | 
| 17632 | 359 | |
| 360 | end |