| author | haftmann | 
| Wed, 13 Jan 2010 08:56:25 +0100 | |
| changeset 34889 | dcaf6ec84e28 | 
| parent 33639 | 603320b93668 | 
| child 37601 | 2a4fb776ca53 | 
| permissions | -rw-r--r-- | 
| 31596 | 1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
| 26348 | 2  | 
|
3  | 
header {* Finite types as explicit enumerations *}
 | 
|
4  | 
||
5  | 
theory Enum  | 
|
| 
30663
 
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
 
haftmann 
parents: 
29961 
diff
changeset
 | 
6  | 
imports Map Main  | 
| 26348 | 7  | 
begin  | 
8  | 
||
9  | 
subsection {* Class @{text enum} *}
 | 
|
10  | 
||
| 29797 | 11  | 
class enum =  | 
| 26348 | 12  | 
fixes enum :: "'a list"  | 
| 33635 | 13  | 
assumes UNIV_enum: "UNIV = set enum"  | 
| 26444 | 14  | 
and enum_distinct: "distinct enum"  | 
| 26348 | 15  | 
begin  | 
16  | 
||
| 29797 | 17  | 
subclass finite proof  | 
18  | 
qed (simp add: UNIV_enum)  | 
|
| 26444 | 19  | 
|
20  | 
lemma enum_all: "set enum = UNIV" unfolding UNIV_enum ..  | 
|
21  | 
||
| 26348 | 22  | 
lemma in_enum [intro]: "x \<in> set enum"  | 
23  | 
unfolding enum_all by auto  | 
|
24  | 
||
25  | 
lemma enum_eq_I:  | 
|
26  | 
assumes "\<And>x. x \<in> set xs"  | 
|
27  | 
shows "set enum = set xs"  | 
|
28  | 
proof -  | 
|
29  | 
from assms UNIV_eq_I have "UNIV = set xs" by auto  | 
|
30  | 
with enum_all show ?thesis by simp  | 
|
31  | 
qed  | 
|
32  | 
||
33  | 
end  | 
|
34  | 
||
35  | 
||
36  | 
subsection {* Equality and order on functions *}
 | 
|
37  | 
||
| 26513 | 38  | 
instantiation "fun" :: (enum, eq) eq  | 
39  | 
begin  | 
|
| 26348 | 40  | 
|
| 26513 | 41  | 
definition  | 
| 26732 | 42  | 
"eq_class.eq f g \<longleftrightarrow> (\<forall>x \<in> set enum. f x = g x)"  | 
| 26513 | 43  | 
|
| 31464 | 44  | 
instance proof  | 
45  | 
qed (simp_all add: eq_fun_def enum_all expand_fun_eq)  | 
|
| 26513 | 46  | 
|
47  | 
end  | 
|
| 26348 | 48  | 
|
| 28562 | 49  | 
lemma order_fun [code]:  | 
| 26348 | 50  | 
fixes f g :: "'a\<Colon>enum \<Rightarrow> 'b\<Colon>order"  | 
| 26968 | 51  | 
shows "f \<le> g \<longleftrightarrow> list_all (\<lambda>x. f x \<le> g x) enum"  | 
52  | 
and "f < g \<longleftrightarrow> f \<le> g \<and> \<not> list_all (\<lambda>x. f x = g x) enum"  | 
|
| 28684 | 53  | 
by (simp_all add: list_all_iff enum_all expand_fun_eq le_fun_def order_less_le)  | 
| 26968 | 54  | 
|
55  | 
||
56  | 
subsection {* Quantifiers *}
 | 
|
57  | 
||
| 28562 | 58  | 
lemma all_code [code]: "(\<forall>x. P x) \<longleftrightarrow> list_all P enum"  | 
| 26968 | 59  | 
by (simp add: list_all_iff enum_all)  | 
60  | 
||
| 28562 | 61  | 
lemma exists_code [code]: "(\<exists>x. P x) \<longleftrightarrow> \<not> list_all (Not o P) enum"  | 
| 26968 | 62  | 
by (simp add: list_all_iff enum_all)  | 
| 26348 | 63  | 
|
64  | 
||
65  | 
subsection {* Default instances *}
 | 
|
66  | 
||
| 26444 | 67  | 
primrec n_lists :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list list" where  | 
68  | 
"n_lists 0 xs = [[]]"  | 
|
69  | 
| "n_lists (Suc n) xs = concat (map (\<lambda>ys. map (\<lambda>y. y # ys) xs) (n_lists n xs))"  | 
|
70  | 
||
71  | 
lemma n_lists_Nil [simp]: "n_lists n [] = (if n = 0 then [[]] else [])"  | 
|
72  | 
by (induct n) simp_all  | 
|
73  | 
||
74  | 
lemma length_n_lists: "length (n_lists n xs) = length xs ^ n"  | 
|
| 
33639
 
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
 
