| author | nipkow | 
| Wed, 03 Jul 2024 19:42:13 +0200 | |
| changeset 80484 | 0ca36dcdcbd3 | 
| parent 69250 | 1011f0b46af7 | 
| child 82388 | f1ff9123c62a | 
| permissions | -rw-r--r-- | 
| 69194 | 1 | (* Title: HOL/Library/Sorting_Algorithms.thy | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | theory Sorting_Algorithms | |
| 6 | imports Main Multiset Comparator | |
| 7 | begin | |
| 8 | ||
| 9 | section \<open>Stably sorted lists\<close> | |
| 10 | ||
| 11 | abbreviation (input) stable_segment :: "'a comparator \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" | |
| 12 | where "stable_segment cmp x \<equiv> filter (\<lambda>y. compare cmp x y = Equiv)" | |
| 13 | ||
| 14 | fun sorted :: "'a comparator \<Rightarrow> 'a list \<Rightarrow> bool" | |
| 15 | where sorted_Nil: "sorted cmp [] \<longleftrightarrow> True" | |
| 16 | | sorted_single: "sorted cmp [x] \<longleftrightarrow> True" | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 17 | | sorted_rec: "sorted cmp (y # x # xs) \<longleftrightarrow> compare cmp y x \<noteq> Greater \<and> sorted cmp (x # xs)" | 
| 69194 | 18 | |
| 19 | lemma sorted_ConsI: | |
| 20 | "sorted cmp (x # xs)" if "sorted cmp xs" | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 21 | and "\<And>y ys. xs = y # ys \<Longrightarrow> compare cmp x y \<noteq> Greater" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 22 | using that by (cases xs) simp_all | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 23 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 24 | lemma sorted_Cons_imp_sorted: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 25 | "sorted cmp xs" if "sorted cmp (x # xs)" | 
| 69194 | 26 | using that by (cases xs) simp_all | 
| 27 | ||
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 28 | lemma sorted_Cons_imp_not_less: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 29 | "compare cmp y x \<noteq> Greater" if "sorted cmp (y # xs)" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 30 | and "x \<in> set xs" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 31 | using that by (induction xs arbitrary: y) (auto dest: compare.trans_not_greater) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 32 | |
| 69194 | 33 | lemma sorted_induct [consumes 1, case_names Nil Cons, induct pred: sorted]: | 
| 34 | "P xs" if "sorted cmp xs" and "P []" | |
| 35 | and *: "\<And>x xs. sorted cmp xs \<Longrightarrow> P xs | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 36 | \<Longrightarrow> (\<And>y. y \<in> set xs \<Longrightarrow> compare cmp x y \<noteq> Greater) \<Longrightarrow> P (x # xs)" | 
| 69194 | 37 | using \<open>sorted cmp xs\<close> proof (induction xs) | 
| 38 | case Nil | |
| 39 | show ?case | |
| 40 | by (rule \<open>P []\<close>) | |
| 41 | next | |
| 42 | case (Cons x xs) | |
| 43 | from \<open>sorted cmp (x # xs)\<close> have "sorted cmp xs" | |
| 44 | by (cases xs) simp_all | |
| 45 | moreover have "P xs" using \<open>sorted cmp xs\<close> | |
| 46 | by (rule Cons.IH) | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 47 | moreover have "compare cmp x y \<noteq> Greater" if "y \<in> set xs" for y | 
| 69194 | 48 | using that \<open>sorted cmp (x # xs)\<close> proof (induction xs) | 
| 49 | case Nil | |
| 50 | then show ?case | |
| 51 | by simp | |
| 52 | next | |
| 53 | case (Cons z zs) | |
| 54 | then show ?case | |
| 55 | proof (cases zs) | |
| 56 | case Nil | |
| 57 | with Cons.prems show ?thesis | |
| 58 | by simp | |
| 59 | next | |
| 60 | case (Cons w ws) | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 61 | with Cons.prems have "compare cmp z w \<noteq> Greater" "compare cmp x z \<noteq> Greater" | 
| 69194 | 62 | by auto | 
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 63 | then have "compare cmp x w \<noteq> Greater" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 64 | by (auto dest: compare.trans_not_greater) | 
| 69194 | 65 | with Cons show ?thesis | 
| 66 | using Cons.prems Cons.IH by auto | |
| 67 | qed | |
| 68 | qed | |
| 69 | ultimately show ?case | |
| 70 | by (rule *) | |
| 71 | qed | |
| 72 | ||
| 73 | lemma sorted_induct_remove1 [consumes 1, case_names Nil minimum]: | |
| 74 | "P xs" if "sorted cmp xs" and "P []" | |
| 75 | and *: "\<And>x xs. sorted cmp xs \<Longrightarrow> P (remove1 x xs) | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 76 | \<Longrightarrow> x \<in> set xs \<Longrightarrow> hd (stable_segment cmp x xs) = x \<Longrightarrow> (\<And>y. y \<in> set xs \<Longrightarrow> compare cmp x y \<noteq> Greater) | 
| 69194 | 77 | \<Longrightarrow> P xs" | 
| 78 | using \<open>sorted cmp xs\<close> proof (induction xs) | |
| 79 | case Nil | |
| 80 | show ?case | |
| 81 | by (rule \<open>P []\<close>) | |
| 82 | next | |
| 83 | case (Cons x xs) | |
| 84 | then have "sorted cmp (x # xs)" | |
| 85 | by (simp add: sorted_ConsI) | |
| 86 | moreover note Cons.IH | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 87 | moreover have "\<And>y. compare cmp x y = Greater \<Longrightarrow> y \<in> set xs \<Longrightarrow> False" | 
| 69194 | 88 | using Cons.hyps by simp | 
| 89 | ultimately show ?case | |
| 90 | by (auto intro!: * [of "x # xs" x]) blast | |
| 91 | qed | |
| 92 | ||
| 93 | lemma sorted_remove1: | |
| 94 | "sorted cmp (remove1 x xs)" if "sorted cmp xs" | |
| 95 | proof (cases "x \<in> set xs") | |
| 96 | case False | |
| 97 | with that show ?thesis | |
| 98 | by (simp add: remove1_idem) | |
| 99 | next | |
| 100 | case True | |
| 101 | with that show ?thesis proof (induction xs) | |
| 102 | case Nil | |
| 103 | then show ?case | |
| 104 | by simp | |
| 105 | next | |
| 106 | case (Cons y ys) | |
| 107 | show ?case proof (cases "x = y") | |
| 108 | case True | |
| 109 | with Cons.hyps show ?thesis | |
| 110 | by simp | |
| 111 | next | |
| 112 | case False | |
| 113 | then have "sorted cmp (remove1 x ys)" | |
| 114 | using Cons.IH Cons.prems by auto | |
| 115 | then have "sorted cmp (y # remove1 x ys)" | |
| 116 | proof (rule sorted_ConsI) | |
| 117 | fix z zs | |
| 118 | assume "remove1 x ys = z # zs" | |
| 119 | with \<open>x \<noteq> y\<close> have "z \<in> set ys" | |
| 120 | using notin_set_remove1 [of z ys x] by auto | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 121 | then show "compare cmp y z \<noteq> Greater" | 
| 69194 | 122 | by (rule Cons.hyps(2)) | 
| 123 | qed | |
| 124 | with False show ?thesis | |
| 125 | by simp | |
| 126 | qed | |
| 127 | qed | |
| 128 | qed | |
| 129 | ||
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 130 | lemma sorted_stable_segment: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 131 | "sorted cmp (stable_segment cmp x xs)" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 132 | proof (induction xs) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 133 | case Nil | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 134 | show ?case | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 135 | by simp | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 136 | next | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 137 | case (Cons y ys) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 138 | then show ?case | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 139 | by (auto intro!: sorted_ConsI simp add: filter_eq_Cons_iff compare.sym) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 140 | (auto dest: compare.trans_equiv simp add: compare.sym compare.greater_iff_sym_less) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 141 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 142 | qed | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 143 | |
| 69194 | 144 | primrec insort :: "'a comparator \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 145 | where "insort cmp y [] = [y]" | |
| 146 | | "insort cmp y (x # xs) = (if compare cmp y x \<noteq> Greater | |
| 147 | then y # x # xs | |
| 148 | else x # insort cmp y xs)" | |
| 149 | ||
| 150 | lemma mset_insort [simp]: | |
| 151 | "mset (insort cmp x xs) = add_mset x (mset xs)" | |
| 152 | by (induction xs) simp_all | |
| 153 | ||
| 154 | lemma length_insort [simp]: | |
| 155 | "length (insort cmp x xs) = Suc (length xs)" | |
| 156 | by (induction xs) simp_all | |
| 157 | ||
| 158 | lemma sorted_insort: | |
| 159 | "sorted cmp (insort cmp x xs)" if "sorted cmp xs" | |
| 160 | using that proof (induction xs) | |
| 161 | case Nil | |
| 162 | then show ?case | |
| 163 | by simp | |
| 164 | next | |
| 165 | case (Cons y ys) | |
| 166 | then show ?case by (cases ys) | |
| 167 | (auto, simp_all add: compare.greater_iff_sym_less) | |
| 168 | qed | |
| 169 | ||
| 170 | lemma stable_insort_equiv: | |
| 171 | "stable_segment cmp y (insort cmp x xs) = x # stable_segment cmp y xs" | |
| 172 | if "compare cmp y x = Equiv" | |
| 173 | proof (induction xs) | |
| 174 | case Nil | |
| 175 | from that show ?case | |
| 176 | by simp | |
| 177 | next | |
| 178 | case (Cons z xs) | |
| 179 | moreover from that have "compare cmp y z = Equiv \<Longrightarrow> compare cmp z x = Equiv" | |
| 180 | by (auto intro: compare.trans_equiv simp add: compare.sym) | |
| 181 | ultimately show ?case | |
| 182 | using that by (auto simp add: compare.greater_iff_sym_less) | |
| 183 | qed | |
| 184 | ||
| 185 | lemma stable_insort_not_equiv: | |
| 186 | "stable_segment cmp y (insort cmp x xs) = stable_segment cmp y xs" | |
| 187 | if "compare cmp y x \<noteq> Equiv" | |
| 188 | using that by (induction xs) simp_all | |
| 189 | ||
| 190 | lemma remove1_insort_same_eq [simp]: | |
| 191 | "remove1 x (insort cmp x xs) = xs" | |
| 192 | by (induction xs) simp_all | |
| 193 | ||
| 194 | lemma insort_eq_ConsI: | |
| 195 | "insort cmp x xs = x # xs" | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 196 | if "sorted cmp xs" "\<And>y. y \<in> set xs \<Longrightarrow> compare cmp x y \<noteq> Greater" | 
| 69194 | 197 | using that by (induction xs) (simp_all add: compare.greater_iff_sym_less) | 
| 198 | ||
| 199 | lemma remove1_insort_not_same_eq [simp]: | |
| 200 | "remove1 y (insort cmp x xs) = insort cmp x (remove1 y xs)" | |
| 201 | if "sorted cmp xs" "x \<noteq> y" | |
| 202 | using that proof (induction xs) | |
| 203 | case Nil | |
| 204 | then show ?case | |
| 205 | by simp | |
| 206 | next | |
| 207 | case (Cons z zs) | |
| 208 | show ?case | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 209 | proof (cases "compare cmp x z = Greater") | 
| 69194 | 210 | case True | 
| 211 | with Cons show ?thesis | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 212 | by simp | 
| 69194 | 213 | next | 
| 214 | case False | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 215 | then have "compare cmp x y \<noteq> Greater" if "y \<in> set zs" for y | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 216 | using that Cons.hyps | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 217 | by (auto dest: compare.trans_not_greater) | 
| 69194 | 218 | with Cons show ?thesis | 
| 219 | by (simp add: insort_eq_ConsI) | |
| 220 | qed | |
| 221 | qed | |
| 222 | ||
| 223 | lemma insort_remove1_same_eq: | |
| 224 | "insort cmp x (remove1 x xs) = xs" | |
| 225 | if "sorted cmp xs" and "x \<in> set xs" and "hd (stable_segment cmp x xs) = x" | |
| 226 | using that proof (induction xs) | |
| 227 | case Nil | |
| 228 | then show ?case | |
| 229 | by simp | |
| 230 | next | |
| 231 | case (Cons y ys) | |
| 232 | then have "compare cmp x y \<noteq> Less" | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 233 | by (auto simp add: compare.greater_iff_sym_less) | 
| 69194 | 234 | then consider "compare cmp x y = Greater" | "compare cmp x y = Equiv" | 
| 235 | by (cases "compare cmp x y") auto | |
| 236 | then show ?case proof cases | |
| 237 | case 1 | |
| 238 | with Cons.prems Cons.IH show ?thesis | |
| 239 | by auto | |
| 240 | next | |
| 241 | case 2 | |
| 242 | with Cons.prems have "x = y" | |
| 243 | by simp | |
| 244 | with Cons.hyps show ?thesis | |
| 245 | by (simp add: insort_eq_ConsI) | |
| 246 | qed | |
| 247 | qed | |
| 248 | ||
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 249 | lemma sorted_append_iff: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 250 | "sorted cmp (xs @ ys) \<longleftrightarrow> sorted cmp xs \<and> sorted cmp ys | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 251 | \<and> (\<forall>x \<in> set xs. \<forall>y \<in> set ys. compare cmp x y \<noteq> Greater)" (is "?P \<longleftrightarrow> ?R \<and> ?S \<and> ?Q") | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 252 | proof | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 253 | assume ?P | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 254 | have ?R | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 255 | using \<open>?P\<close> by (induction xs) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 256 | (auto simp add: sorted_Cons_imp_not_less, | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 257 | auto simp add: sorted_Cons_imp_sorted intro: sorted_ConsI) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 258 | moreover have ?S | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 259 | using \<open>?P\<close> by (induction xs) (auto dest: sorted_Cons_imp_sorted) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 260 | moreover have ?Q | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 261 | using \<open>?P\<close> by (induction xs) (auto simp add: sorted_Cons_imp_not_less, | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 262 | simp add: sorted_Cons_imp_sorted) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 263 | ultimately show "?R \<and> ?S \<and> ?Q" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 264 | by simp | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 265 | next | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 266 | assume "?R \<and> ?S \<and> ?Q" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 267 | then have ?R ?S ?Q | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 268 | by simp_all | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 269 | then show ?P | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 270 | by (induction xs) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 271 | (auto simp add: append_eq_Cons_conv intro!: sorted_ConsI) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 272 | qed | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 273 | |
| 69194 | 274 | definition sort :: "'a comparator \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 275 | where "sort cmp xs = foldr (insort cmp) xs []" | |
| 276 | ||
| 277 | lemma sort_simps [simp]: | |
| 278 | "sort cmp [] = []" | |
| 279 | "sort cmp (x # xs) = insort cmp x (sort cmp xs)" | |
| 280 | by (simp_all add: sort_def) | |
| 281 | ||
| 282 | lemma mset_sort [simp]: | |
| 283 | "mset (sort cmp xs) = mset xs" | |
| 284 | by (induction xs) simp_all | |
| 285 | ||
| 286 | lemma length_sort [simp]: | |
| 287 | "length (sort cmp xs) = length xs" | |
| 288 | by (induction xs) simp_all | |
| 289 | ||
| 290 | lemma sorted_sort [simp]: | |
| 291 | "sorted cmp (sort cmp xs)" | |
| 292 | by (induction xs) (simp_all add: sorted_insort) | |
| 293 | ||
| 294 | lemma stable_sort: | |
| 295 | "stable_segment cmp x (sort cmp xs) = stable_segment cmp x xs" | |
| 296 | by (induction xs) (simp_all add: stable_insort_equiv stable_insort_not_equiv) | |
| 297 | ||
| 298 | lemma sort_remove1_eq [simp]: | |
| 299 | "sort cmp (remove1 x xs) = remove1 x (sort cmp xs)" | |
| 300 | by (induction xs) simp_all | |
| 301 | ||
| 302 | lemma set_insort [simp]: | |
| 303 | "set (insort cmp x xs) = insert x (set xs)" | |
| 304 | by (induction xs) auto | |
| 305 | ||
| 306 | lemma set_sort [simp]: | |
| 307 | "set (sort cmp xs) = set xs" | |
| 308 | by (induction xs) auto | |
| 309 | ||
| 310 | lemma sort_eqI: | |
| 311 | "sort cmp ys = xs" | |
| 312 | if permutation: "mset ys = mset xs" | |
| 313 | and sorted: "sorted cmp xs" | |
| 314 | and stable: "\<And>y. y \<in> set ys \<Longrightarrow> | |
| 315 | stable_segment cmp y ys = stable_segment cmp y xs" | |
| 316 | proof - | |
| 317 | have stable': "stable_segment cmp y ys = | |
| 318 | stable_segment cmp y xs" for y | |
| 319 | proof (cases "\<exists>x\<in>set ys. compare cmp y x = Equiv") | |
| 320 | case True | |
| 321 | then obtain z where "z \<in> set ys" and "compare cmp y z = Equiv" | |
| 322 | by auto | |
| 323 | then have "compare cmp y x = Equiv \<longleftrightarrow> compare cmp z x = Equiv" for x | |
| 324 | by (meson compare.sym compare.trans_equiv) | |
| 325 | moreover have "stable_segment cmp z ys = | |
| 326 | stable_segment cmp z xs" | |
| 327 | using \<open>z \<in> set ys\<close> by (rule stable) | |
| 328 | ultimately show ?thesis | |
| 329 | by simp | |
| 330 | next | |
| 331 | case False | |
| 332 | moreover from permutation have "set ys = set xs" | |
| 333 | by (rule mset_eq_setD) | |
| 334 | ultimately show ?thesis | |
| 335 | by simp | |
| 336 | qed | |
| 337 | show ?thesis | |
| 338 | using sorted permutation stable' proof (induction xs arbitrary: ys rule: sorted_induct_remove1) | |
| 339 | case Nil | |
| 340 | then show ?case | |
| 341 | by simp | |
| 342 | next | |
| 343 | case (minimum x xs) | |
| 344 | from \<open>mset ys = mset xs\<close> have ys: "set ys = set xs" | |
| 345 | by (rule mset_eq_setD) | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 346 | then have "compare cmp x y \<noteq> Greater" if "y \<in> set ys" for y | 
| 69194 | 347 | using that minimum.hyps by simp | 
| 348 | from minimum.prems have stable: "stable_segment cmp x ys = stable_segment cmp x xs" | |
| 349 | by simp | |
| 350 | have "sort cmp (remove1 x ys) = remove1 x xs" | |
| 351 | by (rule minimum.IH) (simp_all add: minimum.prems filter_remove1) | |
| 352 | then have "remove1 x (sort cmp ys) = remove1 x xs" | |
| 353 | by simp | |
| 354 | then have "insort cmp x (remove1 x (sort cmp ys)) = | |
| 355 | insort cmp x (remove1 x xs)" | |
| 356 | by simp | |
| 357 | also from minimum.hyps ys stable have "insort cmp x (remove1 x (sort cmp ys)) = sort cmp ys" | |
| 358 | by (simp add: stable_sort insort_remove1_same_eq) | |
| 359 | also from minimum.hyps have "insort cmp x (remove1 x xs) = xs" | |
| 360 | by (simp add: insort_remove1_same_eq) | |
| 361 | finally show ?case . | |
| 362 | qed | |
| 363 | qed | |
| 364 | ||
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 365 | lemma filter_insort: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 366 | "filter P (insort cmp x xs) = insort cmp x (filter P xs)" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 367 | if "sorted cmp xs" and "P x" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 368 | using that by (induction xs) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 369 | (auto simp add: compare.trans_not_greater insort_eq_ConsI) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 370 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 371 | lemma filter_insort_triv: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 372 | "filter P (insort cmp x xs) = filter P xs" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 373 | if "\<not> P x" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 374 | using that by (induction xs) simp_all | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 375 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 376 | lemma filter_sort: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 377 | "filter P (sort cmp xs) = sort cmp (filter P xs)" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 378 | by (induction xs) (auto simp add: filter_insort filter_insort_triv) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 379 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 380 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 381 | section \<open>Alternative sorting algorithms\<close> | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 382 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 383 | subsection \<open>Quicksort\<close> | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 384 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 385 | definition quicksort :: "'a comparator \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 386 | where quicksort_is_sort [simp]: "quicksort = sort" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 387 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 388 | lemma sort_by_quicksort: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 389 | "sort = quicksort" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 390 | by simp | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 391 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 392 | lemma sort_by_quicksort_rec: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 393 | "sort cmp xs = sort cmp [x\<leftarrow>xs. compare cmp x (xs ! (length xs div 2)) = Less] | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 394 | @ stable_segment cmp (xs ! (length xs div 2)) xs | 
| 69250 | 395 | @ sort cmp [x\<leftarrow>xs. compare cmp x (xs ! (length xs div 2)) = Greater]" (is "_ = ?rhs") | 
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 396 | proof (rule sort_eqI) | 
| 69250 | 397 | show "mset xs = mset ?rhs" | 
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 398 | by (rule multiset_eqI) (auto simp add: compare.sym intro: comp.exhaust) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 399 | next | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 400 | show "sorted cmp ?rhs" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 401 | by (auto simp add: sorted_append_iff sorted_stable_segment compare.equiv_subst_right dest: compare.trans_greater) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 402 | next | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 403 | let ?pivot = "xs ! (length xs div 2)" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 404 | fix l | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 405 | have "compare cmp x ?pivot = comp \<and> compare cmp l x = Equiv | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 406 | \<longleftrightarrow> compare cmp l ?pivot = comp \<and> compare cmp l x = Equiv" for x comp | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 407 | proof - | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 408 | have "compare cmp x ?pivot = comp \<longleftrightarrow> compare cmp l ?pivot = comp" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 409 | if "compare cmp l x = Equiv" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 410 | using that by (simp add: compare.equiv_subst_left compare.sym) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 411 | then show ?thesis by blast | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 412 | qed | 
| 69250 | 413 | then show "stable_segment cmp l xs = stable_segment cmp l ?rhs" | 
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 414 | by (simp add: stable_sort compare.sym [of _ ?pivot]) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 415 | (cases "compare cmp l ?pivot", simp_all) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 416 | qed | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 417 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 418 | context | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 419 | begin | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 420 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 421 | qualified definition partition :: "'a comparator \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list \<times> 'a list" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 422 | where "partition cmp pivot xs = | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 423 | ([x \<leftarrow> xs. compare cmp x pivot = Less], stable_segment cmp pivot xs, [x \<leftarrow> xs. compare cmp x pivot = Greater])" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 424 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 425 | qualified lemma partition_code [code]: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 426 | "partition cmp pivot [] = ([], [], [])" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 427 | "partition cmp pivot (x # xs) = | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 428 | (let (lts, eqs, gts) = partition cmp pivot xs | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 429 | in case compare cmp x pivot of | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 430 | Less \<Rightarrow> (x # lts, eqs, gts) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 431 | | Equiv \<Rightarrow> (lts, x # eqs, gts) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 432 | | Greater \<Rightarrow> (lts, eqs, x # gts))" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 433 | using comp.exhaust by (auto simp add: partition_def Let_def compare.sym [of _ pivot]) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 434 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 435 | lemma quicksort_code [code]: | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 436 | "quicksort cmp xs = | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 437 | (case xs of | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 438 | [] \<Rightarrow> [] | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 439 | | [x] \<Rightarrow> xs | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 440 | | [x, y] \<Rightarrow> (if compare cmp x y \<noteq> Greater then xs else [y, x]) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 441 | | _ \<Rightarrow> | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 442 | let (lts, eqs, gts) = partition cmp (xs ! (length xs div 2)) xs | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 443 | in quicksort cmp lts @ eqs @ quicksort cmp gts)" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 444 | proof (cases "length xs \<ge> 3") | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 445 | case False | 
| 69250 | 446 |   then have "length xs \<in> {0, 1, 2}"
 | 
| 447 | by (auto simp add: not_le le_less less_antisym) | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 448 | then consider "xs = []" | x where "xs = [x]" | x y where "xs = [x, y]" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 449 | by (auto simp add: length_Suc_conv numeral_2_eq_2) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 450 | then show ?thesis | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 451 | by cases simp_all | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 452 | next | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 453 | case True | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 454 | then obtain x y z zs where "xs = x # y # z # zs" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 455 | by (metis le_0_eq length_0_conv length_Cons list.exhaust not_less_eq_eq numeral_3_eq_3) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 456 | moreover have "quicksort cmp xs = | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 457 | (let (lts, eqs, gts) = partition cmp (xs ! (length xs div 2)) xs | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 458 | in quicksort cmp lts @ eqs @ quicksort cmp gts)" | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 459 | using sort_by_quicksort_rec [of cmp xs] by (simp add: partition_def) | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 460 | ultimately show ?thesis | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 461 | by simp | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 462 | qed | 
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 463 | |
| 69194 | 464 | end | 
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 465 | |
| 69250 | 466 | |
| 467 | subsection \<open>Mergesort\<close> | |
| 468 | ||
| 469 | definition mergesort :: "'a comparator \<Rightarrow> 'a list \<Rightarrow> 'a list" | |
| 470 | where mergesort_is_sort [simp]: "mergesort = sort" | |
| 471 | ||
| 472 | lemma sort_by_mergesort: | |
| 473 | "sort = mergesort" | |
| 474 | by simp | |
| 475 | ||
| 476 | context | |
| 477 | fixes cmp :: "'a comparator" | |
| 478 | begin | |
| 479 | ||
| 480 | qualified function merge :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" | |
| 481 | where "merge [] ys = ys" | |
| 482 | | "merge xs [] = xs" | |
| 483 | | "merge (x # xs) (y # ys) = (if compare cmp x y = Greater | |
| 484 | then y # merge (x # xs) ys else x # merge xs (y # ys))" | |
| 485 | by pat_completeness auto | |
| 486 | ||
| 487 | qualified termination by lexicographic_order | |
| 488 | ||
| 489 | lemma mset_merge: | |
| 490 | "mset (merge xs ys) = mset xs + mset ys" | |
| 491 | by (induction xs ys rule: merge.induct) simp_all | |
| 492 | ||
| 493 | lemma merge_eq_Cons_imp: | |
| 494 | "xs \<noteq> [] \<and> z = hd xs \<or> ys \<noteq> [] \<and> z = hd ys" | |
| 495 | if "merge xs ys = z # zs" | |
| 496 | using that by (induction xs ys rule: merge.induct) (auto split: if_splits) | |
| 497 | ||
| 498 | lemma filter_merge: | |
| 499 | "filter P (merge xs ys) = merge (filter P xs) (filter P ys)" | |
| 500 | if "sorted cmp xs" and "sorted cmp ys" | |
| 501 | using that proof (induction xs ys rule: merge.induct) | |
| 502 | case (1 ys) | |
| 503 | then show ?case | |
| 504 | by simp | |
| 505 | next | |
| 506 | case (2 xs) | |
| 507 | then show ?case | |
| 508 | by simp | |
| 509 | next | |
| 510 | case (3 x xs y ys) | |
| 511 | show ?case | |
| 512 | proof (cases "compare cmp x y = Greater") | |
| 513 | case True | |
| 514 | with 3 have hyp: "filter P (merge (x # xs) ys) = | |
| 515 | merge (filter P (x # xs)) (filter P ys)" | |
| 516 | by (simp add: sorted_Cons_imp_sorted) | |
| 517 | show ?thesis | |
| 518 | proof (cases "\<not> P x \<and> P y") | |
| 519 | case False | |
| 520 | with \<open>compare cmp x y = Greater\<close> show ?thesis | |
| 521 | by (auto simp add: hyp) | |
| 522 | next | |
| 523 | case True | |
| 524 | from \<open>compare cmp x y = Greater\<close> "3.prems" | |
| 525 | have *: "compare cmp z y = Greater" if "z \<in> set (filter P xs)" for z | |
| 526 | using that by (auto dest: compare.trans_not_greater sorted_Cons_imp_not_less) | |
| 527 | from \<open>compare cmp x y = Greater\<close> show ?thesis | |
| 528 | by (cases "filter P xs") (simp_all add: hyp *) | |
| 529 | qed | |
| 530 | next | |
| 531 | case False | |
| 532 | with 3 have hyp: "filter P (merge xs (y # ys)) = | |
| 533 | merge (filter P xs) (filter P (y # ys))" | |
| 534 | by (simp add: sorted_Cons_imp_sorted) | |
| 535 | show ?thesis | |
| 536 | proof (cases "P x \<and> \<not> P y") | |
| 537 | case False | |
| 538 | with \<open>compare cmp x y \<noteq> Greater\<close> show ?thesis | |
| 539 | by (auto simp add: hyp) | |
| 540 | next | |
| 541 | case True | |
| 542 | from \<open>compare cmp x y \<noteq> Greater\<close> "3.prems" | |
| 543 | have *: "compare cmp x z \<noteq> Greater" if "z \<in> set (filter P ys)" for z | |
| 544 | using that by (auto dest: compare.trans_not_greater sorted_Cons_imp_not_less) | |
| 545 | from \<open>compare cmp x y \<noteq> Greater\<close> show ?thesis | |
| 546 | by (cases "filter P ys") (simp_all add: hyp *) | |
| 547 | qed | |
| 548 | qed | |
| 549 | qed | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 550 | |
| 69250 | 551 | lemma sorted_merge: | 
| 552 | "sorted cmp (merge xs ys)" if "sorted cmp xs" and "sorted cmp ys" | |
| 553 | using that proof (induction xs ys rule: merge.induct) | |
| 554 | case (1 ys) | |
| 555 | then show ?case | |
| 556 | by simp | |
| 557 | next | |
| 558 | case (2 xs) | |
| 559 | then show ?case | |
| 560 | by simp | |
| 561 | next | |
| 562 | case (3 x xs y ys) | |
| 563 | show ?case | |
| 564 | proof (cases "compare cmp x y = Greater") | |
| 565 | case True | |
| 566 | with 3 have "sorted cmp (merge (x # xs) ys)" | |
| 567 | by (simp add: sorted_Cons_imp_sorted) | |
| 568 | then have "sorted cmp (y # merge (x # xs) ys)" | |
| 569 | proof (rule sorted_ConsI) | |
| 570 | fix z zs | |
| 571 | assume "merge (x # xs) ys = z # zs" | |
| 572 | with 3(4) True show "compare cmp y z \<noteq> Greater" | |
| 573 | by (clarsimp simp add: sorted_Cons_imp_sorted dest!: merge_eq_Cons_imp) | |
| 574 | (auto simp add: compare.asym_greater sorted_Cons_imp_not_less) | |
| 575 | qed | |
| 576 | with True show ?thesis | |
| 577 | by simp | |
| 578 | next | |
| 579 | case False | |
| 580 | with 3 have "sorted cmp (merge xs (y # ys))" | |
| 581 | by (simp add: sorted_Cons_imp_sorted) | |
| 582 | then have "sorted cmp (x # merge xs (y # ys))" | |
| 583 | proof (rule sorted_ConsI) | |
| 584 | fix z zs | |
| 585 | assume "merge xs (y # ys) = z # zs" | |
| 586 | with 3(3) False show "compare cmp x z \<noteq> Greater" | |
| 587 | by (clarsimp simp add: sorted_Cons_imp_sorted dest!: merge_eq_Cons_imp) | |
| 588 | (auto simp add: compare.asym_greater sorted_Cons_imp_not_less) | |
| 589 | qed | |
| 590 | with False show ?thesis | |
| 591 | by simp | |
| 592 | qed | |
| 593 | qed | |
| 594 | ||
| 595 | lemma merge_eq_appendI: | |
| 596 | "merge xs ys = xs @ ys" | |
| 597 | if "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set ys \<Longrightarrow> compare cmp x y \<noteq> Greater" | |
| 598 | using that by (induction xs ys rule: merge.induct) simp_all | |
| 599 | ||
| 600 | lemma merge_stable_segments: | |
| 601 | "merge (stable_segment cmp l xs) (stable_segment cmp l ys) = | |
| 602 | stable_segment cmp l xs @ stable_segment cmp l ys" | |
| 603 | by (rule merge_eq_appendI) (auto dest: compare.trans_equiv_greater) | |
| 604 | ||
| 605 | lemma sort_by_mergesort_rec: | |
| 606 | "sort cmp xs = | |
| 607 | merge (sort cmp (take (length xs div 2) xs)) | |
| 608 | (sort cmp (drop (length xs div 2) xs))" (is "_ = ?rhs") | |
| 609 | proof (rule sort_eqI) | |
| 610 | have "mset (take (length xs div 2) xs) + mset (drop (length xs div 2) xs) = | |
| 611 | mset (take (length xs div 2) xs @ drop (length xs div 2) xs)" | |
| 612 | by (simp only: mset_append) | |
| 613 | then show "mset xs = mset ?rhs" | |
| 614 | by (simp add: mset_merge) | |
| 615 | next | |
| 616 | show "sorted cmp ?rhs" | |
| 617 | by (simp add: sorted_merge) | |
| 618 | next | |
| 619 | fix l | |
| 620 | have "stable_segment cmp l (take (length xs div 2) xs) @ stable_segment cmp l (drop (length xs div 2) xs) | |
| 621 | = stable_segment cmp l xs" | |
| 622 | by (simp only: filter_append [symmetric] append_take_drop_id) | |
| 623 | have "merge (stable_segment cmp l (take (length xs div 2) xs)) | |
| 624 | (stable_segment cmp l (drop (length xs div 2) xs)) = | |
| 625 | stable_segment cmp l (take (length xs div 2) xs) @ stable_segment cmp l (drop (length xs div 2) xs)" | |
| 626 | by (rule merge_eq_appendI) (auto simp add: compare.trans_equiv_greater) | |
| 627 | also have "\<dots> = stable_segment cmp l xs" | |
| 628 | by (simp only: filter_append [symmetric] append_take_drop_id) | |
| 629 | finally show "stable_segment cmp l xs = stable_segment cmp l ?rhs" | |
| 630 | by (simp add: stable_sort filter_merge) | |
| 631 | qed | |
| 632 | ||
| 633 | lemma mergesort_code [code]: | |
| 634 | "mergesort cmp xs = | |
| 635 | (case xs of | |
| 636 | [] \<Rightarrow> [] | |
| 637 | | [x] \<Rightarrow> xs | |
| 638 | | [x, y] \<Rightarrow> (if compare cmp x y \<noteq> Greater then xs else [y, x]) | |
| 639 | | _ \<Rightarrow> | |
| 640 | let | |
| 641 | half = length xs div 2; | |
| 642 | ys = take half xs; | |
| 643 | zs = drop half xs | |
| 644 | in merge (mergesort cmp ys) (mergesort cmp zs))" | |
| 645 | proof (cases "length xs \<ge> 3") | |
| 646 | case False | |
| 647 |   then have "length xs \<in> {0, 1, 2}"
 | |
| 648 | by (auto simp add: not_le le_less less_antisym) | |
| 649 | then consider "xs = []" | x where "xs = [x]" | x y where "xs = [x, y]" | |
| 650 | by (auto simp add: length_Suc_conv numeral_2_eq_2) | |
| 651 | then show ?thesis | |
| 652 | by cases simp_all | |
| 653 | next | |
| 654 | case True | |
| 655 | then obtain x y z zs where "xs = x # y # z # zs" | |
| 656 | by (metis le_0_eq length_0_conv length_Cons list.exhaust not_less_eq_eq numeral_3_eq_3) | |
| 657 | moreover have "mergesort cmp xs = | |
| 658 | (let | |
| 659 | half = length xs div 2; | |
| 660 | ys = take half xs; | |
| 661 | zs = drop half xs | |
| 662 | in merge (mergesort cmp ys) (mergesort cmp zs))" | |
| 663 | using sort_by_mergesort_rec [of xs] by (simp add: Let_def) | |
| 664 | ultimately show ?thesis | |
| 665 | by simp | |
| 666 | qed | |
| 69246 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 667 | |
| 
c1fe9dcc274a
concrecte sorting algorithms beyond insertion sort
 haftmann parents: 
69194diff
changeset | 668 | end | 
| 69250 | 669 | |
| 670 | end |