| author | chaieb | 
| Fri, 27 Mar 2009 14:44:18 +0000 | |
| changeset 30747 | b8ca7e450de3 | 
| parent 30663 | 0b6aff7451b2 | 
| child 31193 | f8d4ac84334f | 
| permissions | -rw-r--r-- | 
| 26348 | 1 | (* Title: HOL/Library/Enum.thy | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Finite types as explicit enumerations *}
 | |
| 6 | ||
| 7 | theory Enum | |
| 30663 
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
 haftmann parents: 
29961diff
changeset | 8 | imports Map Main | 
| 26348 | 9 | begin | 
| 10 | ||
| 11 | subsection {* Class @{text enum} *}
 | |
| 12 | ||
| 29797 | 13 | class enum = | 
| 26348 | 14 | fixes enum :: "'a list" | 
| 28562 | 15 | assumes UNIV_enum [code]: "UNIV = set enum" | 
| 26444 | 16 | and enum_distinct: "distinct enum" | 
| 26348 | 17 | begin | 
| 18 | ||
| 29797 | 19 | subclass finite proof | 
| 20 | qed (simp add: UNIV_enum) | |
| 26444 | 21 | |
| 22 | lemma enum_all: "set enum = UNIV" unfolding UNIV_enum .. | |
| 23 | ||
| 26348 | 24 | lemma in_enum [intro]: "x \<in> set enum" | 
| 25 | unfolding enum_all by auto | |
| 26 | ||
| 27 | lemma enum_eq_I: | |
| 28 | assumes "\<And>x. x \<in> set xs" | |
| 29 | shows "set enum = set xs" | |
| 30 | proof - | |
| 31 | from assms UNIV_eq_I have "UNIV = set xs" by auto | |
| 32 | with enum_all show ?thesis by simp | |
| 33 | qed | |
| 34 | ||
| 35 | end | |
| 36 | ||
| 37 | ||
| 38 | subsection {* Equality and order on functions *}
 | |
| 39 | ||
| 26513 | 40 | instantiation "fun" :: (enum, eq) eq | 
| 41 | begin | |
| 26348 | 42 | |
| 26513 | 43 | definition | 
| 26732 | 44 | "eq_class.eq f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)" | 
| 26513 | 45 | |
| 46 | instance by default | |
| 47 | (simp_all add: eq_fun_def enum_all expand_fun_eq) | |
| 48 | ||
| 49 | end | |
| 26348 | 50 | |
| 28562 | 51 | lemma order_fun [code]: | 
| 26348 | 52 | fixes f g :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>order" | 
| 26968 | 53 | shows "f \<le> g \<longleftrightarrow> list_all (\<lambda>x. f x \<le> g x) enum" | 
| 54 | and "f < g \<longleftrightarrow> f \<le> g \<and> \<not> list_all (\<lambda>x. f x = g x) enum" | |
| 28684 | 55 | by (simp_all add: list_all_iff enum_all expand_fun_eq le_fun_def order_less_le) | 
| 26968 | 56 | |
| 57 | ||
| 58 | subsection {* Quantifiers *}
 | |
| 59 | ||
| 28562 | 60 | lemma all_code [code]: "(\<forall>x. P x) \<longleftrightarrow> list_all P enum" | 
| 26968 | 61 | by (simp add: list_all_iff enum_all) | 
| 62 | ||
| 28562 | 63 | lemma exists_code [code]: "(\<exists>x. P x) \<longleftrightarrow> \<not> list_all (Not o P) enum" | 
| 26968 | 64 | by (simp add: list_all_iff enum_all) | 
| 26348 | 65 | |
| 66 | ||
| 67 | subsection {* Default instances *}
 | |
| 68 | ||
| 26444 | 69 | primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where | 
| 70 | "n_lists 0 xs = [[]]" | |
| 71 | | "n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))" | |
| 72 | ||
| 73 | lemma n_lists_Nil [simp]: "n_lists n [] = (if n = 0 then [[]] else [])" | |
| 74 | by (induct n) simp_all | |
| 75 | ||
| 76 | lemma length_n_lists: "length (n_lists n xs) = length xs ^ n" | |
| 77 | by (induct n) (auto simp add: length_concat map_compose [symmetric] o_def listsum_triv) | |
| 78 | ||
| 79 | lemma length_n_lists_elem: "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n" | |
| 80 | by (induct n arbitrary: ys) auto | |
| 81 | ||
| 82 | lemma set_n_lists: "set (n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
 | |
