| author | paulson | 
| Mon, 11 Jan 2016 22:14:15 +0000 | |
| changeset 62131 | 1baed43f453e | 
| parent 61115 | 3a4400985780 | 
| child 62139 | 519362f817c7 | 
| permissions | -rw-r--r-- | 
| 35303 | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 2 | ||
| 60500 | 3 | section \<open>Lists with elements distinct as canonical example for datatype invariants\<close> | 
| 35303 | 4 | |
| 5 | theory Dlist | |
| 45990 
b7b905b23b2a
incorporated More_Set and More_List into the Main body -- to be consolidated later
 haftmann parents: 
45927diff
changeset | 6 | imports Main | 
| 35303 | 7 | begin | 
| 8 | ||
| 60500 | 9 | subsection \<open>The type of distinct lists\<close> | 
| 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: 
43764diff
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: 
38857diff
changeset | 17 | lemma dlist_eq_iff: | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 18 | "dxs = dys \<longleftrightarrow> list_of_dlist dxs = list_of_dlist dys" | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 19 | by (simp add: list_of_dlist_inject) | 
| 36274 | 20 | |
| 39380 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 21 | lemma dlist_eqI: | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 22 | "list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys" | 
| 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
38857diff
changeset | 23 | by (simp add: dlist_eq_iff) | 
| 36112 
7fa17a225852
user interface for abstract datatypes is an attribute, not a command
 haftmann parents: 
35688diff
changeset | 24 | |
| 60500 | 25 | text \<open>Formal, totalized constructor for @{typ "'a dlist"}:\<close>
 | 
| 35303 | 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: 
38857diff
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: 
35688diff
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 | ||
| 60500 | 47 | text \<open>Fundamental operations:\<close> | 
| 35303 | 48 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 49 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 50 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 51 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 52 | qualified definition empty :: "'a dlist" where | 
| 35303 | 53 | "empty = Dlist []" | 
| 54 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 55 | qualified definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where | 
| 35303 | 56 | "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))" | 
| 57 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 58 | qualified definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where | 
| 35303 | 59 | "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))" | 
| 60 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 61 | qualified definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
 | 
| 35303 | 62 | "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))" | 
| 63 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 64 | qualified definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
 | 
| 35303 | 65 | "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))" | 
| 66 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 67 | end | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 68 | |
| 35303 | 69 | |
| 60500 | 70 | text \<open>Derived operations:\<close> | 
| 35303 | 71 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 72 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 73 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 74 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 75 | qualified definition null :: "'a dlist \<Rightarrow> bool" where | 
| 35303 | 76 | "null dxs = List.null (list_of_dlist dxs)" | 
| 77 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 78 | qualified definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 35303 | 79 | "member dxs = List.member (list_of_dlist dxs)" | 
| 80 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 81 | qualified definition length :: "'a dlist \<Rightarrow> nat" where | 
| 35303 | 82 | "length dxs = List.length (list_of_dlist dxs)" | 
| 83 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 84 | 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 | 85 | "fold f dxs = List.fold f (list_of_dlist dxs)" | 
| 37022 
f9681d9d1d56
moved generic List operations to theory More_List
 haftmann parents: 
