| author | wenzelm | 
| Wed, 03 May 2017 23:04:25 +0200 | |
| changeset 65703 | cead65c19f2e | 
| parent 63375 | 59803048b0e8 | 
| child 67399 | eab6ce8368fa | 
| permissions | -rw-r--r-- | 
| 62139 | 1  | 
(* Author: Florian Haftmann, TU Muenchen  | 
2  | 
Author: Andreas Lochbihler, ETH Zurich *)  | 
|
| 35303 | 3  | 
|
| 60500 | 4  | 
section \<open>Lists with elements distinct as canonical example for datatype invariants\<close>  | 
| 35303 | 5  | 
|
6  | 
theory Dlist  | 
|
| 
45990
 
b7b905b23b2a
incorporated More_Set and More_List into the Main body -- to be consolidated later
 
haftmann 
parents: 
45927 
diff
changeset
 | 
7  | 
imports Main  | 
| 35303 | 8  | 
begin  | 
9  | 
||
| 60500 | 10  | 
subsection \<open>The type of distinct lists\<close>  | 
| 35303 | 11  | 
|
| 49834 | 12  | 
typedef 'a dlist = "{xs::'a list. distinct xs}"
 | 
| 35303 | 13  | 
morphisms list_of_dlist Abs_dlist  | 
14  | 
proof  | 
|
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
43764 
diff
changeset
 | 
15  | 
  show "[] \<in> {xs. distinct xs}" by simp
 | 
| 35303 | 16  | 
qed  | 
17  | 
||
| 62139 | 18  | 
setup_lifting type_definition_dlist  | 
19  | 
||
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
20  | 
lemma dlist_eq_iff:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
21  | 
"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
 | 
22  | 
by (simp add: list_of_dlist_inject)  | 
| 36274 | 23  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
24  | 
lemma dlist_eqI:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
25  | 
"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
 | 
26  | 
by (simp add: dlist_eq_iff)  | 
| 
36112
 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 
haftmann 
parents: 
35688 
diff
changeset
 | 
27  | 
|
| 60500 | 28  | 
text \<open>Formal, totalized constructor for @{typ "'a dlist"}:\<close>
 | 
| 35303 | 29  | 
|
30  | 
definition Dlist :: "'a list \<Rightarrow> 'a dlist" where  | 
|
| 37765 | 31  | 
"Dlist xs = Abs_dlist (remdups xs)"  | 
| 35303 | 32  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
38857 
diff
changeset
 | 
33  | 
lemma distinct_list_of_dlist [simp, intro]:  | 
| 35303 | 34  | 
"distinct (list_of_dlist dxs)"  | 
35  | 
using list_of_dlist [of dxs] by simp  | 
|
36  | 
||
37  | 
lemma list_of_dlist_Dlist [simp]:  | 
|
38  | 
"list_of_dlist (Dlist xs) = remdups xs"  | 
|
39  | 
by (simp add: Dlist_def Abs_dlist_inverse)  | 
|
40  | 
||
| 39727 | 41  | 
lemma remdups_list_of_dlist [simp]:  | 
42  | 
"remdups (list_of_dlist dxs) = list_of_dlist dxs"  | 
|
43  | 
by simp  | 
|
44  | 
||
| 
36112
 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 
haftmann 
parents: 
35688 
diff
changeset
 | 
45  | 
lemma Dlist_list_of_dlist [simp, code abstype]:  | 
| 35303 | 46  | 
"Dlist (list_of_dlist dxs) = dxs"  | 
47  | 
by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)  | 
|
48  | 
||
49  | 
||
| 60500 | 50  | 
text \<open>Fundamental operations:\<close>  | 
| 35303 | 51  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
52  | 
context  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
53  | 
begin  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
54  | 
|
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
55  | 
qualified definition empty :: "'a dlist" where  | 
| 35303 | 56  | 
"empty = Dlist []"  | 
57  | 
||
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
58  | 
qualified definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where  | 
| 35303 | 59  | 
"insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"  | 
60  | 
||
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
61  | 
qualified definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where  | 
| 35303 | 62  | 
"remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"  | 
63  | 
||
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
64  | 
qualified definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
 | 
