| author | haftmann | 
| Wed, 19 May 2010 09:20:36 +0200 | |
| changeset 36980 | 1a4cc941171a | 
| parent 36274 | 42bd879dc1b0 | 
| child 37022 | f9681d9d1d56 | 
| permissions | -rw-r--r-- | 
| 35303 | 1  | 
(* Author: Florian Haftmann, TU Muenchen *)  | 
2  | 
||
3  | 
header {* Lists with elements distinct as canonical example for datatype invariants *}
 | 
|
4  | 
||
5  | 
theory Dlist  | 
|
6  | 
imports Main Fset  | 
|
7  | 
begin  | 
|
8  | 
||
9  | 
section {* Prelude *}
 | 
|
10  | 
||
11  | 
text {* Without canonical argument order, higher-order things tend to get confusing quite fast: *}
 | 
|
12  | 
||
13  | 
setup {* Sign.map_naming (Name_Space.add_path "List") *}
 | 
|
14  | 
||
15  | 
primrec member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where  | 
|
16  | 
"member [] y \<longleftrightarrow> False"  | 
|
17  | 
| "member (x#xs) y \<longleftrightarrow> x = y \<or> member xs y"  | 
|
18  | 
||
19  | 
lemma member_set:  | 
|
20  | 
"member = set"  | 
|
21  | 
proof (rule ext)+  | 
|
22  | 
fix xs :: "'a list" and x :: 'a  | 
|
23  | 
have "member xs x \<longleftrightarrow> x \<in> set xs" by (induct xs) auto  | 
|
24  | 
then show "member xs x = set xs x" by (simp add: mem_def)  | 
|
25  | 
qed  | 
|
26  | 
||
27  | 
lemma not_set_compl:  | 
|
28  | 
"Not \<circ> set xs = - set xs"  | 
|
29  | 
by (simp add: fun_Compl_def bool_Compl_def comp_def expand_fun_eq)  | 
|
30  | 
||
31  | 
primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
 | 
|
32  | 
"fold f [] s = s"  | 
|
33  | 
| "fold f (x#xs) s = fold f xs (f x s)"  | 
|
34  | 
||
35  | 
lemma foldl_fold:  | 
|
36  | 
"foldl f s xs = List.fold (\<lambda>x s. f s x) xs s"  | 
|
37  | 
by (induct xs arbitrary: s) simp_all  | 
|
38  | 
||
39  | 
setup {* Sign.map_naming Name_Space.parent_path *}
 | 
|
40  | 
||
41  | 
||
42  | 
section {* The type of distinct lists *}
 | 
|
43  | 
||
44  | 
typedef (open) 'a dlist = "{xs::'a list. distinct xs}"
 | 
|
45  | 
morphisms list_of_dlist Abs_dlist  | 
|
46  | 
proof  | 
|
47  | 
show "[] \<in> ?dlist" by simp  | 
|
48  | 
qed  | 
|
49  | 
||
| 36274 | 50  | 
lemma dlist_ext:  | 
51  | 
assumes "list_of_dlist xs = list_of_dlist ys"  | 
|
52  | 
shows "xs = ys"  | 
|
53  | 
using assms by (simp add: list_of_dlist_inject)  | 
|
54  | 
||
| 
36112
 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 
haftmann 
parents: 
35688 
diff
changeset
 | 
55  | 
|
| 35303 | 56  | 
text {* Formal, totalized constructor for @{typ "'a dlist"}: *}
 | 
57  | 
||
58  | 
definition Dlist :: "'a list \<Rightarrow> 'a dlist" where  | 
|
59  | 
[code del]: "Dlist xs = Abs_dlist (remdups xs)"  | 
|
60  | 
||
61  | 
lemma distinct_list_of_dlist [simp]:  | 
|
62  | 
"distinct (list_of_dlist dxs)"  | 
|
63  | 
using list_of_dlist [of dxs] by simp  | 
|
64  | 
||
65  | 
lemma list_of_dlist_Dlist [simp]:  | 
|
66  | 
"list_of_dlist (Dlist xs) = remdups xs"  | 
|
67  | 
by (simp add: Dlist_def Abs_dlist_inverse)  | 
|
68  | 
||
| 
36112
 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 