36980diff
changeset | 86 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 87 | 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 | 88 | "foldr f dxs = List.foldr f (list_of_dlist dxs)" | 
| 35303 | 89 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 90 | end | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 91 | |
| 35303 | 92 | |
| 60500 | 93 | subsection \<open>Executable version obeying invariant\<close> | 
| 35303 | 94 | |
| 95 | lemma list_of_dlist_empty [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 96 | "list_of_dlist Dlist.empty = []" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 97 | by (simp add: Dlist.empty_def) | 
| 35303 | 98 | |
| 99 | lemma list_of_dlist_insert [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 100 | "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 | 101 | by (simp add: Dlist.insert_def) | 
| 35303 | 102 | |
| 103 | lemma list_of_dlist_remove [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 104 | "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 | 105 | by (simp add: Dlist.remove_def) | 
| 35303 | 106 | |
| 107 | lemma list_of_dlist_map [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 108 | "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 | 109 | by (simp add: Dlist.map_def) | 
| 35303 | 110 | |
| 111 | lemma list_of_dlist_filter [simp, code abstract]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 112 | "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 | 113 | by (simp add: Dlist.filter_def) | 
| 35303 | 114 | |
| 115 | ||
| 60500 | 116 | 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 | 117 | |
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 118 | 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 | 119 | "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 | 120 | |
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 121 | 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 | 122 | "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 | 123 | by simp | 
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 124 | |
| 
1a4cc941171a
added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
 haftmann parents: 
36274diff
changeset | 125 | |
| 60500 | 126 | text \<open>Equality\<close> | 
| 38512 | 127 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38512diff
changeset | 128 | instantiation dlist :: (equal) equal | 
| 38512 | 129 | begin | 
| 130 | ||
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38512diff
changeset | 131 | definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)" | 
| 38512 | 132 | |
| 60679 | 133 | instance | 
| 134 | by standard (simp add: equal_dlist_def equal list_of_dlist_inject) | |
| 38512 | 135 | |
| 136 | end | |
| 137 | ||
| 43764 | 138 | declare equal_dlist_def [code] | 
| 139 | ||
| 60679 | 140 | 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 | 141 | by (fact equal_refl) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38512diff
changeset | 142 | |
| 38512 | 143 | |
| 60500 | 144 | subsection \<open>Induction principle and case distinction\<close> | 
| 37106 | 145 | |
| 146 | lemma dlist_induct [case_names empty insert, induct type: dlist]: | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 147 | assumes empty: "P Dlist.empty" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 148 | assumes insrt: "\<And>x dxs. \<not> Dlist.member dxs x \<Longrightarrow> P dxs \<Longrightarrow> P (Dlist.insert x dxs)" | 
| 37106 | 149 | shows "P dxs" | 
| 150 | proof (cases dxs) | |
| 151 | case (Abs_dlist xs) | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 152 | then have "distinct xs" and dxs: "dxs = Dlist xs" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 153 | by (simp_all add: Dlist_def distinct_remdups_id) | 
| 60500 | 154 | 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 | 155 | proof (induct xs) | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 156 | case Nil from empty show ?case by (simp add: Dlist.empty_def) | 
| 37106 | 157 | next | 
| 40122 
1d8ad2ff3e01
dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
 haftmann parents: 
39915diff
changeset | 158 | case (Cons x xs) | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 159 | then have "\<not> Dlist.member (Dlist xs) x" and "P (Dlist xs)" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 160 | by (simp_all add: Dlist.member_def List.member_def) | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 161 | with insrt have "P (Dlist.insert x (Dlist xs))" . | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 162 | with Cons show ?case by (simp add: Dlist.insert_def distinct_remdups_id) | 
| 37106 | 163 | qed | 
| 164 | with dxs show "P dxs" by simp | |
| 165 | qed | |
| 166 | ||
| 55913 | 167 | lemma dlist_case [cases type: dlist]: | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 168 | obtains (empty) "dxs = Dlist.empty" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 169 | | (insert) x dys where "\<not> Dlist.member dys x" and "dxs = Dlist.insert x dys" | 
| 37106 | 170 | proof (cases dxs) | 
| 171 | case (Abs_dlist xs) | |
| 172 | then have dxs: "dxs = Dlist xs" and distinct: "distinct xs" | |
| 173 | by (simp_all add: Dlist_def distinct_remdups_id) | |
| 55913 | 174 | show thesis | 
| 175 | proof (cases xs) | |
| 176 | case Nil with dxs | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 177 | have "dxs = Dlist.empty" by (simp add: Dlist.empty_def) | 
| 55913 | 178 | with empty show ?thesis . | 
| 37106 | 179 | next | 
| 180 | case (Cons x xs) | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 181 | with dxs distinct have "\<not> Dlist.member (Dlist xs) x" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 182 | and "dxs = Dlist.insert x (Dlist xs)" | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 183 | by (simp_all add: Dlist.member_def List.member_def Dlist.insert_def distinct_remdups_id) | 
| 55913 | 184 | with insert show ?thesis . | 
| 37106 | 185 | qed | 
| 186 | qed | |
| 187 | ||
| 188 | ||
| 60500 | 189 | subsection \<open>Functorial structure\<close> | 
| 40603 | 190 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
49834diff
changeset | 191 | functor map: map | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 192 | by (simp_all add: remdups_map_remdups fun_eq_iff dlist_eq_iff) | 
| 40603 | 193 | |
| 48282 | 194 | |
| 60500 | 195 | subsection \<open>Quickcheck generators\<close> | 
| 45927 
e0305e4f02c9
adding quickcheck generator for distinct lists; adding examples
 bulwahn parents: 
45694diff
changeset | 196 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60679diff
changeset | 197 | quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert | 
| 35303 | 198 | |
| 199 | end |