| 35303 | 65  | 
"map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"  | 
66  | 
||
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
67  | 
qualified definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
 | 
| 35303 | 68  | 
"filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"  | 
69  | 
||
| 
63375
 
59803048b0e8
basic facts about almost everywhere fix bijections
 
haftmann 
parents: 
62390 
diff
changeset
 | 
70  | 
qualified definition rotate :: "nat \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where  | 
| 
 
59803048b0e8
basic facts about almost everywhere fix bijections
 
haftmann 
parents: 
62390 
diff
changeset
 | 
71  | 
"rotate n dxs = Dlist (List.rotate n (list_of_dlist dxs))"  | 
| 
 
59803048b0e8
basic facts about almost everywhere fix bijections
 
haftmann 
parents: 
62390 
diff
changeset
 | 
72  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
73  | 
end  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
74  | 
|
| 35303 | 75  | 
|
| 60500 | 76  | 
text \<open>Derived operations:\<close>  | 
| 35303 | 77  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
78  | 
context  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
79  | 
begin  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
80  | 
|
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
81  | 
qualified definition null :: "'a dlist \<Rightarrow> bool" where  | 
| 35303 | 82  | 
"null dxs = List.null (list_of_dlist dxs)"  | 
83  | 
||
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
84  | 
qualified definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where  | 
| 35303 | 85  | 
"member dxs = List.member (list_of_dlist dxs)"  | 
86  | 
||
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
87  | 
qualified definition length :: "'a dlist \<Rightarrow> nat" where  | 
| 35303 | 88  | 
"length dxs = List.length (list_of_dlist dxs)"  | 
89  | 
||
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
90  | 
qualified 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
 | 
91  | 
"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
 | 
92  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
93  | 
qualified definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
 | 
| 
37022
 
f9681d9d1d56
moved generic List operations to theory More_List
 
haftmann 
parents: 
36980 
diff
changeset
 | 
94  | 
"foldr f dxs = List.foldr f (list_of_dlist dxs)"  | 
| 35303 | 95  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
96  | 
end  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
97  | 
|
| 35303 | 98  | 
|
| 60500 | 99  | 
subsection \<open>Executable version obeying invariant\<close>  | 
| 35303 | 100  | 
|
101  | 
lemma list_of_dlist_empty [simp, code abstract]:  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
102  | 
"list_of_dlist Dlist.empty = []"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
103  | 
by (simp add: Dlist.empty_def)  | 
| 35303 | 104  | 
|
105  | 
lemma list_of_dlist_insert [simp, code abstract]:  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
106  | 
"list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
107  | 
by (simp add: Dlist.insert_def)  | 
| 35303 | 108  | 
|
109  | 
lemma list_of_dlist_remove [simp, code abstract]:  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
110  | 
"list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
111  | 
by (simp add: Dlist.remove_def)  | 
| 35303 | 112  | 
|
113  | 
lemma list_of_dlist_map [simp, code abstract]:  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
114  | 
"list_of_dlist (Dlist.map f dxs) = remdups (List.map f (list_of_dlist dxs))"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
115  | 
by (simp add: Dlist.map_def)  | 
| 35303 | 116  | 
|
117  | 
lemma list_of_dlist_filter [simp, code abstract]:  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
118  | 
"list_of_dlist (Dlist.filter P dxs) = List.filter P (list_of_dlist dxs)"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
119  | 
by (simp add: Dlist.filter_def)  | 
| 35303 | 120  | 
|
| 
63375
 
59803048b0e8
basic facts about almost everywhere fix bijections
 
haftmann 
parents: 
62390 
diff
changeset
 | 
121  | 
lemma list_of_dlist_rotate [simp, code abstract]:  | 
| 
 
59803048b0e8
basic facts about almost everywhere fix bijections
 
haftmann 
parents: 
62390 
diff
changeset
 | 
122  | 
"list_of_dlist (Dlist.rotate n dxs) = List.rotate n (list_of_dlist dxs)"  | 
| 
 
59803048b0e8
basic facts about almost everywhere fix bijections
 
haftmann 
parents: 
62390 
diff
changeset
 | 
123  | 
by (simp add: Dlist.rotate_def)  | 
| 
 
59803048b0e8
basic facts about almost everywhere fix bijections
 
