8414
|
1 |
(* Title: HOL/ex/Qsort.thy
|
969
|
2 |
ID: $Id$
|
1476
|
3 |
Author: Tobias Nipkow
|
969
|
4 |
Copyright 1994 TU Muenchen
|
|
5 |
*)
|
|
6 |
|
15815
|
7 |
header{*Quicksort*}
|
8590
|
8 |
|
15815
|
9 |
theory Qsort
|
|
10 |
imports Sorting
|
|
11 |
begin
|
|
12 |
|
|
13 |
subsection{*Version 1: higher-order*}
|
|
14 |
|
13159
|
15 |
consts qsort :: "('a \<Rightarrow> 'a => bool) * 'a list \<Rightarrow> 'a list"
|
969
|
16 |
|
8414
|
17 |
recdef qsort "measure (size o snd)"
|
15815
|
18 |
"qsort(le, []) = []"
|
|
19 |
"qsort(le, x#xs) = qsort(le, [y:xs . ~ le x y]) @ [x] @
|
|
20 |
qsort(le, [y:xs . le x y])"
|
15247
|
21 |
(hints recdef_simp: length_filter_le[THEN le_less_trans])
|
13159
|
22 |
|
15815
|
23 |
lemma qsort_permutes [simp]:
|
|
24 |
"multiset_of (qsort(le,xs)) = multiset_of xs"
|
15631
|
25 |
by (induct le xs rule: qsort.induct) (auto simp: union_ac)
|
13159
|
26 |
|
15815
|
27 |
lemma set_qsort [simp]: "set (qsort(le,xs)) = set xs";
|
15631
|
28 |
by(simp add: set_count_greater_0)
|
13159
|
29 |
|
|
30 |
lemma sorted_qsort:
|
15815
|
31 |
"total(le) ==> transf(le) ==> sorted le (qsort(le,xs))"
|
13159
|
32 |
apply (induct le xs rule: qsort.induct)
|
|
33 |
apply simp
|
|
34 |
apply simp
|
|
35 |
apply(unfold Sorting.total_def Sorting.transf_def)
|
|
36 |
apply blast
|
|
37 |
done
|
8414
|
38 |
|
8590
|
39 |
|
15815
|
40 |
subsection{*Version 2:type classes*}
|
13159
|
41 |
|
8590
|
42 |
consts quickSort :: "('a::linorder) list => 'a list"
|
|
43 |
|
|
44 |
recdef quickSort "measure size"
|
15815
|
45 |
"quickSort [] = []"
|
|
46 |
"quickSort (x#l) = quickSort [y:l. ~ x\<le>y] @ [x] @ quickSort [y:l. x\<le>y]"
|
15247
|
47 |
(hints recdef_simp: length_filter_le[THEN le_less_trans])
|
13159
|
48 |
|
|
49 |
lemma quickSort_permutes[simp]:
|
15631
|
50 |
"multiset_of (quickSort xs) = multiset_of xs"
|
|
51 |
by (induct xs rule: quickSort.induct) (auto simp: union_ac)
|
13159
|
52 |
|
|
53 |
lemma set_quickSort[simp]: "set (quickSort xs) = set xs"
|
15631
|
54 |
by(simp add: set_count_greater_0)
|
13159
|
55 |
|
15815
|
56 |
theorem sorted_quickSort: "sorted (op \<le>) (quickSort xs)"
|
|
57 |
by (induct xs rule: quickSort.induct, auto)
|
8590
|
58 |
|
969
|
59 |
end
|