| author | hoelzl | 
| Thu, 02 Sep 2010 21:08:31 +0200 | |
| changeset 39101 | 606432dd1896 | 
| parent 38857 | 97775f3e8722 | 
| child 39190 | a2775776be3f | 
| child 39198 | f967a16dfcdd | 
| permissions | -rw-r--r-- | 
| 31807 | 1 | |
| 2 | (* Author: Florian Haftmann, TU Muenchen *) | |
| 3 | ||
| 4 | header {* Executable finite sets *}
 | |
| 5 | ||
| 31849 | 6 | theory Fset | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 7 | imports More_Set More_List | 
| 31807 | 8 | begin | 
| 9 | ||
| 10 | subsection {* Lifting *}
 | |
| 11 | ||
| 37468 | 12 | typedef (open) 'a fset = "UNIV :: 'a set set" | 
| 13 | morphisms member Fset by rule+ | |
| 31807 | 14 | |
| 37468 | 15 | lemma member_Fset [simp]: | 
| 31807 | 16 | "member (Fset A) = A" | 
| 37468 | 17 | by (rule Fset_inverse) rule | 
| 34048 | 18 | |
| 31807 | 19 | lemma Fset_member [simp]: | 
| 20 | "Fset (member A) = A" | |
| 37699 | 21 | by (fact member_inverse) | 
| 37468 | 22 | |
| 23 | declare member_inject [simp] | |
| 24 | ||
| 25 | lemma Fset_inject [simp]: | |
| 26 | "Fset A = Fset B \<longleftrightarrow> A = B" | |
| 27 | by (simp add: Fset_inject) | |
| 28 | ||
| 37473 | 29 | lemma fset_eqI: | 
| 30 | "member A = member B \<Longrightarrow> A = B" | |
| 31 | by simp | |
| 32 | ||
| 37468 | 33 | declare mem_def [simp] | 
| 31807 | 34 | |
| 35 | definition Set :: "'a list \<Rightarrow> 'a fset" where | |
| 36 | "Set xs = Fset (set xs)" | |
| 37 | ||
| 38 | lemma member_Set [simp]: | |
| 39 | "member (Set xs) = set xs" | |
| 40 | by (simp add: Set_def) | |
| 41 | ||
| 32880 | 42 | definition Coset :: "'a list \<Rightarrow> 'a fset" where | 
| 43 | "Coset xs = Fset (- set xs)" | |
| 44 | ||
| 45 | lemma member_Coset [simp]: | |
| 46 | "member (Coset xs) = - set xs" | |
| 47 | by (simp add: Coset_def) | |
| 48 | ||
| 49 | code_datatype Set Coset | |
| 50 | ||
| 51 | lemma member_code [code]: | |
| 37023 | 52 | "member (Set xs) = List.member xs" | 
| 53 | "member (Coset xs) = Not \<circ> List.member xs" | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37473diff
changeset | 54 | by (simp_all add: expand_fun_eq member_def fun_Compl_def bool_Compl_def) | 
| 32880 | 55 | |
| 56 | lemma member_image_UNIV [simp]: | |
| 57 | "member ` UNIV = UNIV" | |
| 58 | proof - | |
| 59 | have "\<And>A \<Colon> 'a set. \<exists>B \<Colon> 'a fset. A = member B" | |
| 60 | proof | |
| 61 | fix A :: "'a set" | |
| 62 | show "A = member (Fset A)" by simp | |
| 63 | qed | |
| 64 | then show ?thesis by (simp add: image_def) | |
| 65 | qed | |
| 31807 | 66 | |
| 37468 | 67 | definition (in term_syntax) | 
| 68 | setify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term) | |
| 69 | \<Rightarrow> 'a fset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where | |
| 70 |   [code_unfold]: "setify xs = Code_Evaluation.valtermify Set {\<cdot>} xs"
 | |