| 83 | proof (rule set_ext) | |
| 84 | fix ys :: "'a list" | |
| 85 |   show "ys \<in> set (n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
 | |
| 86 | proof - | |
| 87 | have "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n" | |
| 88 | by (induct n arbitrary: ys) auto | |
| 89 | moreover have "\<And>x. ys \<in> set (n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs" | |
| 90 | by (induct n arbitrary: ys) auto | |
| 91 | moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> set (n_lists (length ys) xs)" | |
| 92 | by (induct ys) auto | |
| 93 | ultimately show ?thesis by auto | |
| 94 | qed | |
| 95 | qed | |
| 96 | ||
| 97 | lemma distinct_n_lists: | |
| 98 | assumes "distinct xs" | |
| 99 | shows "distinct (n_lists n xs)" | |
| 100 | proof (rule card_distinct) | |
| 101 | from assms have card_length: "card (set xs) = length xs" by (rule distinct_card) | |
| 102 | have "card (set (n_lists n xs)) = card (set xs) ^ n" | |
| 103 | proof (induct n) | |
| 104 | case 0 then show ?case by simp | |
| 105 | next | |
| 106 | case (Suc n) | |
| 107 | moreover have "card (\<Union>ys\<in>set (n_lists n xs). (\<lambda>y. y # ys) ` set xs) | |
| 108 | = (\<Sum>ys\<in>set (n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))" | |
| 109 | by (rule card_UN_disjoint) auto | |
| 110 | moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)" | |
| 111 | by (rule card_image) (simp add: inj_on_def) | |
| 112 | ultimately show ?case by auto | |
| 113 | qed | |
| 114 | also have "\<dots> = length xs ^ n" by (simp add: card_length) | |
| 115 | finally show "card (set (n_lists n xs)) = length (n_lists n xs)" | |
| 116 | by (simp add: length_n_lists) | |
| 117 | qed | |
| 118 | ||
| 119 | lemma map_of_zip_map: | |
| 120 | fixes f :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>enum" | |
| 121 | shows "map_of (zip xs (map f xs)) = (\<lambda>x. if x \<in> set xs then Some (f x) else None)" | |
| 122 | by (induct xs) (simp_all add: expand_fun_eq) | |
| 123 | ||
| 124 | lemma map_of_zip_enum_is_Some: | |
| 125 | assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 126 | shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y" | |
| 127 | proof - | |
| 128 | from assms have "x \<in> set (enum \<Colon> 'a\<Colon>enum list) \<longleftrightarrow> | |
| 129 | (\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y)" | |
| 130 | by (auto intro!: map_of_zip_is_Some) | |
| 131 | then show ?thesis using enum_all by auto | |
| 132 | qed | |
| 133 | ||
| 134 | lemma map_of_zip_enum_inject: | |
| 135 | fixes xs ys :: "'b\<Colon>enum list" | |
| 136 | assumes length: "length xs = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 137 | "length ys = length (enum \<Colon> 'a\<Colon>enum list)" | |
| 138 | and map_of: "the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys)" | |
| 139 | shows "xs = ys" | |
| 140 | proof - | |
| 141 | have "map_of (zip (enum \<Colon> 'a list) xs) = map_of (zip (enum \<Colon> 'a list) ys)" | |
| 142 | proof | |
| 143 | fix x :: 'a | |
| 144 | from length map_of_zip_enum_is_Some obtain y1 y2 | |
| 145 | where "map_of (zip (enum \<Colon> 'a list) xs) x = Some y1" | |
| 146 | and "map_of (zip (enum \<Colon> 'a list) ys) x = Some y2" by blast | |
| 147 | moreover from map_of have "the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x) = the (map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x)" | |
| 148 | by (auto dest: fun_cong) | |
| 149 | ultimately show "map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x = map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x" | |
| 150 | by simp | |
| 151 | qed | |
| 152 | with length enum_distinct show "xs = ys" by (rule map_of_zip_inject) | |
| 153 | qed | |
| 154 | ||
| 155 | instantiation "fun" :: (enum, enum) enum | |
| 156 | begin | |
| 157 | ||
| 158 | definition | |
| 28562 | 159 | [code del]: "enum = map (\<lambda>ys. the o map_of (zip (enum\<Colon>'a list) ys)) (n_lists (length (enum\<Colon>'a\<Colon>enum list)) enum)" | 
| 26444 | 160 | |
| 161 | instance proof | |
| 162 |   show "UNIV = set (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | |
| 163 | proof (rule UNIV_eq_I) | |
| 164 | fix f :: "'a \<Rightarrow> 'b" | |
| 165 | have "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))" | |
| 166 | by (auto simp add: map_of_zip_map expand_fun_eq) | |
| 167 | then show "f \<in> set enum" | |
| 168 | by (auto simp add: enum_fun_def set_n_lists) | |
| 169 | qed | |
| 170 | next | |
| 171 | from map_of_zip_enum_inject | |
| 172 |   show "distinct (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | |
| 173 | by (auto intro!: inj_onI simp add: enum_fun_def | |
| 174 | distinct_map distinct_n_lists enum_distinct set_n_lists enum_all) | |
| 175 | qed | |
| 176 | ||
| 177 | end | |
| 178 | ||
| 28562 | 179 | lemma enum_fun_code [code]: "enum = (let enum_a = (enum \<Colon> 'a\<Colon>{enum, eq} list)
 | 
| 28245 
9767dd8e1e54
celver code lemma avoid type ambiguity problem with Haskell
 haftmann parents: 
27487diff
changeset | 180 | in map (\<lambda>ys. the o map_of (zip enum_a ys)) (n_lists (length enum_a) enum))" | 
| 
9767dd8e1e54
celver code lemma avoid type ambiguity problem with Haskell
 haftmann parents: 