haftmann 
parents: 
35688 
diff
changeset
 | 
69  | 
lemma Dlist_list_of_dlist [simp, code abstype]:  | 
| 35303 | 70  | 
"Dlist (list_of_dlist dxs) = dxs"  | 
71  | 
by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)  | 
|
72  | 
||
73  | 
||
74  | 
text {* Fundamental operations: *}
 | 
|
75  | 
||
76  | 
definition empty :: "'a dlist" where  | 
|
77  | 
"empty = Dlist []"  | 
|
78  | 
||
79  | 
definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where  | 
|
80  | 
"insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"  | 
|
81  | 
||
82  | 
definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where  | 
|
83  | 
"remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"  | 
|
84  | 
||
85  | 
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
 | 
|
86  | 
"map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"  | 
|
87  | 
||
88  | 
definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
 | 
|
89  | 
"filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"  | 
|
90  | 
||
91  | 
||
92  | 
text {* Derived operations: *}
 | 
|
93  | 
||
94  | 
definition null :: "'a dlist \<Rightarrow> bool" where  | 
|
95  | 
"null dxs = List.null (list_of_dlist dxs)"  | 
|
96  | 
||
97  | 
definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where  | 
|
98  | 
"member dxs = List.member (list_of_dlist dxs)"  | 
|
99  | 
||
100  | 
definition length :: "'a dlist \<Rightarrow> nat" where  | 
|
101  | 
"length dxs = List.length (list_of_dlist dxs)"  | 
|
102  | 
||
103  | 
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
 | 
|
104  | 
"fold f dxs = List.fold f (list_of_dlist dxs)"  | 
|
105  | 
||
106  | 
||
107  | 
section {* Executable version obeying invariant *}
 | 
|
108  | 
||
109  | 
lemma list_of_dlist_empty [simp, code abstract]:  | 
|
110  | 
"list_of_dlist empty = []"  | 
|
111  | 
by (simp add: empty_def)  | 
|
112  | 
||
113  | 
lemma list_of_dlist_insert [simp, code abstract]:  | 
|
114  | 
"list_of_dlist (insert x dxs) = List.insert x (list_of_dlist dxs)"  | 
|
115  | 
by (simp add: insert_def)  | 
|
116  | 
||
117  | 
lemma list_of_dlist_remove [simp, code abstract]:  | 
|
118  | 
"list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)"  | 
|
119  | 
by (simp add: remove_def)  | 
|
120  | 
||
121  | 
lemma list_of_dlist_map [simp, code abstract]:  | 
|
122  | 
"list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))"  | 
|
123  | 
by (simp add: map_def)  | 
|
124  | 
||
125  | 
lemma list_of_dlist_filter [simp, code abstract]:  | 
|
126  | 
"list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)"  | 
|
127  | 
by (simp add: filter_def)  | 
|
128  | 
||
129  | 
||
| 
36980
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
130  | 
text {* Explicit executable conversion *}
 | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
131  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
132  | 
definition dlist_of_list [simp]:  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
133  | 
"dlist_of_list = Dlist"  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
134  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
135  | 
lemma [code abstract]:  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
136  | 
"list_of_dlist (dlist_of_list xs) = remdups xs"  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
137  | 
by simp  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
138  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
139  | 
|
| 35303 | 140  | 
section {* Implementation of sets by distinct lists -- canonical! *}
 | 