| 71 | ||
| 37751 | 72 | notation fcomp (infixl "\<circ>>" 60) | 
| 73 | notation scomp (infixl "\<circ>\<rightarrow>" 60) | |
| 37468 | 74 | |
| 75 | instantiation fset :: (random) random | |
| 76 | begin | |
| 77 | ||
| 78 | definition | |
| 37751 | 79 | "Quickcheck.random i = Quickcheck.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (setify xs))" | 
| 37468 | 80 | |
| 81 | instance .. | |
| 82 | ||
| 83 | end | |
| 84 | ||
| 37751 | 85 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 86 | no_notation scomp (infixl "\<circ>\<rightarrow>" 60) | |
| 37468 | 87 | |
| 31807 | 88 | |
| 34048 | 89 | subsection {* Lattice instantiation *}
 | 
| 90 | ||
| 91 | instantiation fset :: (type) boolean_algebra | |
| 92 | begin | |
| 93 | ||
| 94 | definition less_eq_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where | |
| 95 | [simp]: "A \<le> B \<longleftrightarrow> member A \<subseteq> member B" | |
| 96 | ||
| 97 | definition less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" where | |
| 98 | [simp]: "A < B \<longleftrightarrow> member A \<subset> member B" | |
| 99 | ||
| 100 | definition inf_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where | |
| 101 | [simp]: "inf A B = Fset (member A \<inter> member B)" | |
| 102 | ||
| 103 | definition sup_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where | |
| 104 | [simp]: "sup A B = Fset (member A \<union> member B)" | |
| 105 | ||
| 106 | definition bot_fset :: "'a fset" where | |
| 107 |   [simp]: "bot = Fset {}"
 | |
| 108 | ||
| 109 | definition top_fset :: "'a fset" where | |
| 110 | [simp]: "top = Fset UNIV" | |
| 111 | ||
| 112 | definition uminus_fset :: "'a fset \<Rightarrow> 'a fset" where | |
| 113 | [simp]: "- A = Fset (- (member A))" | |
| 114 | ||
| 115 | definition minus_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where | |
| 116 | [simp]: "A - B = Fset (member A - member B)" | |
| 117 | ||
| 118 | instance proof | |
| 119 | qed auto | |
| 120 | ||
| 121 | end | |
| 122 | ||
| 123 | instantiation fset :: (type) complete_lattice | |
| 124 | begin | |
| 125 | ||
| 126 | definition Inf_fset :: "'a fset set \<Rightarrow> 'a fset" where | |
| 37765 | 127 | [simp]: "Inf_fset As = Fset (Inf (image member As))" | 
| 34048 | 128 | |
| 129 | definition Sup_fset :: "'a fset set \<Rightarrow> 'a fset" where | |
| 37765 | 130 | [simp]: "Sup_fset As = Fset (Sup (image member As))" | 
| 34048 | 131 | |
| 132 | instance proof | |
| 133 | qed (auto simp add: le_fun_def le_bool_def) | |
| 134 | ||
| 135 | end | |
| 136 | ||
| 37023 | 137 | |
| 31807 | 138 | subsection {* Basic operations *}
 | 
| 139 | ||
| 140 | definition is_empty :: "'a fset \<Rightarrow> bool" where | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 141 | [simp]: "is_empty A \<longleftrightarrow> More_Set.is_empty (member A)" | 
| 31807 | 142 | |
| 143 | lemma is_empty_Set [code]: | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37473diff
changeset | 144 | "is_empty (Set xs) \<longleftrightarrow> List.null xs" | 
| 31846 | 145 | by (simp add: is_empty_set) | 
| 31807 | 146 | |
| 34048 | 147 | lemma empty_Set [code]: | 
| 148 | "bot = Set []" | |
| 37468 | 149 | by (simp add: Set_def) | 
| 31807 | 150 | |
| 34048 | 151 | lemma UNIV_Set [code]: | 
| 152 | "top = Coset []" | |
| 37468 | 153 | by (simp add: Coset_def) | 
| 31807 | 154 | |
| 155 | definition insert :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where | |
| 31846 | 156 | [simp]: "insert x A = Fset (Set.insert x (member A))" | 
| 31807 | 157 | |
| 158 | lemma insert_Set [code]: | |
| 34976 
06df18c9a091
generalized lemma foldl_apply_inv to foldl_apply
 haftmann parents: 
