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