| author | nipkow | 
| Sat, 19 Jan 2013 21:05:05 +0100 | |
| changeset 50986 | c54ea7f5418f | 
| parent 49834 | b27bbb021df1 | 
| child 55467 | a5c9002bc54d | 
| 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  | 
|
| 
45990
 
b7b905b23b2a
incorporated More_Set and More_List into the Main body -- to be consolidated later
 
haftmann 
parents: 
45927 
diff
changeset
 | 
6  | 
imports Main  | 
| 35303 | 7  | 
begin  | 
8  | 
||
| 43146 | 9  | 
subsection {* The type of distinct lists *}
 | 
| 35303 | 10  | 
|
| 49834 | 11  | 
typedef 'a dlist = "{xs::'a list. distinct xs}"
 | 
| 35303 | 12  | 
morphisms list_of_dlist Abs_dlist  | 
13  | 
proof  | 
|
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
43764 
diff
changeset
 | 
14  | 
  show "[] \<in> {xs. distinct xs}" by simp
 | 
| 35303 | 15  | 
qed  | 
16  | 
||
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
17  | 
lemma dlist_eq_iff:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
18  | 
"dxs = dys \<longleftrightarrow> list_of_dlist dxs = list_of_dlist dys"  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
19  | 
by (simp add: list_of_dlist_inject)  | 
| 36274 | 20  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
21  | 
lemma dlist_eqI:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
22  | 
"list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys"  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
23  | 
by (simp add: dlist_eq_iff)  | 
| 
36112
 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 
haftmann 
parents: 
35688 
diff
changeset
 | 
24  | 
|
| 35303 | 25  | 
text {* Formal, totalized constructor for @{typ "'a dlist"}: *}
 | 
26  | 
||
27  | 
definition Dlist :: "'a list \<Rightarrow> 'a dlist" where  | 
|
| 37765 | 28  | 
"Dlist xs = Abs_dlist (remdups xs)"  | 
| 35303 | 29  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
30  | 
lemma distinct_list_of_dlist [simp, intro]:  | 
| 35303 | 31  | 
"distinct (list_of_dlist dxs)"  | 
32  | 
using list_of_dlist [of dxs] by simp  | 
|
33  | 
||
34  | 
lemma list_of_dlist_Dlist [simp]:  | 
|
35  | 
"list_of_dlist (Dlist xs) = remdups xs"  | 
|
36  | 
by (simp add: Dlist_def Abs_dlist_inverse)  | 
|
37  | 
||
| 39727 | 38  | 
lemma remdups_list_of_dlist [simp]:  | 
39  | 
"remdups (list_of_dlist dxs) = list_of_dlist dxs"  | 
|
40  | 
by simp  | 
|
41  | 
||
| 
36112
 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 
haftmann 
parents: 
35688 
diff
changeset
 | 
42  | 
lemma Dlist_list_of_dlist [simp, code abstype]:  | 
| 35303 | 43  | 
"Dlist (list_of_dlist dxs) = dxs"  | 
44  | 
by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)  | 
|
45  | 
||
46  | 
||
47  | 
text {* Fundamental operations: *}
 | 
|
48  | 
||
49  | 
definition empty :: "'a dlist" where  | 
|
50  | 
"empty = Dlist []"  | 
|
51  | 
||
52  | 
definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where  | 
|
53  | 
"insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"  | 
|
54  | 
||
55  | 
definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where  | 
|
56  | 
"remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"  | 
|
57  | 
||
58  | 
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
 | 
|
59  | 
"map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"  | 
|
60  | 
||
61  | 
definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
 | 
|
62  | 
"filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"  | 
|
63  | 
||
64  | 
||
65  | 
text {* Derived operations: *}
 | 
|
66  | 
||
67  | 
definition null :: "'a dlist \<Rightarrow> bool" where  | 
|
68  | 
"null dxs = List.null (list_of_dlist dxs)"  | 
|
69  | 
||
70  | 
definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where  | 
|
71  | 
"member dxs = List.member (list_of_dlist dxs)"  | 
|
72  | 
||
73  | 
definition length :: "'a dlist \<Rightarrow> nat" where  | 
|
74  | 
"length dxs = List.length (list_of_dlist dxs)"  | 
|
75  | 
||
76  | 
definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
 | 
