| author | wenzelm | 
| Sat, 03 Oct 2020 14:06:00 +0200 | |
| changeset 72367 | d3069e7e1175 | 
| parent 71494 | cbe0b6b0bed8 | 
| child 73398 | 180981b87929 | 
| permissions | -rw-r--r-- | 
| 71494 | 1 | (* Author: Florian Haftmann, TU Muenchen | 
| 62139 | 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 | |
| 71494 | 7 | imports Confluent_Quotient | 
| 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 | ||
| 71494 | 18 | context begin | 
| 19 | ||
| 20 | qualified definition dlist_eq where "dlist_eq = BNF_Def.vimage2p remdups remdups (=)" | |
| 21 | ||
| 22 | qualified lemma equivp_dlist_eq: "equivp dlist_eq" | |
| 23 | unfolding dlist_eq_def by(rule equivp_vimage2p)(rule identity_equivp) | |
| 24 | ||
| 25 | qualified definition abs_dlist :: "'a list \<Rightarrow> 'a dlist" where "abs_dlist = Abs_dlist o remdups" | |
| 26 | ||
| 27 | definition qcr_dlist :: "'a list \<Rightarrow> 'a dlist \<Rightarrow> bool" where "qcr_dlist x y \<longleftrightarrow> y = abs_dlist x" | |
| 28 | ||
| 29 | qualified lemma Quotient_dlist_remdups: "Quotient dlist_eq abs_dlist list_of_dlist qcr_dlist" | |
| 30 | unfolding Quotient_def dlist_eq_def qcr_dlist_def vimage2p_def abs_dlist_def | |
| 31 | by (auto simp add: fun_eq_iff Abs_dlist_inject | |
| 32 | list_of_dlist[simplified] list_of_dlist_inverse distinct_remdups_id) | |
| 33 | ||
| 34 | end | |
| 35 | ||
| 36 | locale Quotient_dlist begin | |
| 37 | setup_lifting Dlist.Quotient_dlist_remdups Dlist.equivp_dlist_eq[THEN equivp_reflp2] | |
| 38 | end | |
| 39 | ||
| 62139 | 40 | setup_lifting type_definition_dlist | 
| 41 | ||
| 39380 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 42 | lemma dlist_eq_iff: | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 43 | "dxs = dys \<longleftrightarrow> list_of_dlist dxs = list_of_dlist dys" | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 44 | by (simp add: list_of_dlist_inject) | 
| 36274 | 45 | |
| 39380 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 46 | lemma dlist_eqI: | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 47 | "list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys" | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 48 | by (simp add: dlist_eq_iff) | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35688diff
changeset | 49 | |
| 69593 | 50 | text \<open>Formal, totalized constructor for \<^typ>\<open>'a dlist\<close>:\<close> | 
| 35303 | 51 | |
| 52 | definition Dlist :: "'a list \<Rightarrow> 'a dlist" where | |
| 37765 | 53 | "Dlist xs = Abs_dlist (remdups xs)" | 
| 35303 | 54 | |
| 39380 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 55 | lemma distinct_list_of_dlist [simp, intro]: | 
| 35303 | 56 | "distinct (list_of_dlist dxs)" | 
| 57 | using list_of_dlist [of dxs] by simp | |
| 58 | ||
| 59 | lemma list_of_dlist_Dlist [simp]: | |
| 60 | "list_of_dlist (Dlist xs) = remdups xs" | |
| 61 | by (simp add: Dlist_def Abs_dlist_inverse) | |
| 62 | ||
| 39727 | 63 | lemma remdups_list_of_dlist [simp]: | 
| 64 | "remdups (list_of_dlist dxs) = list_of_dlist dxs" | |
| 65 | by simp | |
| 66 | ||
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35688diff
changeset | 67 | lemma Dlist_list_of_dlist [simp, code abstype]: | 
| 35303 | 68 | "Dlist (list_of_dlist dxs) = dxs" | 
| 69 | by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id) | |
| 70 | ||
| 71 | ||
| 60500 | 72 | text \<open>Fundamental operations:\<close> | 
| 35303 | 73 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 74 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 75 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 76 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 77 | qualified definition empty :: "'a dlist" where | 
| 35303 | 78 | "empty = Dlist []" | 
| 79 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 80 | qualified definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where | 
| 35303 | 81 | "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))" | 
| 82 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 83 | qualified definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where | 
| 35303 | 84 | "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))" | 
| 85 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 86 | qualified definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
 | 