34048diff
changeset | 159 | "insert x (Set xs) = Set (List.insert x xs)" | 
| 
06df18c9a091
generalized lemma foldl_apply_inv to foldl_apply
 haftmann parents: 
34048diff
changeset | 160 | "insert x (Coset xs) = Coset (removeAll x xs)" | 
| 37023 | 161 | by (simp_all add: Set_def Coset_def) | 
| 31807 | 162 | |
| 163 | definition remove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 164 | [simp]: "remove x A = Fset (More_Set.remove x (member A))" | 
| 31807 | 165 | |
| 166 | lemma remove_Set [code]: | |
| 34976 
06df18c9a091
generalized lemma foldl_apply_inv to foldl_apply
 haftmann parents: 
34048diff
changeset | 167 | "remove x (Set xs) = Set (removeAll x xs)" | 
| 
06df18c9a091
generalized lemma foldl_apply_inv to foldl_apply
 haftmann parents: 
34048diff
changeset | 168 | "remove x (Coset xs) = Coset (List.insert x xs)" | 
| 
06df18c9a091
generalized lemma foldl_apply_inv to foldl_apply
 haftmann parents: 
34048diff
changeset | 169 | by (simp_all add: Set_def Coset_def remove_set_compl) | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 170 | (simp add: More_Set.remove_def) | 
| 31807 | 171 | |
| 172 | definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" where
 | |
| 31846 | 173 | [simp]: "map f A = Fset (image f (member A))" | 
| 31807 | 174 | |
| 175 | lemma map_Set [code]: | |
| 176 | "map f (Set xs) = Set (remdups (List.map f xs))" | |
| 31846 | 177 | by (simp add: Set_def) | 
| 31807 | 178 | |
| 31847 | 179 | definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
 | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 180 | [simp]: "filter P A = Fset (More_Set.project P (member A))" | 
| 31807 | 181 | |
| 31847 | 182 | lemma filter_Set [code]: | 
| 183 | "filter P (Set xs) = Set (List.filter P xs)" | |
| 31846 | 184 | by (simp add: Set_def project_set) | 
| 31807 | 185 | |
| 186 | definition forall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
 | |
| 31846 | 187 | [simp]: "forall P A \<longleftrightarrow> Ball (member A) P" | 
| 31807 | 188 | |
| 189 | lemma forall_Set [code]: | |
| 190 | "forall P (Set xs) \<longleftrightarrow> list_all P xs" | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37473diff
changeset | 191 | by (simp add: Set_def list_all_iff) | 
| 31807 | 192 | |
| 193 | definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> bool" where
 | |
| 31846 | 194 | [simp]: "exists P A \<longleftrightarrow> Bex (member A) P" | 
| 31807 | 195 | |
| 196 | lemma exists_Set [code]: | |
| 197 | "exists P (Set xs) \<longleftrightarrow> list_ex P xs" | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37473diff
changeset | 198 | by (simp add: Set_def list_ex_iff) | 
| 31846 | 199 | |
| 31849 | 200 | definition card :: "'a fset \<Rightarrow> nat" where | 
| 201 | [simp]: "card A = Finite_Set.card (member A)" | |
| 202 | ||
| 203 | lemma card_Set [code]: | |
| 204 | "card (Set xs) = length (remdups xs)" | |
| 205 | proof - | |
| 206 | have "Finite_Set.card (set (remdups xs)) = length (remdups xs)" | |
| 207 | by (rule distinct_card) simp | |
| 37023 | 208 | then show ?thesis by (simp add: Set_def) | 
| 31849 | 209 | qed | 
| 210 | ||
| 37023 | 211 | lemma compl_Set [simp, code]: | 
| 212 | "- Set xs = Coset xs" | |
| 213 | by (simp add: Set_def Coset_def) | |
| 214 | ||
| 215 | lemma compl_Coset [simp, code]: | |
| 216 | "- Coset xs = Set xs" | |
| 217 | by (simp add: Set_def Coset_def) | |
| 218 | ||
| 31846 | 219 | |
| 220 | subsection {* Derived operations *}
 | |