|
| 
46133
 
d9fe85d3d2cd
incorporated canonical fold combinator on lists into body of List theory; refactored passages on List.fold(l/r)
 
haftmann 
parents: 
45990 
diff
changeset
 | 
77  | 
"fold f dxs = List.fold f (list_of_dlist dxs)"  | 
| 
37022
 
f9681d9d1d56
moved generic List operations to theory More_List
 
haftmann 
parents: 
36980 
diff
changeset
 | 
78  | 
|
| 
 
f9681d9d1d56
moved generic List operations to theory More_List
 
haftmann 
parents: 
36980 
diff
changeset
 | 
79  | 
definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
 | 
| 
 
f9681d9d1d56
moved generic List operations to theory More_List
 
haftmann 
parents: 
36980 
diff
changeset
 | 
80  | 
"foldr f dxs = List.foldr f (list_of_dlist dxs)"  | 
| 35303 | 81  | 
|
82  | 
||
| 43146 | 83  | 
subsection {* Executable version obeying invariant *}
 | 
| 35303 | 84  | 
|
85  | 
lemma list_of_dlist_empty [simp, code abstract]:  | 
|
86  | 
"list_of_dlist empty = []"  | 
|
87  | 
by (simp add: empty_def)  | 
|
88  | 
||
89  | 
lemma list_of_dlist_insert [simp, code abstract]:  | 
|
90  | 
"list_of_dlist (insert x dxs) = List.insert x (list_of_dlist dxs)"  | 
|
91  | 
by (simp add: insert_def)  | 
|
92  | 
||
93  | 
lemma list_of_dlist_remove [simp, code abstract]:  | 
|
94  | 
"list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)"  | 
|
95  | 
by (simp add: remove_def)  | 
|
96  | 
||
97  | 
lemma list_of_dlist_map [simp, code abstract]:  | 
|
98  | 
"list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))"  | 
|
99  | 
by (simp add: map_def)  | 
|
100  | 
||
101  | 
lemma list_of_dlist_filter [simp, code abstract]:  | 
|
102  | 
"list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)"  | 
|
103  | 
by (simp add: filter_def)  | 
|
104  | 
||
105  | 
||
| 
36980
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
106  | 
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
 | 
107  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
108  | 
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
 | 
109  | 
"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
 | 
110  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
111  | 
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
 | 
112  | 
"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
 | 
113  | 
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
 | 
114  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
115  | 
|
| 38512 | 116  | 
text {* Equality *}
 | 
117  | 
||
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
118  | 
instantiation dlist :: (equal) equal  | 
| 38512 | 119  | 
begin  | 
120  | 
||
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
121  | 
definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)"  | 
| 38512 | 122  | 
|
123  | 
instance proof  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
124  | 
qed (simp add: equal_dlist_def equal list_of_dlist_inject)  | 
| 38512 | 125  | 
|
126  | 
end  | 
|
127  | 
||
| 43764 | 128  | 
declare equal_dlist_def [code]  | 
129  | 
||
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
130  | 
lemma [code nbe]:  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
131  | 
"HOL.equal (dxs :: 'a::equal dlist) dxs \<longleftrightarrow> True"  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
132  | 
by (fact equal_refl)  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
133  | 
|
| 38512 | 134  | 
|
| 43146 | 135  | 
subsection {* Induction principle and case distinction *}
 | 
| 37106 | 136  | 
|
137  | 
lemma dlist_induct [case_names empty insert, induct type: dlist]:  | 
|
138  | 
assumes empty: "P empty"  | 
|
139  | 
assumes insrt: "\<And>x dxs. \<not> member dxs x \<Longrightarrow> P dxs \<Longrightarrow> P (insert x dxs)"  | 
|
140  | 
shows "P dxs"  | 
|
141  | 
proof (cases dxs)  | 
|
142  | 
case (Abs_dlist xs)  | 
|
143  | 
then have "distinct xs" and dxs: "dxs = Dlist xs" by (simp_all add: Dlist_def distinct_remdups_id)  | 
|
144  | 
from `distinct xs` have "P (Dlist xs)"  | 
|
| 
39915
 
