author | haftmann |
Mon, 06 Feb 2017 20:56:32 +0100 | |
changeset 64988 | 93aaff2b0ae0 |
parent 62430 | 9527ff088c15 |
child 66453 | cc19f7ca2ed6 |
permissions | -rw-r--r-- |
30738 | 1 |
(* Author: Tobias Nipkow |
24615 | 2 |
Copyright 1994 TU Muenchen |
3 |
*) |
|
4 |
||
61343 | 5 |
section \<open>Quicksort with function package\<close> |
24615 | 6 |
|
7 |
theory Quicksort |
|
44604 | 8 |
imports "~~/src/HOL/Library/Multiset" |
24615 | 9 |
begin |
10 |
||
11 |
context linorder |
|
12 |
begin |
|
13 |
||
28041 | 14 |
fun quicksort :: "'a list \<Rightarrow> 'a list" where |
37075
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
15 |
"quicksort [] = []" |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
16 |
| "quicksort (x#xs) = quicksort [y\<leftarrow>xs. \<not> x\<le>y] @ [x] @ quicksort [y\<leftarrow>xs. x\<le>y]" |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
17 |
|
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
18 |
lemma [code]: |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
19 |
"quicksort [] = []" |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
20 |
"quicksort (x#xs) = quicksort [y\<leftarrow>xs. y<x] @ [x] @ quicksort [y\<leftarrow>xs. x\<le>y]" |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
21 |
by (simp_all add: not_le) |
24615 | 22 |
|
23 |
lemma quicksort_permutes [simp]: |
|
60515 | 24 |
"mset (quicksort xs) = mset xs" |
37075
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
25 |
by (induct xs rule: quicksort.induct) (simp_all add: ac_simps) |
24615 | 26 |
|
27 |
lemma set_quicksort [simp]: "set (quicksort xs) = set xs" |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61343
diff
changeset
|
28 |
proof - |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61343
diff
changeset
|
29 |
have "set_mset (mset (quicksort xs)) = set_mset (mset xs)" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61343
diff
changeset
|
30 |
by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61343
diff
changeset
|
31 |
then show ?thesis by (simp only: set_mset_mset) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61343
diff
changeset
|
32 |
qed |
24615 | 33 |
|
37075
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
34 |
lemma sorted_quicksort: "sorted (quicksort xs)" |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
35 |
by (induct xs rule: quicksort.induct) (auto simp add: sorted_Cons sorted_append not_le less_imp_le) |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
36 |
|
40349 | 37 |
theorem sort_quicksort: |
37075
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
38 |
"sort = quicksort" |
a680ce27aa56
modernized sorting algorithms; quicksort implements sort
haftmann
parents:
30738
diff
changeset
|
39 |
by (rule ext, rule properties_for_sort) (fact quicksort_permutes sorted_quicksort)+ |
24615 | 40 |
|
41 |
end |
|
42 |
||
43 |
end |