| 35303 | 87 | "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))" | 
| 88 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 89 | qualified definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
 | 
| 35303 | 90 | "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))" | 
| 91 | ||
| 63375 
59803048b0e8
basic facts about almost everywhere fix bijections
 haftmann parents: 
62390diff
changeset | 92 | qualified definition rotate :: "nat \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where | 
| 
59803048b0e8
basic facts about almost everywhere fix bijections
 haftmann parents: 
62390diff
changeset | 93 | "rotate n dxs = Dlist (List.rotate n (list_of_dlist dxs))" | 
| 
59803048b0e8
basic facts about almost everywhere fix bijections
 haftmann parents: 
62390diff
changeset | 94 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 95 | end | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 96 | |
| 35303 | 97 | |
| 60500 | 98 | text \<open>Derived operations:\<close> | 
| 35303 | 99 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 100 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 101 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 102 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 103 | qualified definition null :: "'a dlist \<Rightarrow> bool" where | 
| 35303 | 104 | "null dxs = List.null (list_of_dlist dxs)" | 
| 105 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 106 | qualified definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 35303 | 107 | "member dxs = List.member (list_of_dlist dxs)" | 
| 108 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 109 | qualified definition length :: "'a dlist \<Rightarrow> nat" where | 
| 35303 | 110 | "length dxs = List.length (list_of_dlist dxs)" | 
| 111 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 112 | 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 | 113 | "fold f dxs = List.fold f (list_of_dlist dxs)" | 
| 37022 
f9681d9d1d56
moved generic List operations to theory More_List
 haftmann parents: 