ecf97cf3d248
turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
 
haftmann 
parents: 
39727 
diff
changeset
 | 
145  | 
proof (induct xs)  | 
| 37106 | 146  | 
case Nil from empty show ?case by (simp add: empty_def)  | 
147  | 
next  | 
|
| 
40122
 
1d8ad2ff3e01
dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
 
haftmann 
parents: 
39915 
diff
changeset
 | 
148  | 
case (Cons x xs)  | 
| 37106 | 149  | 
then have "\<not> member (Dlist xs) x" and "P (Dlist xs)"  | 
| 
37595
 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 
haftmann 
parents: 
37473 
diff
changeset
 | 
150  | 
by (simp_all add: member_def List.member_def)  | 
| 37106 | 151  | 
with insrt have "P (insert x (Dlist xs))" .  | 
| 
40122
 
1d8ad2ff3e01
dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
 
haftmann 
parents: 
39915 
diff
changeset
 | 
152  | 
with Cons show ?case by (simp add: insert_def distinct_remdups_id)  | 
| 37106 | 153  | 
qed  | 
154  | 
with dxs show "P dxs" by simp  | 
|
155  | 
qed  | 
|
156  | 
||
157  | 
lemma dlist_case [case_names empty insert, cases type: dlist]:  | 
|
158  | 
assumes empty: "dxs = empty \<Longrightarrow> P"  | 
|
159  | 
assumes insert: "\<And>x dys. \<not> member dys x \<Longrightarrow> dxs = insert x dys \<Longrightarrow> P"  | 
|
160  | 
shows P  | 
|
161  | 
proof (cases dxs)  | 
|
162  | 
case (Abs_dlist xs)  | 
|
163  | 
then have dxs: "dxs = Dlist xs" and distinct: "distinct xs"  | 
|
164  | 
by (simp_all add: Dlist_def distinct_remdups_id)  | 
|
165  | 
show P proof (cases xs)  | 
|
166  | 
case Nil with dxs have "dxs = empty" by (simp add: empty_def)  | 
|
167  | 
with empty show P .  | 
|
168  | 
next  | 
|
169  | 
case (Cons x xs)  | 
|
170  | 
with dxs distinct have "\<not> member (Dlist xs) x"  | 
|
171  | 
and "dxs = insert x (Dlist xs)"  | 
|
| 
37595
 
9591362629e3
dropped ancient infix mem; refined code generation operations in List.thy
 
haftmann 
parents: 
37473 
diff
changeset
 | 
172  | 
by (simp_all add: member_def List.member_def insert_def distinct_remdups_id)  | 
| 37106 | 173  | 
with insert show P .  | 
174  | 
qed  | 
|
175  | 
qed  | 
|
176  | 
||
177  | 
||
| 43146 | 178  | 
subsection {* Functorial structure *}
 | 
| 40603 | 179  | 
|
| 
41505
 
6d19301074cf
"enriched_type" replaces less specific "type_lifting"
 
haftmann 
parents: 
41372 
diff
changeset
 | 
180  | 
enriched_type map: map  | 
| 41372 | 181  | 
by (simp_all add: List.map.id remdups_map_remdups fun_eq_iff dlist_eq_iff)  | 
| 40603 | 182  | 
|
| 48282 | 183  | 
|
| 
45927
 
e0305e4f02c9
adding quickcheck generator for distinct lists; adding examples
 
bulwahn 
parents: 
45694 
diff
changeset
 | 
184  | 
subsection {* Quickcheck generators *}
 | 
| 
 
e0305e4f02c9
adding quickcheck generator for distinct lists; adding examples
 
bulwahn 
parents: 
45694 
diff
changeset
 | 
185  | 
|
| 46565 | 186  | 
quickcheck_generator dlist predicate: distinct constructors: empty, insert  | 
| 40603 | 187  | 
|
| 48282 | 188  | 
|
| 
37022
 
f9681d9d1d56
moved generic List operations to theory More_List
 
haftmann 
parents: 
36980 
diff
changeset
 | 
189  | 
hide_const (open) member fold foldr empty insert remove map filter null member length fold  | 
| 35303 | 190  | 
|
191  | 
end  |