| 23854 |      1 | (*  Title:      HOL/Library/Executable_Set.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Stefan Berghofer, TU Muenchen
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Implementation of finite sets by lists *}
 | 
|  |      7 | 
 | 
|  |      8 | theory Executable_Set
 | 
| 27487 |      9 | imports Plain "~~/src/HOL/List"
 | 
| 23854 |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | subsection {* Definitional rewrites *}
 | 
|  |     13 | 
 | 
| 28522 |     14 | definition subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
 | 
|  |     15 |   "subset = op \<le>"
 | 
|  |     16 | 
 | 
|  |     17 | declare subset_def [symmetric, code unfold]
 | 
|  |     18 | 
 | 
| 28939 |     19 | lemma [code]: "subset A B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)"
 | 
| 28522 |     20 |   unfolding subset_def subset_eq ..
 | 
| 23854 |     21 | 
 | 
| 28522 |     22 | definition is_empty :: "'a set \<Rightarrow> bool" where
 | 
|  |     23 |   "is_empty A \<longleftrightarrow> A = {}"
 | 
|  |     24 | 
 | 
|  |     25 | definition eq_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
 | 
|  |     26 |   [code del]: "eq_set = op ="
 | 
|  |     27 | 
 | 
|  |     28 | lemma [code]: "eq_set A B \<longleftrightarrow> A \<subseteq> B \<and> B \<subseteq> A"
 | 
|  |     29 |   unfolding eq_set_def by auto
 | 
| 26816 |     30 | 
 | 
| 23854 |     31 | lemma [code]:
 | 
|  |     32 |   "a \<in> A \<longleftrightarrow> (\<exists>x\<in>A. x = a)"
 | 
|  |     33 |   unfolding bex_triv_one_point1 ..
 | 
|  |     34 | 
 | 
| 28939 |     35 | definition filter_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
 | 
| 23854 |     36 |   "filter_set P xs = {x\<in>xs. P x}"
 | 
|  |     37 | 
 | 
|  |     38 | 
 | 
|  |     39 | subsection {* Operations on lists *}
 | 
|  |     40 | 
 | 
|  |     41 | subsubsection {* Basic definitions *}
 | 
|  |     42 | 
 | 
|  |     43 | definition
 | 
|  |     44 |   flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
 | 
|  |     45 |   "flip f a b = f b a"
 | 
|  |     46 | 
 | 
|  |     47 | definition
 | 
|  |     48 |   member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
 | 
|  |     49 |   "member xs x \<longleftrightarrow> x \<in> set xs"
 | 
|  |     50 | 
 | 
|  |     51 | definition
 | 
|  |     52 |   insertl :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
 | 
|  |     53 |   "insertl x xs = (if member xs x then xs else x#xs)"
 | 
|  |     54 | 
 | 
|  |     55 | lemma [code target: List]: "member [] y \<longleftrightarrow> False"
 | 
|  |     56 |   and [code target: List]: "member (x#xs) y \<longleftrightarrow> y = x \<or> member xs y"
 | 
|  |     57 |   unfolding member_def by (induct xs) simp_all
 | 
|  |     58 | 
 | 
|  |     59 | fun
 | 
|  |     60 |   drop_first :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
 | 
|  |     61 |   "drop_first f [] = []"
 | 
|  |     62 | | "drop_first f (x#xs) = (if f x then xs else x # drop_first f xs)"
 | 
|  |     63 | declare drop_first.simps [code del]
 | 
|  |     64 | declare drop_first.simps [code target: List]
 | 
|  |     65 | 
 | 
|  |     66 | declare remove1.simps [code del]
 | 
|  |     67 | lemma [code target: List]:
 | 
|  |     68 |   "remove1 x xs = (if member xs x then drop_first (\<lambda>y. y = x) xs else xs)"
 | 
|  |     69 | proof (cases "member xs x")
 | 
|  |     70 |   case False thus ?thesis unfolding member_def by (induct xs) auto
 | 
|  |     71 | next
 | 