| 221 | ||
| 222 | lemma subfset_eq_forall [code]: | |
| 34048 | 223 | "A \<le> B \<longleftrightarrow> forall (member B) A" | 
| 31846 | 224 | by (simp add: subset_eq) | 
| 225 | ||
| 226 | lemma subfset_subfset_eq [code]: | |
| 34048 | 227 | "A < B \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> (A :: 'a fset)" | 
| 228 | by (fact less_le_not_le) | |
| 31846 | 229 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 230 | instantiation fset :: (type) equal | 
| 37468 | 231 | begin | 
| 232 | ||
| 233 | definition | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 234 | "HOL.equal A B \<longleftrightarrow> A \<le> B \<and> B \<le> (A :: 'a fset)" | 
| 37468 | 235 | |
| 236 | instance proof | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 237 | qed (simp add: equal_fset_def set_eq [symmetric]) | 
| 37468 | 238 | |
| 239 | end | |
| 31846 | 240 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 241 | lemma [code nbe]: | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 242 | "HOL.equal (A :: 'a fset) A \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 243 | by (fact equal_refl) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 244 | |
| 31807 | 245 | |
| 246 | subsection {* Functorial operations *}
 | |
| 247 | ||
| 32880 | 248 | lemma inter_project [code]: | 
| 34048 | 249 | "inf A (Set xs) = Set (List.filter (member A) xs)" | 
| 37023 | 250 | "inf A (Coset xs) = foldr remove xs A" | 
| 31807 | 251 | proof - | 
| 34048 | 252 | show "inf A (Set xs) = Set (List.filter (member A) xs)" | 
| 32880 | 253 | by (simp add: inter project_def Set_def) | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 254 | have *: "\<And>x::'a. remove = (\<lambda>x. Fset \<circ> More_Set.remove x \<circ> member)" | 
| 37023 | 255 | by (simp add: expand_fun_eq) | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 256 | have "member \<circ> fold (\<lambda>x. Fset \<circ> More_Set.remove x \<circ> member) xs = | 
| 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 257 | fold More_Set.remove xs \<circ> member" | 
| 37023 | 258 | by (rule fold_apply) (simp add: expand_fun_eq) | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 259 | then have "fold More_Set.remove xs (member A) = | 
| 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 260 | member (fold (\<lambda>x. Fset \<circ> More_Set.remove x \<circ> member) xs A)" | 
| 37023 | 261 | by (simp add: expand_fun_eq) | 
| 262 | then have "inf A (Coset xs) = fold remove xs A" | |
| 263 | by (simp add: Diff_eq [symmetric] minus_set *) | |
| 264 | moreover have "\<And>x y :: 'a. Fset.remove y \<circ> Fset.remove x = Fset.remove x \<circ> Fset.remove y" | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 265 | by (auto simp add: More_Set.remove_def * intro: ext) | 
| 37023 | 266 | ultimately show "inf A (Coset xs) = foldr remove xs A" | 
| 267 | by (simp add: foldr_fold) | |
| 31807 | 268 | qed | 
| 269 | ||
| 270 | lemma subtract_remove [code]: | |
| 37023 | 271 | "A - Set xs = foldr remove xs A" | 
| 34048 | 272 | "A - Coset xs = Set (List.filter (member A) xs)" | 
| 37023 | 273 | by (simp_all only: diff_eq compl_Set compl_Coset inter_project) | 
| 32880 | 274 | |
| 275 | lemma union_insert [code]: | |
| 37023 | 276 | "sup (Set xs) A = foldr insert xs A" | 
| 34048 | 277 | "sup (Coset xs) A = Coset (List.filter (Not \<circ> member A) xs)" | 
| 32880 | 278 | proof - | 
| 37023 | 279 | have *: "\<And>x::'a. insert = (\<lambda>x. Fset \<circ> Set.insert x \<circ> member)" | 
| 280 | by (simp add: expand_fun_eq) | |
| 281 | have "member \<circ> fold (\<lambda>x. Fset \<circ> Set.insert x \<circ> member) xs = | |
| 282 | fold Set.insert xs \<circ> member" | |
| 283 | by (rule fold_apply) (simp add: expand_fun_eq) | |
| 284 | then have "fold Set.insert xs (member A) = | |
| 285 | member (fold (\<lambda>x. Fset \<circ> Set.insert x \<circ> member) xs A)" | |
| 286 | by (simp add: expand_fun_eq) | |
| 287 | then have "sup (Set xs) A = fold insert xs A" | |
| 288 | by (simp add: union_set *) | |
| 289 | moreover have "\<And>x y :: 'a. Fset.insert y \<circ> Fset.insert x = Fset.insert x \<circ> Fset.insert y" | |
| 290 | by (auto simp add: * intro: ext) | |
| 291 | ultimately show "sup (Set xs) A = foldr insert xs A" | |
| 292 | by (simp add: foldr_fold) | |
| 34048 | 293 | show "sup (Coset xs) A = Coset (List.filter (Not \<circ> member A) xs)" | 
| 32880 | 294 | by (auto simp add: Coset_def) | 
| 31807 | 295 | qed | 
| 296 | ||
| 34048 | 297 | context complete_lattice | 
| 298 | begin | |
| 31807 | 299 | |
| 34048 | 300 | definition Infimum :: "'a fset \<Rightarrow> 'a" where | 
| 301 | [simp]: "Infimum A = Inf (member A)" | |
| 31807 | 302 | |
| 34048 | 303 | lemma Infimum_inf [code]: | 
| 37023 | 304 | "Infimum (Set As) = foldr inf As top" | 
| 34048 | 305 | "Infimum (Coset []) = bot" | 
| 37023 | 306 | by (simp_all add: Inf_set_foldr Inf_UNIV) | 
| 31807 | 307 | |
| 34048 | 308 | definition Supremum :: "'a fset \<Rightarrow> 'a" where | 
| 309 | [simp]: "Supremum A = Sup (member A)" | |
| 310 | ||
| 311 | lemma Supremum_sup [code]: | |
| 37023 | 312 | "Supremum (Set As) = foldr sup As bot" | 
| 34048 | 313 | "Supremum (Coset []) = top" | 
| 37023 | 314 | by (simp_all add: Sup_set_foldr Sup_UNIV) | 
| 34048 | 315 | |
| 316 | end | |
| 31807 | 317 | |
| 318 | ||
| 31846 | 319 | subsection {* Simplified simprules *}
 | 
| 320 | ||
| 321 | lemma is_empty_simp [simp]: | |
| 322 |   "is_empty A \<longleftrightarrow> member A = {}"
 | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 323 | by (simp add: More_Set.is_empty_def) | 
| 31846 | 324 | declare is_empty_def [simp del] | 
| 325 | ||
| 326 | lemma remove_simp [simp]: | |
| 327 |   "remove x A = Fset (member A - {x})"
 | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 328 | by (simp add: More_Set.remove_def) | 
| 31846 | 329 | declare remove_def [simp del] | 
| 330 | ||
| 31847 | 331 | lemma filter_simp [simp]: | 
| 332 |   "filter P A = Fset {x \<in> member A. P x}"
 | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 333 | by (simp add: More_Set.project_def) | 
| 31847 | 334 | declare filter_def [simp del] | 
| 31846 | 335 | |
| 336 | declare mem_def [simp del] | |
| 337 | ||
| 31849 | 338 | |
| 37468 | 339 | hide_const (open) setify is_empty insert remove map filter forall exists card | 
| 34048 | 340 | Inter Union | 
| 31849 | 341 | |
| 31807 | 342 | end |