27487diff
changeset | 181 | by (simp add: enum_fun_def Let_def) | 
| 26444 | 182 | |
| 26348 | 183 | instantiation unit :: enum | 
| 184 | begin | |
| 185 | ||
| 186 | definition | |
| 187 | "enum = [()]" | |
| 188 | ||
| 189 | instance by default | |
| 26444 | 190 | (simp_all add: enum_unit_def UNIV_unit) | 
| 26348 | 191 | |
| 192 | end | |
| 193 | ||
| 194 | instantiation bool :: enum | |
| 195 | begin | |
| 196 | ||
| 197 | definition | |
| 198 | "enum = [False, True]" | |
| 199 | ||
| 200 | instance by default | |
| 26444 | 201 | (simp_all add: enum_bool_def UNIV_bool) | 
| 26348 | 202 | |
| 203 | end | |
| 204 | ||
| 205 | primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
 | |
| 206 | "product [] _ = []" | |
| 207 | | "product (x#xs) ys = map (Pair x) ys @ product xs ys" | |
| 208 | ||
| 209 | lemma product_list_set: | |
| 210 | "set (product xs ys) = set xs \<times> set ys" | |
| 211 | by (induct xs) auto | |
| 212 | ||
| 26444 | 213 | lemma distinct_product: | 
| 214 | assumes "distinct xs" and "distinct ys" | |
| 215 | shows "distinct (product xs ys)" | |
| 216 | using assms by (induct xs) | |
| 217 | (auto intro: inj_onI simp add: product_list_set distinct_map) | |
| 218 | ||
| 26348 | 219 | instantiation * :: (enum, enum) enum | 
| 220 | begin | |
| 221 | ||
| 222 | definition | |
| 223 | "enum = product enum enum" | |
| 224 | ||
| 225 | instance by default | |
| 26444 | 226 | (simp_all add: enum_prod_def product_list_set distinct_product enum_all enum_distinct) | 
| 26348 | 227 | |
| 228 | end | |
| 229 | ||
| 230 | instantiation "+" :: (enum, enum) enum | |
| 231 | begin | |
| 232 | ||
| 233 | definition | |
| 234 | "enum = map Inl enum @ map Inr enum" | |
| 235 | ||
| 236 | instance by default | |
| 26444 | 237 | (auto simp add: enum_all enum_sum_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct) | 
| 26348 | 238 | |
| 239 | end | |
| 240 | ||
| 241 | primrec sublists :: "'a list \<Rightarrow> 'a list list" where | |
| 242 | "sublists [] = [[]]" | |
| 243 | | "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)" | |
| 244 | ||
| 26444 | 245 | lemma length_sublists: | 
| 246 | "length (sublists xs) = Suc (Suc (0\<Colon>nat)) ^ length xs" | |
| 247 | by (induct xs) (simp_all add: Let_def) | |
| 248 | ||
| 26348 | 249 | lemma sublists_powset: | 
| 26444 | 250 | "set ` set (sublists xs) = Pow (set xs)" | 
| 26348 | 251 | proof - | 
| 252 | have aux: "\<And>x A. set ` Cons x ` A = insert x ` set ` A" | |
| 253 | by (auto simp add: image_def) | |
| 26444 | 254 | have "set (map set (sublists xs)) = Pow (set xs)" | 
| 26348 | 255 | by (induct xs) | 
| 26444 | 256 | (simp_all add: aux Let_def Pow_insert Un_commute) | 
| 257 | then show ?thesis by simp | |
| 258 | qed | |
| 259 | ||
| 260 | lemma distinct_set_sublists: | |
| 261 | assumes "distinct xs" | |
| 262 | shows "distinct (map set (sublists xs))" | |
| 263 | proof (rule card_distinct) | |
| 264 | have "finite (set xs)" by rule | |
| 265 | then have "card (Pow (set xs)) = Suc (Suc 0) ^ card (set xs)" by (rule card_Pow) | |
| 266 | with assms distinct_card [of xs] | |
| 267 | have "card (Pow (set xs)) = Suc (Suc 0) ^ length xs" by simp | |
| 268 | then show "card (set (map set (sublists xs))) = length (map set (sublists xs))" | |
| 269 | by (simp add: sublists_powset length_sublists) | |
| 26348 | 270 | qed | 
| 271 | ||
| 272 | instantiation nibble :: enum | |
| 273 | begin | |
| 274 | ||
| 275 | definition | |
| 276 | "enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7, | |
| 277 | Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]" | |
| 278 | ||
| 279 | instance by default | |
| 26444 | 280 | (simp_all add: enum_nibble_def UNIV_nibble) | 
| 26348 | 281 | |
| 282 | end | |
| 283 | ||
| 284 | instantiation char :: enum | |
| 285 | begin | |
| 286 | ||
| 287 | definition | |
| 28562 | 288 | [code del]: "enum = map (split Char) (product enum enum)" | 
| 26444 | 289 | |
| 28562 | 290 | lemma enum_char [code]: | 
| 26444 | 291 | "enum = [Char Nibble0 Nibble0, Char Nibble0 Nibble1, Char Nibble0 Nibble2, | 
| 292 | Char Nibble0 Nibble3, Char Nibble0 Nibble4, Char Nibble0 Nibble5, | |
| 293 | Char Nibble0 Nibble6, Char Nibble0 Nibble7, Char Nibble0 Nibble8, | |
| 294 | Char Nibble0 Nibble9, Char Nibble0 NibbleA, Char Nibble0 NibbleB, | |
| 295 | Char Nibble0 NibbleC, Char Nibble0 NibbleD, Char Nibble0 NibbleE, | |
| 296 | Char Nibble0 NibbleF, Char Nibble1 Nibble0, Char Nibble1 Nibble1, | |
| 297 | Char Nibble1 Nibble2, Char Nibble1 Nibble3, Char Nibble1 Nibble4, | |
| 298 | Char Nibble1 Nibble5, Char Nibble1 Nibble6, Char Nibble1 Nibble7, | |
| 299 | Char Nibble1 Nibble8, Char Nibble1 Nibble9, Char Nibble1 NibbleA, | |
| 300 | Char Nibble1 NibbleB, Char Nibble1 NibbleC, Char Nibble1 NibbleD, | |
| 301 | Char Nibble1 NibbleE, Char Nibble1 NibbleF, CHR '' '', CHR ''!'', | |
| 302 | Char Nibble2 Nibble2, CHR ''#'', CHR ''$'', CHR ''%'', CHR ''&'', | |
| 303 |   Char Nibble2 Nibble7, CHR ''('', CHR '')'', CHR ''*'', CHR ''+'', CHR '','',
 | |