|  |     72 |   case True
 | 
|  |     73 |   have "remove1 x xs = drop_first (\<lambda>y. y = x) xs" by (induct xs) simp_all
 | 
|  |     74 |   with True show ?thesis by simp
 | 
|  |     75 | qed
 | 
|  |     76 | 
 | 
|  |     77 | lemma member_nil [simp]:
 | 
|  |     78 |   "member [] = (\<lambda>x. False)"
 | 
| 26816 |     79 | proof (rule ext)
 | 
| 23854 |     80 |   fix x
 | 
|  |     81 |   show "member [] x = False" unfolding member_def by simp
 | 
|  |     82 | qed
 | 
|  |     83 | 
 | 
|  |     84 | lemma member_insertl [simp]:
 | 
|  |     85 |   "x \<in> set (insertl x xs)"
 | 
|  |     86 |   unfolding insertl_def member_def mem_iff by simp
 | 
|  |     87 | 
 | 
|  |     88 | lemma insertl_member [simp]:
 | 
|  |     89 |   fixes xs x
 | 
|  |     90 |   assumes member: "member xs x"
 | 
|  |     91 |   shows "insertl x xs = xs"
 | 
|  |     92 |   using member unfolding insertl_def by simp
 | 
|  |     93 | 
 | 
|  |     94 | lemma insertl_not_member [simp]:
 | 
|  |     95 |   fixes xs x
 | 
|  |     96 |   assumes member: "\<not> (member xs x)"
 | 
|  |     97 |   shows "insertl x xs = x # xs"
 | 
|  |     98 |   using member unfolding insertl_def by simp
 | 
|  |     99 | 
 | 
|  |    100 | lemma foldr_remove1_empty [simp]:
 | 
|  |    101 |   "foldr remove1 xs [] = []"
 | 
|  |    102 |   by (induct xs) simp_all
 | 
|  |    103 | 
 | 
|  |    104 | 
 | 
|  |    105 | subsubsection {* Derived definitions *}
 | 
|  |    106 | 
 | 
|  |    107 | function unionl :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 | 
|  |    108 | where
 | 
|  |    109 |   "unionl [] ys = ys"
 | 
|  |    110 | | "unionl xs ys = foldr insertl xs ys"
 | 
|  |    111 | by pat_completeness auto
 | 
|  |    112 | termination by lexicographic_order
 | 
|  |    113 | 
 | 
| 26312 |    114 | lemmas unionl_eq = unionl.simps(2)
 | 
| 23854 |    115 | 
 | 
|  |    116 | function intersect :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 | 
|  |    117 | where
 | 
|  |    118 |   "intersect [] ys = []"
 | 
|  |    119 | | "intersect xs [] = []"
 | 
|  |    120 | | "intersect xs ys = filter (member xs) ys"
 | 
|  |    121 | by pat_completeness auto
 | 
|  |    122 | termination by lexicographic_order
 | 
|  |    123 | 
 | 
| 26312 |    124 | lemmas intersect_eq = intersect.simps(3)
 | 
| 23854 |    125 | 
 | 
|  |    126 | function subtract :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 | 
|  |    127 | where
 | 
|  |    128 |   "subtract [] ys = ys"
 | 
|  |    129 | | "subtract xs [] = []"
 | 
|  |    130 | | "subtract xs ys = foldr remove1 xs ys"
 | 
|  |    131 | by pat_completeness auto
 | 
|  |    132 | termination by lexicographic_order
 | 
|  |    133 | 
 | 
| 26312 |    134 | lemmas subtract_eq = subtract.simps(3)
 | 
| 23854 |    135 | 
 | 
|  |    136 | function map_distinct :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"
 | 
|  |    137 | where
 | 
|  |    138 |   "map_distinct f [] = []"
 | 
|  |    139 | | "map_distinct f xs = foldr (insertl o f) xs []"
 | 
|  |    140 | by pat_completeness auto
 | 
|  |    141 | termination by lexicographic_order
 | 
|  |    142 | 
 | 
