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