hoelzl 
parents: 
33635 
diff
changeset
 | 
75  | 
by (induct n) (auto simp add: length_concat o_def listsum_triv)  | 
| 26444 | 76  | 
|
77  | 
lemma length_n_lists_elem: "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n"  | 
|
78  | 
by (induct n arbitrary: ys) auto  | 
|
79  | 
||
80  | 
lemma set_n_lists: "set (n_lists n xs) = {ys. length ys = n \<and> set ys \<subseteq> set xs}"
 | 
|
81  | 
proof (rule set_ext)  | 
|
82  | 
fix ys :: "'a list"  | 
|
83  | 
  show "ys \<in> set (n_lists n xs) \<longleftrightarrow> ys \<in> {ys. length ys = n \<and> set ys \<subseteq> set xs}"
 | 
|
84  | 
proof -  | 
|
85  | 
have "ys \<in> set (n_lists n xs) \<Longrightarrow> length ys = n"  | 
|
86  | 
by (induct n arbitrary: ys) auto  | 
|
87  | 
moreover have "\<And>x. ys \<in> set (n_lists n xs) \<Longrightarrow> x \<in> set ys \<Longrightarrow> x \<in> set xs"  | 
|
88  | 
by (induct n arbitrary: ys) auto  | 
|
89  | 
moreover have "set ys \<subseteq> set xs \<Longrightarrow> ys \<in> set (n_lists (length ys) xs)"  | 
|
90  | 
by (induct ys) auto  | 
|
91  | 
ultimately show ?thesis by auto  | 
|
92  | 
qed  | 
|
93  | 
qed  | 
|
94  | 
||
95  | 
lemma distinct_n_lists:  | 
|
96  | 
assumes "distinct xs"  | 
|
97  | 
shows "distinct (n_lists n xs)"  | 
|
98  | 
proof (rule card_distinct)  | 
|
99  | 
from assms have card_length: "card (set xs) = length xs" by (rule distinct_card)  | 
|
100  | 
have "card (set (n_lists n xs)) = card (set xs) ^ n"  | 
|
101  | 
proof (induct n)  | 
|
102  | 
case 0 then show ?case by simp  | 
|
103  | 
next  | 
|
104  | 
case (Suc n)  | 
|
105  | 
moreover have "card (\<Union>ys\<in>set (n_lists n xs). (\<lambda>y. y # ys) ` set xs)  | 
|
106  | 
= (\<Sum>ys\<in>set (n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))"  | 
|
107  | 
by (rule card_UN_disjoint) auto  | 
|
108  | 
moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)"  | 
|
109  | 
by (rule card_image) (simp add: inj_on_def)  | 
|
110  | 
ultimately show ?case by auto  | 
|
111  | 
qed  | 
|
112  | 
also have "\<dots> = length xs ^ n" by (simp add: card_length)  | 
|
113  | 
finally show "card (set (n_lists n xs)) = length (n_lists n xs)"  | 
|
114  | 
by (simp add: length_n_lists)  | 
|
115  | 
qed  | 
|
116  | 
||
117  | 
lemma map_of_zip_enum_is_Some:  | 
|
118  | 
assumes "length ys = length (enum \<Colon> 'a\<Colon>enum list)"  | 
|
119  | 
shows "\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y"  | 
|
120  | 
proof -  | 
|
121  | 
from assms have "x \<in> set (enum \<Colon> 'a\<Colon>enum list) \<longleftrightarrow>  | 
|
122  | 
(\<exists>y. map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x = Some y)"  | 
|
123  | 
by (auto intro!: map_of_zip_is_Some)  | 
|
124  | 
then show ?thesis using enum_all by auto  | 
|
125  | 
qed  | 
|
126  | 
||
127  | 
lemma map_of_zip_enum_inject:  | 
|
128  | 
fixes xs ys :: "'b\<Colon>enum list"  | 
|
129  | 
assumes length: "length xs = length (enum \<Colon> 'a\<Colon>enum list)"  | 
|
130  | 
"length ys = length (enum \<Colon> 'a\<Colon>enum list)"  | 
|
131  | 
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)"  | 
|
132  | 
shows "xs = ys"  | 
|
133  | 
proof -  | 
|
134  | 
have "map_of (zip (enum \<Colon> 'a list) xs) = map_of (zip (enum \<Colon> 'a list) ys)"  | 
|
135  | 
proof  | 
|
136  | 
fix x :: 'a  | 
|
137  | 
from length map_of_zip_enum_is_Some obtain y1 y2  | 
|
138  | 
where "map_of (zip (enum \<Colon> 'a list) xs) x = Some y1"  | 
|
139  | 
and "map_of (zip (enum \<Colon> 'a list) ys) x = Some y2" by blast  | 
|
140  | 
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)"  | 
|
141  | 
by (auto dest: fun_cong)  | 
|
142  | 
ultimately show "map_of (zip (enum \<Colon> 'a\<Colon>enum list) xs) x = map_of (zip (enum \<Colon> 'a\<Colon>enum list) ys) x"  | 
|
143  | 
by simp  | 
|
144  | 
qed  | 
|
145  | 
with length enum_distinct show "xs = ys" by (rule map_of_zip_inject)  | 
|
146  | 
qed  | 
|
147  | 
||
148  | 
instantiation "fun" :: (enum, enum) enum  | 
|
149  | 
begin  | 
|
150  | 
||
151  | 
definition  | 
|
| 28562 | 152  | 
[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 | 153  | 
|
154  | 
instance proof  | 
|
155  | 
  show "UNIV = set (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | 
|
156  | 
proof (rule UNIV_eq_I)  | 
|
157  | 
fix f :: "'a \<Rightarrow> 'b"  | 
|
158  | 
have "f = the \<circ> map_of (zip (enum \<Colon> 'a\<Colon>enum list) (map f enum))"  | 
|
159  | 
by (auto simp add: map_of_zip_map expand_fun_eq)  | 
|
160  | 
then show "f \<in> set enum"  | 
|
161  | 
by (auto simp add: enum_fun_def set_n_lists)  | 
|
162  | 
qed  | 
|
163  | 
next  | 
|
164  | 
from map_of_zip_enum_inject  | 
|
165  | 
  show "distinct (enum \<Colon> ('a \<Rightarrow> 'b) list)"
 | 
|
166  | 
by (auto intro!: inj_onI simp add: enum_fun_def  | 
|
167  | 
distinct_map distinct_n_lists enum_distinct set_n_lists enum_all)  | 
|
168  | 
qed  | 
|
169  | 
||
170  | 
end  | 
|
171  | 
||
| 28562 | 172  | 
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: 
27487 
diff
changeset
 | 
173  | 
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: 
27487 
diff
changeset
 | 
174  | 
by (simp add: enum_fun_def Let_def)  | 
| 26444 | 175  | 
|
| 26348 | 176  | 
instantiation unit :: enum  | 
177  | 
begin  | 
|
178  | 
||
179  | 
definition  | 
|
180  | 
"enum = [()]"  | 
|
181  | 
||
| 31464 | 182  | 
instance proof  | 
183  | 
qed (simp_all add: enum_unit_def UNIV_unit)  | 
|
| 26348 | 184  | 
|
185  | 
end  | 
|
186  | 
||
187  | 
instantiation bool :: enum  | 
|
188  | 
begin  | 
|
189  | 
||
190  | 
definition  | 
|
191  | 
"enum = [False, True]"  | 
|
192  | 
||
| 31464 | 193  | 
instance proof  | 
194  | 
qed (simp_all add: enum_bool_def UNIV_bool)  | 
|
| 26348 | 195  | 
|
196  | 
end  | 
|
197  | 
||
198  | 
primrec product :: "'a list \<Rightarrow> 'b list \<Rightarrow> ('a \<times> 'b) list" where
 | 
|
199  | 
"product [] _ = []"  | 
|
200  | 
| "product (x#xs) ys = map (Pair x) ys @ product xs ys"  | 
|
201  | 
||
202  | 
lemma product_list_set:  | 
|
203  | 
"set (product xs ys) = set xs \<times> set ys"  | 
|
204  | 
by (induct xs) auto  | 
|
205  | 
||
| 26444 | 206  | 
lemma distinct_product:  | 
207  | 
assumes "distinct xs" and "distinct ys"  | 
|
208  | 
shows "distinct (product xs ys)"  | 
|
209  | 
using assms by (induct xs)  | 
|
210  | 
(auto intro: inj_onI simp add: product_list_set distinct_map)  | 
|
211  | 
||
| 26348 | 212  | 
instantiation * :: (enum, enum) enum  | 
213  | 
begin  | 
|
214  | 
||
215  | 
definition  | 
|
216  | 
"enum = product enum enum"  | 
|
217  | 
||
218  | 
instance by default  | 
|
| 26444 | 219  | 
(simp_all add: enum_prod_def product_list_set distinct_product enum_all enum_distinct)  | 
| 26348 | 220  | 
|
221  | 
end  | 
|
222  | 
||
223  | 
instantiation "+" :: (enum, enum) enum  | 
|
224  | 
begin  | 
|
225  | 
||
226  | 
definition  | 
|
227  | 
"enum = map Inl enum @ map Inr enum"  | 
|
228  | 
||
229  | 
instance by default  | 
|
| 26444 | 230  | 
(auto simp add: enum_all enum_sum_def, case_tac x, auto intro: inj_onI simp add: distinct_map enum_distinct)  | 
| 26348 | 231  | 
|
232  | 
end  | 
|
233  | 
||
234  | 
primrec sublists :: "'a list \<Rightarrow> 'a list list" where  | 
|
235  | 
"sublists [] = [[]]"  | 
|
236  | 
| "sublists (x#xs) = (let xss = sublists xs in map (Cons x) xss @ xss)"  | 
|
237  | 
||
| 26444 | 238  | 
lemma length_sublists:  | 
239  | 
"length (sublists xs) = Suc (Suc (0\<Colon>nat)) ^ length xs"  | 
|
240  | 
by (induct xs) (simp_all add: Let_def)  | 
|
241  | 
||
| 26348 | 242  | 
lemma sublists_powset:  | 
| 26444 | 243  | 
"set ` set (sublists xs) = Pow (set xs)"  | 
| 26348 | 244  | 
proof -  | 
245  | 
have aux: "\<And>x A. set ` Cons x ` A = insert x ` set ` A"  | 
|
246  | 
by (auto simp add: image_def)  | 
|
| 26444 | 247  | 
have "set (map set (sublists xs)) = Pow (set xs)"  | 
| 26348 | 248  | 
by (induct xs)  | 
| 
33639
 
603320b93668
New list theorems; added map_map to simpset, this is the prefered direction; allow sorting by a key
 
hoelzl 
parents: 
33635 
diff
changeset
 | 
249  | 
(simp_all add: aux Let_def Pow_insert Un_commute comp_def del: map_map)  | 
| 26444 | 250  | 
then show ?thesis by simp  | 
251  | 
qed  | 
|
252  | 
||
253  | 
lemma distinct_set_sublists:  | 
|
254  | 
assumes "distinct xs"  | 
|
255  | 
shows "distinct (map set (sublists xs))"  | 
|
256  | 
proof (rule card_distinct)  | 
|
257  | 
have "finite (set xs)" by rule  | 
|
258  | 
then have "card (Pow (set xs)) = Suc (Suc 0) ^ card (set xs)" by (rule card_Pow)  | 
|
259  | 
with assms distinct_card [of xs]  | 
|
260  | 
have "card (Pow (set xs)) = Suc (Suc 0) ^ length xs" by simp  | 
|
261  | 
then show "card (set (map set (sublists xs))) = length (map set (sublists xs))"  | 
|
262  | 
by (simp add: sublists_powset length_sublists)  | 
|
| 26348 | 263  | 
qed  | 
264  | 
||
265  | 
instantiation nibble :: enum  | 
|
266  | 
begin  | 
|
267  | 
||
268  | 
definition  | 
|
269  | 
"enum = [Nibble0, Nibble1, Nibble2, Nibble3, Nibble4, Nibble5, Nibble6, Nibble7,  | 
|
270  | 
Nibble8, Nibble9, NibbleA, NibbleB, NibbleC, NibbleD, NibbleE, NibbleF]"  | 
|
271  | 
||
| 31464 | 272  | 
instance proof  | 
273  | 
qed (simp_all add: enum_nibble_def UNIV_nibble)  | 
|
| 26348 | 274  | 
|
275  | 
end  | 
|
276  | 
||
277  | 
instantiation char :: enum  | 
|
278  | 
begin  | 
|
279  | 
||
280  | 
definition  | 
|
| 28562 | 281  | 
[code del]: "enum = map (split Char) (product enum enum)"  | 
| 26444 | 282  | 
|
| 31482 | 283  | 
lemma enum_chars [code]:  | 
284  | 
"enum = chars"  | 
|
285  | 
unfolding enum_char_def chars_def enum_nibble_def by simp  | 
|
| 26348 | 286  | 
|
| 31464 | 287  | 
instance proof  | 
288  | 
qed (auto intro: char.exhaust injI simp add: enum_char_def product_list_set enum_all full_SetCompr_eq [symmetric]  | 
|
289  | 
distinct_map distinct_product enum_distinct)  | 
|
| 26348 | 290  | 
|
291  | 
end  | 
|
292  | 
||
| 29024 | 293  | 
instantiation option :: (enum) enum  | 
294  | 
begin  | 
|
295  | 
||
296  | 
definition  | 
|
297  | 
"enum = None # map Some enum"  | 
|
298  | 
||
| 31464 | 299  | 
instance proof  | 
300  | 
qed (auto simp add: enum_all enum_option_def, rule option.exhaust, auto intro: simp add: distinct_map enum_distinct)  | 
|
| 29024 | 301  | 
|
302  | 
end  | 
|
303  | 
||
304  | 
end  |