| 26312 |    143 | lemmas map_distinct_eq = map_distinct.simps(2)
 | 
| 23854 |    144 | 
 | 
|  |    145 | function unions :: "'a list list \<Rightarrow> 'a list"
 | 
|  |    146 | where
 | 
|  |    147 |   "unions [] = []"
 | 
|  |    148 | | "unions xs = foldr unionl xs []"
 | 
|  |    149 | by pat_completeness auto
 | 
|  |    150 | termination by lexicographic_order
 | 
|  |    151 | 
 | 
| 26312 |    152 | lemmas unions_eq = unions.simps(2)
 | 
| 23854 |    153 | 
 | 
|  |    154 | consts intersects :: "'a list list \<Rightarrow> 'a list"
 | 
|  |    155 | primrec
 | 
|  |    156 |   "intersects (x#xs) = foldr intersect xs x"
 | 
|  |    157 | 
 | 
|  |    158 | definition
 | 
|  |    159 |   map_union :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
 | 
|  |    160 |   "map_union xs f = unions (map f xs)"
 | 
|  |    161 | 
 | 
|  |    162 | definition
 | 
|  |    163 |   map_inter :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
 | 
|  |    164 |   "map_inter xs f = intersects (map f xs)"
 | 
|  |    165 | 
 | 
|  |    166 | 
 | 
|  |    167 | subsection {* Isomorphism proofs *}
 | 
|  |    168 | 
 | 
|  |    169 | lemma iso_member:
 | 
|  |    170 |   "member xs x \<longleftrightarrow> x \<in> set xs"
 | 
|  |    171 |   unfolding member_def mem_iff ..
 | 
|  |    172 | 
 | 
|  |    173 | lemma iso_insert:
 | 
|  |    174 |   "set (insertl x xs) = insert x (set xs)"
 | 
|  |    175 |   unfolding insertl_def iso_member by (simp add: Set.insert_absorb)
 | 
|  |    176 | 
 | 
|  |    177 | lemma iso_remove1:
 | 
|  |    178 |   assumes distnct: "distinct xs"
 | 
|  |    179 |   shows "set (remove1 x xs) = set xs - {x}"
 | 
|  |    180 |   using distnct set_remove1_eq by auto
 | 
|  |    181 | 
 | 
|  |    182 | lemma iso_union:
 | 
|  |    183 |   "set (unionl xs ys) = set xs \<union> set ys"
 | 
| 26312 |    184 |   unfolding unionl_eq
 | 
| 23854 |    185 |   by (induct xs arbitrary: ys) (simp_all add: iso_insert)
 | 
|  |    186 | 
 | 
|  |    187 | lemma iso_intersect:
 | 
|  |    188 |   "set (intersect xs ys) = set xs \<inter> set ys"
 | 
| 26312 |    189 |   unfolding intersect_eq Int_def by (simp add: Int_def iso_member) auto
 | 
| 23854 |    190 | 
 | 
|  |    191 | definition
 | 
|  |    192 |   subtract' :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
 | 
|  |    193 |   "subtract' = flip subtract"
 | 
|  |    194 | 
 | 
|  |    195 | lemma iso_subtract:
 | 
|  |    196 |   fixes ys
 | 
|  |    197 |   assumes distnct: "distinct ys"
 | 
|  |    198 |   shows "set (subtract' ys xs) = set ys - set xs"
 | 
|  |    199 |     and "distinct (subtract' ys xs)"
 | 
| 26312 |    200 |   unfolding subtract'_def flip_def subtract_eq
 | 
| 23854 |    201 |   using distnct by (induct xs arbitrary: ys) auto
 | 
|  |    202 | 
 | 
|  |    203 | lemma iso_map_distinct:
 | 
|  |    204 |   "set (map_distinct f xs) = image f (set xs)"
 | 
| 26312 |    205 |   unfolding map_distinct_eq by (induct xs) (simp_all add: iso_insert)
 | 
| 23854 |    206 | 
 | 
|  |    207 | lemma iso_unions:
 | 
