ex/qsort.thy
author clasohm
Wed, 02 Mar 1994 12:26:55 +0100
changeset 48 21291189b51e
parent 46 a73f8a7784bd
permissions -rw-r--r--
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     1
(*  Title: 	HOL/ex/qsort.thy
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     2
    ID:         $Id$
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     3
    Author: 	Tobias Nipkow
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     4
    Copyright   1994 TU Muenchen
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     5
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     6
Quicksort
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     7
*)
65a546c2b8ed changed header
nipkow
parents: 36
diff changeset
     8
46
a73f8a7784bd some more sorting algorithms
nipkow
parents: 37
diff changeset
     9
Qsort = Sorting +
36
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    10
consts
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    11
  qsort  :: "[['a,'a] => bool, 'a list] => 'a list"
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    12
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    13
rules
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    14
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    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
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    18
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    19
(* computational induction.
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    20
   The dependence of p on x but not xs is intentional.
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    21
*)
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    22
qsort_ind
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    23
 "[| P([]); \
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    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
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    26
\ ==> P(xs)"
b503da67d2f7 Verification of quicksort
nipkow
parents:
diff changeset
    27
end