| 304 | CHR ''-'', CHR ''.'', CHR ''/'', CHR ''0'', CHR ''1'', CHR ''2'', CHR ''3'', | |
| 305 | CHR ''4'', CHR ''5'', CHR ''6'', CHR ''7'', CHR ''8'', CHR ''9'', CHR '':'', | |
| 306 | CHR '';'', CHR ''<'', CHR ''='', CHR ''>'', CHR ''?'', CHR ''@'', CHR ''A'', | |
| 307 | CHR ''B'', CHR ''C'', CHR ''D'', CHR ''E'', CHR ''F'', CHR ''G'', CHR ''H'', | |
| 308 | CHR ''I'', CHR ''J'', CHR ''K'', CHR ''L'', CHR ''M'', CHR ''N'', CHR ''O'', | |
| 309 | CHR ''P'', CHR ''Q'', CHR ''R'', CHR ''S'', CHR ''T'', CHR ''U'', CHR ''V'', | |
| 310 | CHR ''W'', CHR ''X'', CHR ''Y'', CHR ''Z'', CHR ''['', Char Nibble5 NibbleC, | |
| 311 | CHR '']'', CHR ''^'', CHR ''_'', Char Nibble6 Nibble0, CHR ''a'', CHR ''b'', | |
| 312 | CHR ''c'', CHR ''d'', CHR ''e'', CHR ''f'', CHR ''g'', CHR ''h'', CHR ''i'', | |
| 313 | CHR ''j'', CHR ''k'', CHR ''l'', CHR ''m'', CHR ''n'', CHR ''o'', CHR ''p'', | |
| 314 | CHR ''q'', CHR ''r'', CHR ''s'', CHR ''t'', CHR ''u'', CHR ''v'', CHR ''w'', | |
| 315 |   CHR ''x'', CHR ''y'', CHR ''z'', CHR ''{'', CHR ''|'', CHR ''}'', CHR ''~'',
 | |
