author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 46 | a73f8a7784bd |
permissions | -rw-r--r-- |
37 | 1 |
(* Title: HOL/ex/qsort.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1994 TU Muenchen |
|
5 |
||
6 |
Quicksort |
|
7 |
*) |
|
8 |
||
46 | 9 |
Qsort = Sorting + |
36 | 10 |
consts |
11 |
qsort :: "[['a,'a] => bool, 'a list] => 'a list" |
|
12 |
||
13 |
rules |
|
14 |
||
15 |
qsort_Nil "qsort(le,[]) = []" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
46
diff
changeset
|
16 |
qsort_Cons "qsort(le,x#xs) = qsort(le,[y:xs . ~le(x,y)]) @ \ |
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
46
diff
changeset
|
17 |
\ (x# qsort(le,[y:xs . le(x,y)]))" |
36 | 18 |
|
19 |
(* computational induction. |
|
20 |
The dependence of p on x but not xs is intentional. |
|
21 |
*) |
|
22 |
qsort_ind |
|
23 |
"[| P([]); \ |
|
24 |
\ !!x xs. [| P([y:xs . ~p(x,y)]); P([y:xs . p(x,y)]) |] ==> \ |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
46
diff
changeset
|
25 |
\ P(x#xs) |] \ |
36 | 26 |
\ ==> P(xs)" |
27 |
end |