141  | 
||
142  | 
definition Set :: "'a dlist \<Rightarrow> 'a fset" where  | 
|
143  | 
"Set dxs = Fset.Set (list_of_dlist dxs)"  | 
|
144  | 
||
145  | 
definition Coset :: "'a dlist \<Rightarrow> 'a fset" where  | 
|
146  | 
"Coset dxs = Fset.Coset (list_of_dlist dxs)"  | 
|
147  | 
||
148  | 
code_datatype Set Coset  | 
|
149  | 
||
150  | 
declare member_code [code del]  | 
|
151  | 
declare is_empty_Set [code del]  | 
|
152  | 
declare empty_Set [code del]  | 
|
153  | 
declare UNIV_Set [code del]  | 
|
154  | 
declare insert_Set [code del]  | 
|
155  | 
declare remove_Set [code del]  | 
|
156  | 
declare map_Set [code del]  | 
|
157  | 
declare filter_Set [code del]  | 
|
158  | 
declare forall_Set [code del]  | 
|
159  | 
declare exists_Set [code del]  | 
|
160  | 
declare card_Set [code del]  | 
|
161  | 
declare inter_project [code del]  | 
|
162  | 
declare subtract_remove [code del]  | 
|
163  | 
declare union_insert [code del]  | 
|
164  | 
declare Infimum_inf [code del]  | 
|
165  | 
declare Supremum_sup [code del]  | 
|
166  | 
||
167  | 
lemma Set_Dlist [simp]:  | 
|
168  | 
"Set (Dlist xs) = Fset (set xs)"  | 
|
169  | 
by (simp add: Set_def Fset.Set_def)  | 
|
170  | 
||
171  | 
lemma Coset_Dlist [simp]:  | 
|
172  | 
"Coset (Dlist xs) = Fset (- set xs)"  | 
|
173  | 
by (simp add: Coset_def Fset.Coset_def)  | 
|
174  | 
||
175  | 
lemma member_Set [simp]:  | 
|
176  | 
"Fset.member (Set dxs) = List.member (list_of_dlist dxs)"  | 
|
177  | 
by (simp add: Set_def member_set)  | 
|
178  | 
||
179  | 
lemma member_Coset [simp]:  | 
|
180  | 
"Fset.member (Coset dxs) = Not \<circ> List.member (list_of_dlist dxs)"  | 
|
181  | 
by (simp add: Coset_def member_set not_set_compl)  | 
|
182  | 
||
| 
36980
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
183  | 
lemma Set_dlist_of_list [code]:  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
184  | 
"Fset.Set xs = Set (dlist_of_list xs)"  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
185  | 
by simp  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
186  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
187  | 
lemma Coset_dlist_of_list [code]:  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
188  | 
"Fset.Coset xs = Coset (dlist_of_list xs)"  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
189  | 
by simp  | 
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
190  | 
|
| 35303 | 191  | 
lemma is_empty_Set [code]:  | 
192  | 
"Fset.is_empty (Set dxs) \<longleftrightarrow> null dxs"  | 
|
193  | 
by (simp add: null_def null_empty member_set)  | 
|
194  | 
||
195  | 
lemma bot_code [code]:  | 
|
196  | 
"bot = Set empty"  | 
|
197  | 
by (simp add: empty_def)  | 
|
198  | 
||
199  | 
lemma top_code [code]:  | 
|
200  | 
"top = Coset empty"  | 
|
201  | 
by (simp add: empty_def)  | 
|
202  | 
||
203  | 
lemma insert_code [code]:  | 
|
204  | 
"Fset.insert x (Set dxs) = Set (insert x dxs)"  | 
|
205  | 
"Fset.insert x (Coset dxs) = Coset (remove x dxs)"  | 
|
206  | 
by (simp_all add: insert_def remove_def member_set not_set_compl)  | 
|
207  | 
||
208  | 
lemma remove_code [code]:  | 
|
209  | 
"Fset.remove x (Set dxs) = Set (remove x dxs)"  | 
|
210  | 
"Fset.remove x (Coset dxs) = Coset (insert x dxs)"  | 
|
211  | 
by (auto simp add: insert_def remove_def member_set not_set_compl)  | 
|
212  | 
||
213  | 
lemma member_code [code]:  | 
|
214  | 
"Fset.member (Set dxs) = member dxs"  | 
|
215  | 
"Fset.member (Coset dxs) = Not \<circ> member dxs"  | 
|
216  | 
by (simp_all add: member_def)  | 
|
217  | 
||
218  | 
lemma map_code [code]:  | 
|
219  | 
"Fset.map f (Set dxs) = Set (map f dxs)"  | 
|
220  | 
by (simp add: member_set)  | 
|
221  | 
||
222  | 
lemma filter_code [code]:  | 
|
223  | 
"Fset.filter f (Set dxs) = Set (filter f dxs)"  | 
|
224  | 
by (simp add: member_set)  | 
|
225  | 
||
226  | 
lemma forall_Set [code]:  | 
|
227  | 
"Fset.forall P (Set xs) \<longleftrightarrow> list_all P (list_of_dlist xs)"  | 
|
228  | 
by (simp add: member_set list_all_iff)  | 
|
229  | 
||
230  | 
lemma exists_Set [code]:  | 
|
231  | 
"Fset.exists P (Set xs) \<longleftrightarrow> list_ex P (list_of_dlist xs)"  | 
|
232  | 
by (simp add: member_set list_ex_iff)  | 
|
233  | 
||
234  | 
lemma card_code [code]:  | 
|
235  | 
"Fset.card (Set dxs) = length dxs"  | 
|
236  | 
by (simp add: length_def member_set distinct_card)  | 
|
237  | 
||
238  | 
lemma foldl_list_of_dlist:  | 
|
239  | 
"foldl f s (list_of_dlist dxs) = fold (\<lambda>x s. f s x) dxs s"  | 
|
240  | 
by (simp add: foldl_fold fold_def)  | 
|
241  | 
||
242  | 
lemma inter_code [code]:  | 
|
243  | 
"inf A (Set xs) = Set (filter (Fset.member A) xs)"  | 
|
244  | 
"inf A (Coset xs) = fold Fset.remove xs A"  | 
|
245  | 
by (simp_all only: Set_def Coset_def foldl_list_of_dlist inter_project list_of_dlist_filter)  | 
|
246  | 
||
247  | 
lemma subtract_code [code]:  | 
|
248  | 
"A - Set xs = fold Fset.remove xs A"  | 
|
249  | 
"A - Coset xs = Set (filter (Fset.member A) xs)"  | 
|
250  | 
by (simp_all only: Set_def Coset_def foldl_list_of_dlist subtract_remove list_of_dlist_filter)  | 
|
251  | 
||
252  | 
lemma union_code [code]:  | 
|
253  | 
"sup (Set xs) A = fold Fset.insert xs A"  | 
|
254  | 
"sup (Coset xs) A = Coset (filter (Not \<circ> Fset.member A) xs)"  | 
|
255  | 
by (simp_all only: Set_def Coset_def foldl_list_of_dlist union_insert list_of_dlist_filter)  | 
|
256  | 
||
257  | 
context complete_lattice  | 
|
258  | 
begin  | 
|
259  | 
||
260  | 
lemma Infimum_code [code]:  | 
|
261  | 
"Infimum (Set As) = fold inf As top"  | 
|
262  | 
by (simp only: Set_def Infimum_inf foldl_list_of_dlist inf.commute)  | 
|
263  | 
||
264  | 
lemma Supremum_code [code]:  | 
|
265  | 
"Supremum (Set As) = fold sup As bot"  | 
|
266  | 
by (simp only: Set_def Supremum_sup foldl_list_of_dlist sup.commute)  | 
|
267  | 
||
268  | 
end  | 
|
269  | 
||
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36112 
diff
changeset
 | 
270  | 
hide_const (open) member fold empty insert remove map filter null member length fold  | 
| 35303 | 271  | 
|
272  | 
end  |