haftmann 
parents: 
62390 
diff
changeset
 | 
124  | 
|
| 35303 | 125  | 
|
| 60500 | 126  | 
text \<open>Explicit executable conversion\<close>  | 
| 
36980
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
127  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
128  | 
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
 | 
129  | 
"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
 | 
130  | 
|
| 
 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 
haftmann 
parents: 
36274 
diff
changeset
 | 
131  | 
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
 | 
132  | 
"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
 | 
133  | 
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
 | 
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  | 
|
| 60500 | 136  | 
text \<open>Equality\<close>  | 
| 38512 | 137  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
138  | 
instantiation dlist :: (equal) equal  | 
| 38512 | 139  | 
begin  | 
140  | 
||
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
141  | 
definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)"  | 
| 38512 | 142  | 
|
| 60679 | 143  | 
instance  | 
144  | 
by standard (simp add: equal_dlist_def equal list_of_dlist_inject)  | 
|
| 38512 | 145  | 
|
146  | 
end  | 
|
147  | 
||
| 43764 | 148  | 
declare equal_dlist_def [code]  | 
149  | 
||
| 60679 | 150  | 
lemma [code nbe]: "HOL.equal (dxs :: 'a::equal dlist) dxs \<longleftrightarrow> True"  | 
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
151  | 
by (fact equal_refl)  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
38512 
diff
changeset
 | 
152  | 
|
| 38512 | 153  | 
|
| 60500 | 154  | 
subsection \<open>Induction principle and case distinction\<close>  | 
| 37106 | 155  | 
|
156  | 
lemma dlist_induct [case_names empty insert, induct type: dlist]:  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
157  | 
assumes empty: "P Dlist.empty"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
158  | 
assumes insrt: "\<And>x dxs. \<not> Dlist.member dxs x \<Longrightarrow> P dxs \<Longrightarrow> P (Dlist.insert x dxs)"  | 
| 37106 | 159  | 
shows "P dxs"  | 
160  | 
proof (cases dxs)  | 
|
161  | 
case (Abs_dlist xs)  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
162  | 
then have "distinct xs" and dxs: "dxs = Dlist xs"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
163  | 
by (simp_all add: Dlist_def distinct_remdups_id)  | 
| 60500 | 164  | 
from \<open>distinct xs\<close> 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
 | 
165  | 
proof (induct xs)  | 
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
166  | 
case Nil from empty show ?case by (simp add: Dlist.empty_def)  | 
| 37106 | 167  | 
next  | 
| 
40122
 
1d8ad2ff3e01
dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
 
haftmann 
parents: 
39915 
diff
changeset
 | 
168  | 
case (Cons x xs)  | 
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
169  | 
then have "\<not> Dlist.member (Dlist xs) x" and "P (Dlist xs)"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
170  | 
by (simp_all add: Dlist.member_def List.member_def)  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
171  | 
with insrt have "P (Dlist.insert x (Dlist xs))" .  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
172  | 
with Cons show ?case by (simp add: Dlist.insert_def distinct_remdups_id)  | 
| 37106 | 173  | 
qed  | 
174  | 
with dxs show "P dxs" by simp  | 
|
175  | 
qed  | 
|
176  | 
||
| 55913 | 177  | 
lemma dlist_case [cases type: dlist]:  | 
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
178  | 
obtains (empty) "dxs = Dlist.empty"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
179  | 
| (insert) x dys where "\<not> Dlist.member dys x" and "dxs = Dlist.insert x dys"  | 
| 37106 | 180  | 
proof (cases dxs)  | 
181  | 
case (Abs_dlist xs)  | 
|
182  | 
then have dxs: "dxs = Dlist xs" and distinct: "distinct xs"  | 
|
183  | 
by (simp_all add: Dlist_def distinct_remdups_id)  | 
|
| 55913 | 184  | 
show thesis  | 
185  | 
proof (cases xs)  | 
|
186  | 
case Nil with dxs  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
187  | 
have "dxs = Dlist.empty" by (simp add: Dlist.empty_def)  | 
| 55913 | 188  | 
with empty show ?thesis .  | 
| 37106 | 189  | 
next  | 
190  | 
case (Cons x xs)  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
191  | 
with dxs distinct have "\<not> Dlist.member (Dlist xs) x"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
192  | 
and "dxs = Dlist.insert x (Dlist xs)"  | 
| 
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
193  | 
by (simp_all add: Dlist.member_def List.member_def Dlist.insert_def distinct_remdups_id)  | 
| 55913 | 194  | 
with insert show ?thesis .  | 
| 37106 | 195  | 
qed  | 
196  | 
qed  | 
|
197  | 
||
198  | 
||
| 60500 | 199  | 
subsection \<open>Functorial structure\<close>  | 
| 40603 | 200  | 
|
| 
55467
 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 
