| author | blanchet | 
| Mon, 22 Aug 2011 15:02:45 +0200 | |
| changeset 44400 | 01b8b6fcd857 | 
| parent 43363 | eaf8b7f22d39 | 
| child 44860 | 56101fa00193 | 
| permissions | -rw-r--r-- | 
| 34020 | 1 | (* Title: HOL/Library/Executable_Set.thy | 
| 2 | Author: Stefan Berghofer, TU Muenchen | |
| 33947 | 3 | Author: Florian Haftmann, TU Muenchen | 
| 4 | *) | |
| 5 | ||
| 6 | header {* A crude implementation of finite sets by lists -- avoid using this at any cost! *}
 | |
| 7 | ||
| 34020 | 8 | theory Executable_Set | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 9 | imports More_Set | 
| 33947 | 10 | begin | 
| 11 | ||
| 39141 | 12 | text {*
 | 
| 13 | This is just an ad-hoc hack which will rarely give you what you want. | |
| 14 | For the moment, whenever you need executable sets, consider using | |
| 43363 | 15 |   type @{text Cset.set} from theory @{text Cset}.
 | 
| 39141 | 16 | *} | 
| 17 | ||
| 33947 | 18 | declare mem_def [code del] | 
| 19 | declare Collect_def [code del] | |
| 20 | declare insert_code [code del] | |
| 21 | declare vimage_code [code del] | |
| 22 | ||
| 23 | subsection {* Set representation *}
 | |
| 24 | ||
| 25 | setup {*
 | |
| 26 | Code.add_type_cmd "set" | |
| 27 | *} | |
| 28 | ||
| 29 | definition Set :: "'a list \<Rightarrow> 'a set" where | |
| 30 | [simp]: "Set = set" | |
| 31 | ||
| 32 | definition Coset :: "'a list \<Rightarrow> 'a set" where | |
| 33 | [simp]: "Coset xs = - set xs" | |
| 34 | ||
| 35 | setup {*
 | |
| 36 |   Code.add_signature_cmd ("Set", "'a list \<Rightarrow> 'a set")
 | |
| 37 |   #> Code.add_signature_cmd ("Coset", "'a list \<Rightarrow> 'a set")
 | |
| 38 |   #> Code.add_signature_cmd ("set", "'a list \<Rightarrow> 'a set")
 | |
| 39 |   #> Code.add_signature_cmd ("op \<in>", "'a \<Rightarrow> 'a set \<Rightarrow> bool")
 | |
| 40 | *} | |
| 41 | ||
| 42 | code_datatype Set Coset | |
| 43 | ||
| 34020 | 44 | consts_code | 
| 45 |   Coset ("\<module>Coset")
 | |
| 46 |   Set ("\<module>Set")
 | |