36980diff
changeset | 114 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 115 | 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 | 116 | "foldr f dxs = List.foldr f (list_of_dlist dxs)" | 
| 35303 | 117 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 118 | end | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 119 | |
| 35303 | 120 | |
| 60500 | 121 | subsection \<open>Executable version obeying invariant\<close> | 
| 35303 | 122 | |
| 123 | lemma list_of_dlist_empty [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 124 | "list_of_dlist Dlist.empty = []" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 125 | by (simp add: Dlist.empty_def) | 
| 35303 | 126 | |
| 127 | lemma list_of_dlist_insert [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 128 | "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 | 129 | by (simp add: Dlist.insert_def) | 
| 35303 | 130 | |
| 131 | lemma list_of_dlist_remove [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 132 | "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 | 133 | by (simp add: Dlist.remove_def) | 
| 35303 | 134 | |
| 135 | lemma list_of_dlist_map [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 136 | "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 | 137 | by (simp add: Dlist.map_def) | 
| 35303 | 138 | |
| 139 | lemma list_of_dlist_filter [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 140 | "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 | 141 | by (simp add: Dlist.filter_def) | 
| 35303 | 142 | |
| 63375 
59803048b0e8
basic facts about almost everywhere fix bijections
 haftmann parents: 
62390diff
changeset | 143 | lemma list_of_dlist_rotate [simp, code abstract]: | 
| 
59803048b0e8
basic facts about almost everywhere fix bijections
 haftmann parents: 
62390diff
changeset | 144 | "list_of_dlist (Dlist.rotate n dxs) = List.rotate n (list_of_dlist dxs)" | 
| 
59803048b0e8
basic facts about almost everywhere fix bijections
 haftmann parents: 
62390diff
changeset | 145 | by (simp add: Dlist.rotate_def) | 
| 
59803048b0e8
basic facts about almost everywhere fix bijections
 haftmann parents: 
62390diff
changeset | 146 | |
| 35303 | 147 | |
| 60500 | 148 | 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 | 149 | |
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 150 | 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 | 151 | "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 | 152 | |
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 153 | 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 | 154 | "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 | 155 | by simp | 
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 156 | |
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 157 | |
| 60500 | 158 | text \<open>Equality\<close> | 
| 38512 | 159 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38512diff
changeset | 160 | instantiation dlist :: (equal) equal | 
| 38512 | 161 | begin | 
| 162 | ||
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38512diff
changeset | 163 | definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)" | 
| 38512 | 164 | |
| 60679 | 165 | instance | 
| 166 | by standard (simp add: equal_dlist_def equal list_of_dlist_inject) | |
| 38512 | 167 | |
| 168 | end | |
| 169 | ||
| 43764 | 170 | declare equal_dlist_def [code] | 
| 171 | ||
| 60679 | 172 | 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 | 173 | by (fact equal_refl) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38512diff
changeset | 174 | |
| 38512 | 175 | |
| 60500 | 176 | subsection \<open>Induction principle and case distinction\<close> | 
| 37106 | 177 | |
| 178 | lemma dlist_induct [case_names empty insert, induct type: dlist]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 179 | assumes empty: "P Dlist.empty" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 180 | assumes insrt: "\<And>x dxs. \<not> Dlist.member dxs x \<Longrightarrow> P dxs \<Longrightarrow> P (Dlist.insert x dxs)" | 
| 37106 | 181 | shows "P dxs" | 
| 182 | proof (cases dxs) | |
| 183 | case (Abs_dlist xs) | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 184 | then have "distinct xs" and dxs: "dxs = Dlist xs" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 185 | by (simp_all add: Dlist_def distinct_remdups_id) | 
| 60500 | 186 | 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 | 187 | proof (induct xs) | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 188 | case Nil from empty show ?case by (simp add: Dlist.empty_def) | 
| 37106 | 189 | next | 
| 40122 
1d8ad2ff3e01
dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
 haftmann parents: 
39915diff
changeset | 190 | case (Cons x xs) | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 191 | then have "\<not> Dlist.member (Dlist xs) x" and "P (Dlist xs)" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 192 | by (simp_all add: Dlist.member_def List.member_def) | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 193 | with insrt have "P (Dlist.insert x (Dlist xs))" . | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 194 | with Cons show ?case by (simp add: Dlist.insert_def distinct_remdups_id) | 
| 37106 | 195 | qed | 
| 196 | with dxs show "P dxs" by simp | |
| 197 | qed | |
| 198 | ||
| 55913 | 199 | lemma dlist_case [cases type: dlist]: | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 200 | obtains (empty) "dxs = Dlist.empty" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 201 | | (insert) x dys where "\<not> Dlist.member dys x" and "dxs = Dlist.insert x dys" | 
| 37106 | 202 | proof (cases dxs) | 
| 203 | case (Abs_dlist xs) | |
| 204 | then have dxs: "dxs = Dlist xs" and distinct: "distinct xs" | |
| 205 | by (simp_all add: Dlist_def distinct_remdups_id) | |
| 55913 | 206 | show thesis | 
| 207 | proof (cases xs) | |
| 208 | case Nil with dxs | |
| 71494 | 209 | have "dxs = Dlist.empty" by (simp add: Dlist.empty_def) | 
| 55913 | 210 | with empty show ?thesis . | 
| 37106 | 211 | next | 
| 212 | case (Cons x xs) | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 213 | with dxs distinct have "\<not> Dlist.member (Dlist xs) x" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 214 | and "dxs = Dlist.insert x (Dlist xs)" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 215 | by (simp_all add: Dlist.member_def List.member_def Dlist.insert_def distinct_remdups_id) | 
| 55913 | 216 | with insert show ?thesis . | 
| 37106 | 217 | qed | 
| 218 | qed | |
| 219 | ||
| 220 | ||
| 60500 | 221 | subsection \<open>Functorial structure\<close> | 
| 40603 | 222 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
49834diff
changeset | 223 | functor map: map | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 224 | by (simp_all add: remdups_map_remdups fun_eq_iff dlist_eq_iff) | 
| 40603 | 225 | |
| 48282 | 226 | |
| 60500 | 227 | subsection \<open>Quickcheck generators\<close> | 
| 45927 
e0305e4f02c9
adding quickcheck generator for distinct lists; adding examples
 bulwahn parents: 
45694diff
changeset | 228 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 229 | quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert | 
| 35303 | 230 | |
| 62139 | 231 | subsection \<open>BNF instance\<close> | 
| 232 | ||
| 233 | context begin | |
| 234 | ||
| 71494 | 235 | qualified inductive double :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where | 
| 236 | "double (xs @ ys) (xs @ x # ys)" if "x \<in> set ys" | |
| 62139 | 237 | |
| 71494 | 238 | qualified lemma strong_confluentp_double: "strong_confluentp double" | 
| 239 | proof | |
| 240 | fix xs ys zs :: "'a list" | |
| 241 | assume ys: "double xs ys" and zs: "double xs zs" | |
| 242 | consider (left) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ y # bs @ cs" "zs = as @ bs @ z # cs" "y \<in> set (bs @ cs)" "z \<in> set cs" | |
| 243 | | (right) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ bs @ y # cs" "zs = as @ z # bs @ cs" "y \<in> set cs" "z \<in> set (bs @ cs)" | |
| 244 | proof - | |
| 245 | show thesis using ys zs | |
| 246 | by(clarsimp simp add: double.simps append_eq_append_conv2)(auto intro: that) | |
| 247 | qed | |
| 248 | then show "\<exists>us. double\<^sup>*\<^sup>* ys us \<and> double\<^sup>=\<^sup>= zs us" | |
| 62139 | 249 | proof cases | 
| 71494 | 250 | case left | 
| 251 | let ?us = "as @ y # bs @ z # cs" | |
| 252 | have "double ys ?us" "double zs ?us" using left | |
| 253 | by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+ | |
| 254 | then show ?thesis by blast | |
| 62139 | 255 | next | 
| 71494 | 256 | case right | 
| 257 | let ?us = "as @ z # bs @ y # cs" | |
| 258 | have "double ys ?us" "double zs ?us" using right | |
| 259 | by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+ | |
| 260 | then show ?thesis by blast | |
| 62139 | 261 | qed | 
| 262 | qed | |
| 263 | ||
| 71494 | 264 | qualified lemma double_Cons1 [simp]: "double xs (x # xs)" if "x \<in> set xs" | 
| 265 | using double.intros[of x xs "[]"] that by simp | |
| 266 | ||
| 267 | qualified lemma double_Cons_same [simp]: "double xs ys \<Longrightarrow> double (x # xs) (x # ys)" | |
| 268 | by(auto simp add: double.simps Cons_eq_append_conv) | |
| 62139 | 269 | |
| 71494 | 270 | qualified lemma doubles_Cons_same: "double\<^sup>*\<^sup>* xs ys \<Longrightarrow> double\<^sup>*\<^sup>* (x # xs) (x # ys)" | 
| 271 | by(induction rule: rtranclp_induct)(auto intro: rtranclp.rtrancl_into_rtrancl) | |
| 272 | ||
| 273 | qualified lemma remdups_into_doubles: "double\<^sup>*\<^sup>* (remdups xs) xs" | |
| 274 | by(induction xs)(auto intro: doubles_Cons_same rtranclp.rtrancl_into_rtrancl) | |
| 275 | ||
| 276 | qualified lemma dlist_eq_into_doubles: "Dlist.dlist_eq \<le> equivclp double" | |
| 277 | by(auto 4 4 simp add: Dlist.dlist_eq_def vimage2p_def | |
| 278 | intro: equivclp_trans converse_rtranclp_into_equivclp rtranclp_into_equivclp remdups_into_doubles) | |
| 62139 | 279 | |
| 71494 | 280 | qualified lemma factor_double_map: "double (map f xs) ys \<Longrightarrow> \<exists>zs. Dlist.dlist_eq xs zs \<and> ys = map f zs" | 
| 281 | by(auto simp add: double.simps Dlist.dlist_eq_def vimage2p_def map_eq_append_conv) | |
| 282 | (metis list.simps(9) map_append remdups.simps(2) remdups_append2) | |
| 283 | ||
| 284 | qualified lemma dlist_eq_set_eq: "Dlist.dlist_eq xs ys \<Longrightarrow> set xs = set ys" | |
| 285 | by(simp add: Dlist.dlist_eq_def vimage2p_def)(metis set_remdups) | |
| 62139 | 286 | |
| 71494 | 287 | qualified lemma dlist_eq_map_respect: "Dlist.dlist_eq xs ys \<Longrightarrow> Dlist.dlist_eq (map f xs) (map f ys)" | 
| 288 | by(clarsimp simp add: Dlist.dlist_eq_def vimage2p_def)(metis remdups_map_remdups) | |
| 62139 | 289 | |
| 71494 | 290 | qualified lemma confluent_quotient_dlist: | 
| 291 | "confluent_quotient double Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq | |
| 292 | (map fst) (map snd) (map fst) (map snd) list_all2 list_all2 list_all2 set set" | |
| 293 | by(unfold_locales)(auto intro: strong_confluentp_imp_confluentp strong_confluentp_double | |
| 294 | dest: factor_double_map dlist_eq_into_doubles[THEN predicate2D] dlist_eq_set_eq | |
| 295 | simp add: list.in_rel list.rel_compp dlist_eq_map_respect Dlist.equivp_dlist_eq equivp_imp_transp) | |
| 296 | ||
| 62139 | 297 | |
| 298 | lifting_update dlist.lifting | |
| 299 | lifting_forget dlist.lifting | |
| 300 | ||
| 35303 | 301 | end | 
| 62139 | 302 | |
| 71494 | 303 | context begin | 
| 304 | interpretation Quotient_dlist: Quotient_dlist . | |
| 305 | ||
| 306 | lift_bnf (plugins del: code) 'a dlist | |
| 307 | subgoal for A B by(rule confluent_quotient.subdistributivity[OF Dlist.confluent_quotient_dlist]) | |
| 308 | subgoal by(force dest: Dlist.dlist_eq_set_eq intro: equivp_reflp[OF Dlist.equivp_dlist_eq]) | |
| 309 | done | |
| 310 | ||
| 311 | qualified lemma list_of_dlist_transfer[transfer_rule]: | |
| 312 | "bi_unique R \<Longrightarrow> (rel_fun (Quotient_dlist.pcr_dlist R) (list_all2 R)) remdups list_of_dlist" | |
| 313 | unfolding rel_fun_def Quotient_dlist.pcr_dlist_def qcr_dlist_def Dlist.abs_dlist_def | |
| 314 | by (auto simp: Abs_dlist_inverse intro!: remdups_transfer[THEN rel_funD]) | |
| 315 | ||
| 316 | lemma list_of_dlist_map_dlist[simp]: | |
| 317 | "list_of_dlist (map_dlist f xs) = remdups (map f (list_of_dlist xs))" | |
| 318 | by transfer (auto simp: remdups_map_remdups) | |
| 319 | ||
| 62390 | 320 | end | 
| 71494 | 321 | |
| 322 | ||
| 323 | end |