|  |    208 |   "set (unions xss) = \<Union> set (map set xss)"
 | 
| 26312 |    209 |   unfolding unions_eq
 | 
| 23854 |    210 | proof (induct xss)
 | 
|  |    211 |   case Nil show ?case by simp
 | 
|  |    212 | next
 | 
|  |    213 |   case (Cons xs xss) thus ?case by (induct xs) (simp_all add: iso_insert)
 | 
|  |    214 | qed
 | 
|  |    215 | 
 | 
|  |    216 | lemma iso_intersects:
 | 
|  |    217 |   "set (intersects (xs#xss)) = \<Inter> set (map set (xs#xss))"
 | 
|  |    218 |   by (induct xss) (simp_all add: Int_def iso_member, auto)
 | 
|  |    219 | 
 | 
|  |    220 | lemma iso_UNION:
 | 
|  |    221 |   "set (map_union xs f) = UNION (set xs) (set o f)"
 | 
|  |    222 |   unfolding map_union_def iso_unions by simp
 | 
|  |    223 | 
 | 
|  |    224 | lemma iso_INTER:
 | 
|  |    225 |   "set (map_inter (x#xs) f) = INTER (set (x#xs)) (set o f)"
 | 
|  |    226 |   unfolding map_inter_def iso_intersects by (induct xs) (simp_all add: iso_member, auto)
 | 
|  |    227 | 
 | 
|  |    228 | definition
 | 
|  |    229 |   Blall :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
|  |    230 |   "Blall = flip list_all"
 | 
|  |    231 | definition
 | 
|  |    232 |   Blex :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
|  |    233 |   "Blex = flip list_ex"
 | 
|  |    234 | 
 | 
|  |    235 | lemma iso_Ball:
 | 
|  |    236 |   "Blall xs f = Ball (set xs) f"
 | 
|  |    237 |   unfolding Blall_def flip_def by (induct xs) simp_all
 | 
|  |    238 | 
 | 
|  |    239 | lemma iso_Bex:
 | 
|  |    240 |   "Blex xs f = Bex (set xs) f"
 | 
|  |    241 |   unfolding Blex_def flip_def by (induct xs) simp_all
 | 
|  |    242 | 
 | 
|  |    243 | lemma iso_filter:
 | 
|  |    244 |   "set (filter P xs) = filter_set P (set xs)"
 | 
|  |    245 |   unfolding filter_set_def by (induct xs) auto
 | 
|  |    246 | 
 | 
|  |    247 | subsection {* code generator setup *}
 | 
|  |    248 | 
 | 
|  |    249 | ML {*
 | 
|  |    250 | nonfix inter;
 | 
|  |    251 | nonfix union;
 | 
|  |    252 | nonfix subset;
 | 
|  |    253 | *}
 | 
|  |    254 | 
 | 
|  |    255 | subsubsection {* const serializations *}
 | 
|  |    256 | 
 | 
|  |    257 | consts_code
 | 
|  |    258 |   "{}" ("{*[]*}")
 | 
|  |    259 |   insert ("{*insertl*}")
 | 
| 28522 |    260 |   is_empty ("{*null*}")
 | 
| 23854 |    261 |   "op \<union>" ("{*unionl*}")
 | 
|  |    262 |   "op \<inter>" ("{*intersect*}")
 | 
|  |    263 |   "op - \<Colon> 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" ("{* flip subtract *}")
 | 
|  |    264 |   image ("{*map_distinct*}")
 | 
|  |    265 |   Union ("{*unions*}")
 | 
|  |    266 |   Inter ("{*intersects*}")
 | 
|  |    267 |   UNION ("{*map_union*}")
 | 
|  |    268 |   INTER ("{*map_inter*}")
 | 
|  |    269 |   Ball ("{*Blall*}")
 | 
|  |    270 |   Bex ("{*Blex*}")
 | 
|  |    271 |   filter_set ("{*filter*}")
 | 
|  |    272 | 
 | 
|  |    273 | end
 |