blanchet 
parents: 
49834 
diff
changeset
 | 
201  | 
functor map: map  | 
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
202  | 
by (simp_all add: remdups_map_remdups fun_eq_iff dlist_eq_iff)  | 
| 40603 | 203  | 
|
| 48282 | 204  | 
|
| 60500 | 205  | 
subsection \<open>Quickcheck generators\<close>  | 
| 
45927
 
e0305e4f02c9
adding quickcheck generator for distinct lists; adding examples
 
bulwahn 
parents: 
45694 
diff
changeset
 | 
206  | 
|
| 
61115
 
3a4400985780
modernized name space management -- more uniform qualification;
 
wenzelm 
parents: 
60679 
diff
changeset
 | 
207  | 
quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert  | 
| 35303 | 208  | 
|
| 62139 | 209  | 
subsection \<open>BNF instance\<close>  | 
210  | 
||
211  | 
context begin  | 
|
212  | 
||
213  | 
qualified fun wpull :: "('a \<times> 'b) list \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> ('a \<times> 'c) list"
 | 
|
214  | 
where  | 
|
215  | 
"wpull [] ys = []"  | 
|
216  | 
| "wpull xs [] = []"  | 
|
217  | 
| "wpull ((a, b) # xs) ((b', c) # ys) =  | 
|
218  | 
(if b \<in> snd ` set xs then  | 
|
219  | 
(a, the (map_of (rev ((b', c) # ys)) b)) # wpull xs ((b', c) # ys)  | 
|
220  | 
else if b' \<in> fst ` set ys then  | 
|
221  | 
(the (map_of (map prod.swap (rev ((a, b) # xs))) b'), c) # wpull ((a, b) # xs) ys  | 
|
222  | 
else (a, c) # wpull xs ys)"  | 
|
223  | 
||
224  | 
qualified lemma wpull_eq_Nil_iff [simp]: "wpull xs ys = [] \<longleftrightarrow> xs = [] \<or> ys = []"  | 
|
225  | 
by(cases "(xs, ys)" rule: wpull.cases) simp_all  | 
|
226  | 
||
227  | 
qualified lemma wpull_induct  | 
|
228  | 
[consumes 1,  | 
|
229  | 
case_names Nil left[xs eq in_set IH] right[xs ys eq in_set IH] step[xs ys eq IH] ]:  | 
|
230  | 
assumes eq: "remdups (map snd xs) = remdups (map fst ys)"  | 
|
231  | 
and Nil: "P [] []"  | 
|
232  | 
and left: "\<And>a b xs b' c ys.  | 
|
233  | 
\<lbrakk> b \<in> snd ` set xs; remdups (map snd xs) = remdups (map fst ((b', c) # ys));  | 
|
234  | 
(b, the (map_of (rev ((b', c) # ys)) b)) \<in> set ((b', c) # ys); P xs ((b', c) # ys) \<rbrakk>  | 
|
235  | 
\<Longrightarrow> P ((a, b) # xs) ((b', c) # ys)"  | 
|
236  | 
and right: "\<And>a b xs b' c ys.  | 
|
237  | 
\<lbrakk> b \<notin> snd ` set xs; b' \<in> fst ` set ys;  | 
|
238  | 
remdups (map snd ((a, b) # xs)) = remdups (map fst ys);  | 
|
239  | 
(the (map_of (map prod.swap (rev ((a, b) #xs))) b'), b') \<in> set ((a, b) # xs);  | 
|
240  | 
P ((a, b) # xs) ys \<rbrakk>  | 
|
241  | 
\<Longrightarrow> P ((a, b) # xs) ((b', c) # ys)"  | 
|
242  | 
and step: "\<And>a b xs c ys.  | 
|
243  | 
\<lbrakk> b \<notin> snd ` set xs; b \<notin> fst ` set ys; remdups (map snd xs) = remdups (map fst ys);  | 
|
244  | 
P xs ys \<rbrakk>  | 
|
245  | 
\<Longrightarrow> P ((a, b) # xs) ((b, c) # ys)"  | 
|
246  | 
shows "P xs ys"  | 
|
247  | 
using eq  | 
|
248  | 
proof(induction xs ys rule: wpull.induct)  | 
|
249  | 
case 1 thus ?case by(simp add: Nil)  | 
|
250  | 
next  | 
|
| 62390 | 251  | 
case 2 thus ?case by(simp split: if_split_asm)  | 
| 62139 | 252  | 
next  | 
253  | 
case Cons: (3 a b xs b' c ys)  | 
|
254  | 
let ?xs = "(a, b) # xs" and ?ys = "(b', c) # ys"  | 
|
255  | 
consider (xs) "b \<in> snd ` set xs" | (ys) "b \<notin> snd ` set xs" "b' \<in> fst ` set ys"  | 
|
256  | 
| (step) "b \<notin> snd ` set xs" "b' \<notin> fst ` set ys" by auto  | 
|
257  | 
thus ?case  | 
|
258  | 
proof cases  | 
|
259  | 
case xs  | 
|
260  | 
with Cons.prems have eq: "remdups (map snd xs) = remdups (map fst ?ys)" by auto  | 
|
261  | 
from xs eq have "b \<in> fst ` set ?ys" by (metis list.set_map set_remdups)  | 
|
262  | 
hence "map_of (rev ?ys) b \<noteq> None" unfolding map_of_eq_None_iff by auto  | 
|
263  | 
then obtain c' where "map_of (rev ?ys) b = Some c'" by blast  | 
|
| 62390 | 264  | 
then have "(b, the (map_of (rev ?ys) b)) \<in> set ?ys" by(auto dest: map_of_SomeD split: if_split_asm)  | 
| 62139 | 265  | 
from xs eq this Cons.IH(1)[OF xs eq] show ?thesis by(rule left)  | 
266  | 
next  | 
|
267  | 
case ys  | 
|
268  | 
from ys Cons.prems have eq: "remdups (map snd ?xs) = remdups (map fst ys)" by auto  | 
|
269  | 
from ys eq have "b' \<in> snd ` set ?xs" by (metis list.set_map set_remdups)  | 
|
270  | 
hence "map_of (map prod.swap (rev ?xs)) b' \<noteq> None"  | 
|
271  | 
unfolding map_of_eq_None_iff by(auto simp add: image_image)  | 
|
272  | 
then obtain a' where "map_of (map prod.swap (rev ?xs)) b' = Some a'" by blast  | 
|
273  | 
then have "(the (map_of (map prod.swap (rev ?xs)) b'), b') \<in> set ?xs"  | 
|
| 62390 | 274  | 
by(auto dest: map_of_SomeD split: if_split_asm)  | 
| 62139 | 275  | 
from ys eq this Cons.IH(2)[OF ys eq] show ?thesis by(rule right)  | 
276  | 
next  | 
|
277  | 
case *: step  | 
|
278  | 
hence "remdups (map snd xs) = remdups (map fst ys)" "b = b'" using Cons.prems by auto  | 
|
279  | 
from * this(1) Cons.IH(3)[OF * this(1)] show ?thesis unfolding \<open>b = b'\<close> by(rule step)  | 
|
280  | 
qed  | 
|
281  | 
qed  | 
|
282  | 
||
283  | 
qualified lemma set_wpull_subset:  | 
|
284  | 
assumes "remdups (map snd xs) = remdups (map fst ys)"  | 
|
285  | 
shows "set (wpull xs ys) \<subseteq> set xs O set ys"  | 
|
286  | 
using assms by(induction xs ys rule: wpull_induct) auto  | 
|
287  | 
||
288  | 
qualified lemma set_fst_wpull:  | 
|
289  | 
assumes "remdups (map snd xs) = remdups (map fst ys)"  | 
|
290  | 
shows "fst ` set (wpull xs ys) = fst ` set xs"  | 
|
291  | 
using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI)  | 
|
292  | 
||
293  | 
qualified lemma set_snd_wpull:  | 
|
294  | 
assumes "remdups (map snd xs) = remdups (map fst ys)"  | 
|
295  | 
shows "snd ` set (wpull xs ys) = snd ` set ys"  | 
|
296  | 
using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI)  | 
|
297  | 
||
298  | 
qualified lemma wpull:  | 
|
299  | 
assumes "distinct xs"  | 
|
300  | 
and "distinct ys"  | 
|
301  | 
  and "set xs \<subseteq> {(x, y). R x y}"
 | 
|
302  | 
  and "set ys \<subseteq> {(x, y). S x y}"
 | 
|
303  | 
and eq: "remdups (map snd xs) = remdups (map fst ys)"  | 
|
304  | 
  shows "\<exists>zs. distinct zs \<and> set zs \<subseteq> {(x, y). (R OO S) x y} \<and>
 | 
|
305  | 
remdups (map fst zs) = remdups (map fst xs) \<and> remdups (map snd zs) = remdups (map snd ys)"  | 
|
306  | 
proof(intro exI conjI)  | 
|
307  | 
let ?zs = "remdups (wpull xs ys)"  | 
|
308  | 
show "distinct ?zs" by simp  | 
|
309  | 
  show "set ?zs \<subseteq> {(x, y). (R OO S) x y}" using assms(3-4) set_wpull_subset[OF eq] by fastforce
 | 
|
310  | 
show "remdups (map fst ?zs) = remdups (map fst xs)" unfolding remdups_map_remdups using eq  | 
|
311  | 
by(induction xs ys rule: wpull_induct)(auto simp add: set_fst_wpull intro: rev_image_eqI)  | 
|
312  | 
show "remdups (map snd ?zs) = remdups (map snd ys)" unfolding remdups_map_remdups using eq  | 
|
313  | 
by(induction xs ys rule: wpull_induct)(auto simp add: set_snd_wpull intro: rev_image_eqI)  | 
|
314  | 
qed  | 
|
315  | 
||
316  | 
qualified lift_definition set :: "'a dlist \<Rightarrow> 'a set" is List.set .  | 
|
317  | 
||
318  | 
qualified lemma map_transfer [transfer_rule]:  | 
|
319  | 
"(rel_fun op = (rel_fun (pcr_dlist op =) (pcr_dlist op =))) (\<lambda>f x. remdups (List.map f x)) Dlist.map"  | 
|
320  | 
by(simp add: rel_fun_def dlist.pcr_cr_eq cr_dlist_def Dlist.map_def remdups_remdups)  | 
|
321  | 
||
322  | 
bnf "'a dlist"  | 
|
323  | 
map: Dlist.map  | 
|
324  | 
sets: set  | 
|
325  | 
bd: natLeq  | 
|
326  | 
wits: Dlist.empty  | 
|
327  | 
unfolding OO_Grp_alt mem_Collect_eq  | 
|
328  | 
subgoal by(rule ext)(simp add: dlist_eq_iff)  | 
|
329  | 
subgoal by(rule ext)(simp add: dlist_eq_iff remdups_map_remdups)  | 
|
330  | 
subgoal by(simp add: dlist_eq_iff set_def cong: list.map_cong)  | 
|
331  | 
subgoal by(simp add: set_def fun_eq_iff)  | 
|
332  | 
subgoal by(simp add: natLeq_card_order)  | 
|
333  | 
subgoal by(simp add: natLeq_cinfinite)  | 
|
334  | 
subgoal by(rule ordLess_imp_ordLeq)(simp add: finite_iff_ordLess_natLeq[symmetric] set_def)  | 
|
335  | 
subgoal by(rule predicate2I)(transfer; auto simp add: wpull)  | 
|
336  | 
subgoal by(simp add: set_def)  | 
|
337  | 
done  | 
|
338  | 
||
339  | 
lifting_update dlist.lifting  | 
|
340  | 
lifting_forget dlist.lifting  | 
|
341  | 
||
| 35303 | 342  | 
end  | 
| 62139 | 343  | 
|
| 62390 | 344  | 
end  |