| 316 | Char Nibble7 NibbleF, Char Nibble8 Nibble0, Char Nibble8 Nibble1, | |
| 317 | Char Nibble8 Nibble2, Char Nibble8 Nibble3, Char Nibble8 Nibble4, | |
| 318 | Char Nibble8 Nibble5, Char Nibble8 Nibble6, Char Nibble8 Nibble7, | |
| 319 | Char Nibble8 Nibble8, Char Nibble8 Nibble9, Char Nibble8 NibbleA, | |
| 320 | Char Nibble8 NibbleB, Char Nibble8 NibbleC, Char Nibble8 NibbleD, | |
| 321 | Char Nibble8 NibbleE, Char Nibble8 NibbleF, Char Nibble9 Nibble0, | |
| 322 | Char Nibble9 Nibble1, Char Nibble9 Nibble2, Char Nibble9 Nibble3, | |
| 323 | Char Nibble9 Nibble4, Char Nibble9 Nibble5, Char Nibble9 Nibble6, | |
| 324 | Char Nibble9 Nibble7, Char Nibble9 Nibble8, Char Nibble9 Nibble9, | |
| 325 | Char Nibble9 NibbleA, Char Nibble9 NibbleB, Char Nibble9 NibbleC, | |
| 326 | Char Nibble9 NibbleD, Char Nibble9 NibbleE, Char Nibble9 NibbleF, | |
| 327 | Char NibbleA Nibble0, Char NibbleA Nibble1, Char NibbleA Nibble2, | |
| 328 | Char NibbleA Nibble3, Char NibbleA Nibble4, Char NibbleA Nibble5, | |
| 329 | Char NibbleA Nibble6, Char NibbleA Nibble7, Char NibbleA Nibble8, | |
| 330 | Char NibbleA Nibble9, Char NibbleA NibbleA, Char NibbleA NibbleB, | |
| 331 | Char NibbleA NibbleC, Char NibbleA NibbleD, Char NibbleA NibbleE, | |
| 332 | Char NibbleA NibbleF, Char NibbleB Nibble0, Char NibbleB Nibble1, | |
| 333 | Char NibbleB Nibble2, Char NibbleB Nibble3, Char NibbleB Nibble4, | |
| 334 | Char NibbleB Nibble5, Char NibbleB Nibble6, Char NibbleB Nibble7, | |
| 335 | Char NibbleB Nibble8, Char NibbleB Nibble9, Char NibbleB NibbleA, | |
| 336 | Char NibbleB NibbleB, Char NibbleB NibbleC, Char NibbleB NibbleD, | |
| 337 | Char NibbleB NibbleE, Char NibbleB NibbleF, Char NibbleC Nibble0, | |
| 338 | Char NibbleC Nibble1, Char NibbleC Nibble2, Char NibbleC Nibble3, | |
| 339 | Char NibbleC Nibble4, Char NibbleC Nibble5, Char NibbleC Nibble6, | |
| 340 | Char NibbleC Nibble7, Char NibbleC Nibble8, Char NibbleC Nibble9, | |
| 341 | Char NibbleC NibbleA, Char NibbleC NibbleB, Char NibbleC NibbleC, | |
| 342 | Char NibbleC NibbleD, Char NibbleC NibbleE, Char NibbleC NibbleF, | |
| 343 | Char NibbleD Nibble0, Char NibbleD Nibble1, Char NibbleD Nibble2, | |
| 344 | Char NibbleD Nibble3, Char NibbleD Nibble4, Char NibbleD Nibble5, | |
| 345 | Char NibbleD Nibble6, Char NibbleD Nibble7, Char NibbleD Nibble8, | |
| 346 | Char NibbleD Nibble9, Char NibbleD NibbleA, Char NibbleD NibbleB, | |
| 347 | Char NibbleD NibbleC, Char NibbleD NibbleD, Char NibbleD NibbleE, | |
| 348 | Char NibbleD NibbleF, Char NibbleE Nibble0, Char NibbleE Nibble1, | |
| 349 | Char NibbleE Nibble2, Char NibbleE Nibble3, Char NibbleE Nibble4, | |
| 350 | Char NibbleE Nibble5, Char NibbleE Nibble6, Char NibbleE Nibble7, | |
| 351 | Char NibbleE Nibble8, Char NibbleE Nibble9, Char NibbleE NibbleA, | |
| 352 | Char NibbleE NibbleB, Char NibbleE NibbleC, Char NibbleE NibbleD, | |
| 353 | Char NibbleE NibbleE, Char NibbleE NibbleF, Char NibbleF Nibble0, | |
| 354 | Char NibbleF Nibble1, Char NibbleF Nibble2, Char NibbleF Nibble3, | |
| 355 | Char NibbleF Nibble4, Char NibbleF Nibble5, Char NibbleF Nibble6, | |
| 356 | Char NibbleF Nibble7, Char NibbleF Nibble8, Char NibbleF Nibble9, | |
| 357 | Char NibbleF NibbleA, Char NibbleF NibbleB, Char NibbleF NibbleC, | |
| 358 | Char NibbleF NibbleD, Char NibbleF NibbleE, Char NibbleF NibbleF]" | |
| 359 | unfolding enum_char_def enum_nibble_def by simp | |
| 26348 | 360 | |
| 361 | instance by default | |
| 26444 | 362 | (auto intro: char.exhaust injI simp add: enum_char_def product_list_set enum_all full_SetCompr_eq [symmetric] | 
| 363 | distinct_map distinct_product enum_distinct) | |
| 26348 | 364 | |
| 365 | end | |
| 366 | ||
| 29024 | 367 | instantiation option :: (enum) enum | 
| 368 | begin | |
| 369 | ||
| 370 | definition | |
| 371 | "enum = None # map Some enum" | |
| 372 | ||
| 373 | instance by default | |
| 374 | (auto simp add: enum_all enum_option_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct) | |
| 375 | ||
| 376 | end | |
| 377 | ||
| 378 | end |