| 47 | attach {*
 | |
| 48 | datatype 'a set = Set of 'a list | Coset of 'a list; | |
| 49 | *} -- {* This assumes that there won't be a @{text Coset} without a @{text Set} *}
 | |
| 50 | ||
| 33947 | 51 | |
| 52 | subsection {* Basic operations *}
 | |
| 53 | ||
| 54 | lemma [code]: | |
| 55 | "set xs = Set (remdups xs)" | |
| 56 | by simp | |
| 57 | ||
| 58 | lemma [code]: | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 59 | "x \<in> Set xs \<longleftrightarrow> List.member xs x" | 
| 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 60 | "x \<in> Coset xs \<longleftrightarrow> \<not> List.member xs x" | 
| 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 61 | by (simp_all add: member_def) | 
| 33947 | 62 | |
| 63 | definition is_empty :: "'a set \<Rightarrow> bool" where | |
| 64 |   [simp]: "is_empty A \<longleftrightarrow> A = {}"
 | |
| 65 | ||
| 34020 | 66 | lemma [code_unfold]: | 
| 33947 | 67 |   "A = {} \<longleftrightarrow> is_empty A"
 | 
| 68 | by simp | |
| 69 | ||
| 70 | definition empty :: "'a set" where | |
| 71 |   [simp]: "empty = {}"
 | |
| 72 | ||
| 34020 | 73 | lemma [code_unfold]: | 
| 33947 | 74 |   "{} = empty"
 | 
| 75 | by simp | |
| 76 | ||
| 34020 | 77 | lemma [code_unfold, code_inline del]: | 
| 78 | "empty = Set []" | |
| 79 |   by simp -- {* Otherwise @{text \<eta>}-expansion produces funny things. *}
 | |
| 80 | ||
| 33947 | 81 | setup {*
 | 
| 82 |   Code.add_signature_cmd ("is_empty", "'a set \<Rightarrow> bool")
 | |
| 83 |   #> Code.add_signature_cmd ("empty", "'a set")
 | |
| 84 |   #> Code.add_signature_cmd ("insert", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
 | |
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 85 |   #> Code.add_signature_cmd ("More_Set.remove", "'a \<Rightarrow> 'a set \<Rightarrow> 'a set")
 | 
| 33947 | 86 |   #> Code.add_signature_cmd ("image", "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b set")
 | 
| 37024 
e938a0b5286e
renamed List_Set to the now more appropriate More_Set
 haftmann parents: 
37023diff
changeset | 87 |   #> Code.add_signature_cmd ("More_Set.project", "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set")
 | 
| 33947 | 88 |   #> Code.add_signature_cmd ("Ball", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
 | 
| 89 |   #> Code.add_signature_cmd ("Bex", "'a set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool")
 | |
| 90 |   #> Code.add_signature_cmd ("card", "'a set \<Rightarrow> nat")
 | |
| 91 | *} | |
| 92 | ||
| 93 | lemma is_empty_Set [code]: | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 94 | "is_empty (Set xs) \<longleftrightarrow> List.null xs" | 
| 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 95 | by (simp add: null_def) | 
| 33947 | 96 | |
| 97 | lemma empty_Set [code]: | |
| 98 | "empty = Set []" | |
| 99 | by simp | |
| 100 | ||
| 101 | lemma insert_Set [code]: | |
| 34980 | 102 | "insert x (Set xs) = Set (List.insert x xs)" | 
| 103 | "insert x (Coset xs) = Coset (removeAll x xs)" | |
| 104 | by (simp_all add: set_insert) | |
| 33947 | 105 | |
| 106 | lemma remove_Set [code]: | |
| 34980 | 107 | "remove x (Set xs) = Set (removeAll x xs)" | 
| 108 | "remove x (Coset xs) = Coset (List.insert x xs)" | |
| 109 | by (auto simp add: set_insert remove_def) | |
| 33947 | 110 | |
| 111 | lemma image_Set [code]: | |
| 112 | "image f (Set xs) = Set (remdups (map f xs))" | |
| 113 | by simp | |
| 114 | ||
| 115 | lemma project_Set [code]: | |
| 116 | "project P (Set xs) = Set (filter P xs)" | |
| 117 | by (simp add: project_set) | |
| 118 | ||
| 119 | lemma Ball_Set [code]: | |
| 120 | "Ball (Set xs) P \<longleftrightarrow> list_all P xs" | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 121 | by (simp add: list_all_iff) | 
| 33947 | 122 | |
| 123 | lemma Bex_Set [code]: | |
| 124 | "Bex (Set xs) P \<longleftrightarrow> list_ex P xs" | |
| 37595 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 haftmann parents: 
37024diff
changeset | 125 | by (simp add: list_ex_iff) | 
| 33947 | 126 | |
| 127 | lemma card_Set [code]: | |
| 128 | "card (Set xs) = length (remdups xs)" | |
| 129 | proof - | |
| 130 | have "card (set (remdups xs)) = length (remdups xs)" | |
| 131 | by (rule distinct_card) simp | |
| 132 | then show ?thesis by simp | |
| 133 | qed | |
| 134 | ||
| 135 | ||
| 136 | subsection {* Derived operations *}
 | |
| 137 | ||
| 138 | definition set_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where | |
| 139 | [simp]: "set_eq = op =" | |
| 140 | ||
| 34020 | 141 | lemma [code_unfold]: | 
| 33947 | 142 | "op = = set_eq" | 
| 143 | by simp | |
| 144 | ||
| 145 | definition subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where | |
| 146 | [simp]: "subset_eq = op \<subseteq>" | |
| 147 | ||
| 34020 | 148 | lemma [code_unfold]: | 
| 33947 | 149 | "op \<subseteq> = subset_eq" | 
| 150 | by simp | |
| 151 | ||
| 152 | definition subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where | |
| 153 | [simp]: "subset = op \<subset>" | |
| 154 | ||
| 34020 | 155 | lemma [code_unfold]: | 
| 33947 | 156 | "op \<subset> = subset" | 
| 157 | by simp | |
| 158 | ||
| 159 | setup {*
 | |
| 160 |   Code.add_signature_cmd ("set_eq", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
 | |
| 161 |   #> Code.add_signature_cmd ("subset_eq", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
 | |
| 162 |   #> Code.add_signature_cmd ("subset", "'a set \<Rightarrow> 'a set \<Rightarrow> bool")
 | |
| 163 | *} | |
| 164 | ||
| 165 | lemma set_eq_subset_eq [code]: | |
| 166 | "set_eq A B \<longleftrightarrow> subset_eq A B \<and> subset_eq B A" | |
| 167 | by auto | |
| 168 | ||
| 169 | lemma subset_eq_forall [code]: | |
| 170 | "subset_eq A B \<longleftrightarrow> (\<forall>x\<in>A. x \<in> B)" | |
| 171 | by (simp add: subset_eq) | |
| 172 | ||
| 173 | lemma subset_subset_eq [code]: | |
| 174 | "subset A B \<longleftrightarrow> subset_eq A B \<and> \<not> subset_eq B A" | |
| 175 | by (simp add: subset) | |
| 176 | ||
| 177 | ||
| 178 | subsection {* Functorial operations *}
 | |
| 179 | ||
| 180 | definition inter :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where | |
| 181 | [simp]: "inter = op \<inter>" | |
| 182 | ||
| 34020 | 183 | lemma [code_unfold]: | 
| 33947 | 184 | "op \<inter> = inter" | 
| 185 | by simp | |
| 186 | ||
| 187 | definition subtract :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where | |
| 188 | [simp]: "subtract A B = B - A" | |
| 189 | ||
| 34020 | 190 | lemma [code_unfold]: | 
| 33947 | 191 | "B - A = subtract A B" | 
| 192 | by simp | |
| 193 | ||
| 194 | definition union :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where | |
| 195 | [simp]: "union = op \<union>" | |
| 196 | ||
| 34020 | 197 | lemma [code_unfold]: | 
| 33947 | 198 | "op \<union> = union" | 
| 199 | by simp | |
| 200 | ||
| 201 | definition Inf :: "'a::complete_lattice set \<Rightarrow> 'a" where | |
| 202 | [simp]: "Inf = Complete_Lattice.Inf" | |
| 203 | ||
| 34020 | 204 | lemma [code_unfold]: | 
| 33947 | 205 | "Complete_Lattice.Inf = Inf" | 
| 206 | by simp | |
| 207 | ||
| 208 | definition Sup :: "'a::complete_lattice set \<Rightarrow> 'a" where | |
| 209 | [simp]: "Sup = Complete_Lattice.Sup" | |
| 210 | ||
| 34020 | 211 | lemma [code_unfold]: | 
| 33947 | 212 | "Complete_Lattice.Sup = Sup" | 
| 213 | by simp | |
| 214 | ||
| 215 | definition Inter :: "'a set set \<Rightarrow> 'a set" where | |
| 216 | [simp]: "Inter = Inf" | |
| 217 | ||
| 34020 | 218 | lemma [code_unfold]: | 
| 33947 | 219 | "Inf = Inter" | 
| 220 | by simp | |
| 221 | ||
| 222 | definition Union :: "'a set set \<Rightarrow> 'a set" where | |
| 223 | [simp]: "Union = Sup" | |
| 224 | ||
| 34020 | 225 | lemma [code_unfold]: | 
| 33947 | 226 | "Sup = Union" | 
| 227 | by simp | |
| 228 | ||
| 229 | setup {*
 | |
| 230 |   Code.add_signature_cmd ("inter", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
 | |
| 231 |   #> Code.add_signature_cmd ("subtract", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
 | |
| 232 |   #> Code.add_signature_cmd ("union", "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set")
 | |
| 233 |   #> Code.add_signature_cmd ("Inf", "'a set \<Rightarrow> 'a")
 | |
| 234 |   #> Code.add_signature_cmd ("Sup", "'a set \<Rightarrow> 'a")
 | |
| 235 |   #> Code.add_signature_cmd ("Inter", "'a set set \<Rightarrow> 'a set")
 | |
| 236 |   #> Code.add_signature_cmd ("Union", "'a set set \<Rightarrow> 'a set")
 | |
| 237 | *} | |
| 238 | ||
| 239 | lemma inter_project [code]: | |
| 240 | "inter A (Set xs) = Set (List.filter (\<lambda>x. x \<in> A) xs)" | |
| 37023 | 241 | "inter A (Coset xs) = foldr remove xs A" | 
| 242 | by (simp add: inter project_def) (simp add: Diff_eq [symmetric] minus_set_foldr) | |
| 33947 | 243 | |
| 244 | lemma subtract_remove [code]: | |
| 37023 | 245 | "subtract (Set xs) A = foldr remove xs A" | 
| 33947 | 246 | "subtract (Coset xs) A = Set (List.filter (\<lambda>x. x \<in> A) xs)" | 
| 37023 | 247 | by (auto simp add: minus_set_foldr) | 
| 33947 | 248 | |
| 249 | lemma union_insert [code]: | |
| 37023 | 250 | "union (Set xs) A = foldr insert xs A" | 
| 33947 | 251 | "union (Coset xs) A = Coset (List.filter (\<lambda>x. x \<notin> A) xs)" | 
| 37023 | 252 | by (auto simp add: union_set_foldr) | 
| 33947 | 253 | |
| 254 | lemma Inf_inf [code]: | |
| 37023 | 255 | "Inf (Set xs) = foldr inf xs (top :: 'a::complete_lattice)" | 
| 33947 | 256 | "Inf (Coset []) = (bot :: 'a::complete_lattice)" | 
| 37023 | 257 | by (simp_all add: Inf_UNIV Inf_set_foldr) | 
| 33947 | 258 | |
| 259 | lemma Sup_sup [code]: | |
| 37023 | 260 | "Sup (Set xs) = foldr sup xs (bot :: 'a::complete_lattice)" | 
| 33947 | 261 | "Sup (Coset []) = (top :: 'a::complete_lattice)" | 
| 37023 | 262 | by (simp_all add: Sup_UNIV Sup_set_foldr) | 
| 33947 | 263 | |
| 264 | lemma Inter_inter [code]: | |
| 37023 | 265 | "Inter (Set xs) = foldr inter xs (Coset [])" | 
| 33947 | 266 | "Inter (Coset []) = empty" | 
| 267 | unfolding Inter_def Inf_inf by simp_all | |
| 268 | ||
| 269 | lemma Union_union [code]: | |
| 37023 | 270 | "Union (Set xs) = foldr union xs empty" | 
| 33947 | 271 | "Union (Coset []) = Coset []" | 
| 272 | unfolding Union_def Sup_sup by simp_all | |
| 273 | ||
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
34980diff
changeset | 274 | hide_const (open) is_empty empty remove | 
| 33947 | 275 | set_eq subset_eq subset inter union subtract Inf Sup Inter Union | 
| 276 | ||
| 38304 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 277 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 278 | subsection {* Operations on relations *}
 | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 279 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 280 | text {* Initially contributed by Tjark Weber. *}
 | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 281 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 282 | lemma bounded_Collect_code [code_unfold]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 283 |   "{x\<in>S. P x} = project P S"
 | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 284 | by (auto simp add: project_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 285 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 286 | lemma Id_on_code [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 287 | "Id_on (Set xs) = Set [(x,x). x \<leftarrow> xs]" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 288 | by (auto simp add: Id_on_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 289 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 290 | lemma Domain_fst [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 291 | "Domain r = fst ` r" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 292 | by (auto simp add: image_def Bex_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 293 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 294 | lemma Range_snd [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 295 | "Range r = snd ` r" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 296 | by (auto simp add: image_def Bex_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 297 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 298 | lemma irrefl_code [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 299 | "irrefl r \<longleftrightarrow> (\<forall>(x, y)\<in>r. x \<noteq> y)" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 300 | by (auto simp add: irrefl_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 301 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 302 | lemma trans_def [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 303 | "trans r \<longleftrightarrow> (\<forall>(x, y1)\<in>r. \<forall>(y2, z)\<in>r. y1 = y2 \<longrightarrow> (x, z)\<in>r)" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 304 | by (auto simp add: trans_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 305 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 306 | definition "exTimes A B = Sigma A (%_. B)" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 307 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 308 | lemma [code_unfold]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 309 | "Sigma A (%_. B) = exTimes A B" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 310 | by (simp add: exTimes_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 311 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 312 | lemma exTimes_code [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 313 | "exTimes (Set xs) (Set ys) = Set [(x,y). x \<leftarrow> xs, y \<leftarrow> ys]" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 314 | by (auto simp add: exTimes_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 315 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 316 | lemma rel_comp_code [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 317 | "(Set xys) O (Set yzs) = Set (remdups [(fst xy, snd yz). xy \<leftarrow> xys, yz \<leftarrow> yzs, snd xy = fst yz])" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 318 | by (auto simp add: Bex_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 319 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 320 | lemma acyclic_code [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 321 | "acyclic r = irrefl (r^+)" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 322 | by (simp add: acyclic_def irrefl_def) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 323 | |
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 324 | lemma wf_code [code]: | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 325 | "wf (Set xs) = acyclic (Set xs)" | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 326 | by (simp add: wf_iff_acyclic_if_finite) | 
| 
df7d5143db55
executable relation operations contributed by Tjark Weber
 haftmann parents: 
37595diff
changeset | 327 | |
| 33947 | 328 | end |