author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 46 | a73f8a7784bd |
child 195 | df6b3bd14dcb |
permissions | -rw-r--r-- |
46 | 1 |
(* Title: HOL/ex/insort.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1994 TU Muenchen |
|
5 |
||
6 |
Insertion sort |
|
7 |
*) |
|
8 |
||
9 |
InSort = Sorting + |
|
10 |
||
11 |
consts |
|
12 |
ins :: "[['a,'a]=>bool, 'a, 'a list] => 'a list" |
|
13 |
insort :: "[['a,'a]=>bool, 'a list] => 'a list" |
|
14 |
||
15 |
rules |
|
16 |
ins_Nil "ins(f,x,[]) = [x]" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
46
diff
changeset
|
17 |
ins_Cons "ins(f,x,y#ys) = \ |
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
46
diff
changeset
|
18 |
\ if(f(x,y), x#y#ys, y# ins(f,x,ys))" |
46 | 19 |
|
20 |
insort_Nil "insort(f,[]) = []" |
|
48
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
clasohm
parents:
46
diff
changeset
|
21 |
insort_Cons "insort(f,x#xs) = ins(f,x,insort(f,xs))" |
46 | 22 |
end |