| author | nipkow | 
| Wed, 21 Sep 2011 07:31:08 +0200 | |
| changeset 45023 | 76abd26e2e2d | 
| parent 32939 | 1b5a401c78cb | 
| child 70586 | 57df8a85317a | 
| permissions | -rw-r--r-- | 
| 9413 | 1  | 
(* Title: Pure/General/heap.ML  | 
| 29606 | 2  | 
Author: Lawrence C Paulson and Markus Wenzel  | 
| 9095 | 3  | 
|
| 9413 | 4  | 
Heaps over linearly ordered types. See also Chris Okasaki: "Purely  | 
5  | 
Functional Data Structures" (Chapter 3), Cambridge University Press,  | 
|
6  | 
1998.  | 
|
7  | 
*)  | 
|
| 9095 | 8  | 
|
9  | 
signature HEAP =  | 
|
10  | 
sig  | 
|
| 9413 | 11  | 
type elem  | 
12  | 
type T  | 
|
13  | 
val empty: T  | 
|
14  | 
val is_empty: T -> bool  | 
|
15  | 
val merge: T * T -> T  | 
|
| 23179 | 16  | 
val insert: elem -> T -> T  | 
| 28581 | 17  | 
val min: T -> elem (*exception Empty*)  | 
18  | 
val delete_min: T -> T (*exception Empty*)  | 
|
19  | 
val min_elem: T -> elem * T (*exception Empty*)  | 
|
20  | 
val upto: elem -> T -> elem list * T  | 
|
| 9095 | 21  | 
end;  | 
22  | 
||
| 32939 | 23  | 
functor Heap(type elem val ord: elem * elem -> order): HEAP =  | 
| 9095 | 24  | 
struct  | 
| 9413 | 25  | 
|
26  | 
||
27  | 
(* datatype heap *)  | 
|
| 9095 | 28  | 
|
| 9413 | 29  | 
type elem = elem;  | 
30  | 
datatype T = Empty | Heap of int * elem * T * T;  | 
|
31  | 
||
32  | 
||
33  | 
(* empty heaps *)  | 
|
| 9095 | 34  | 
|
| 9413 | 35  | 
val empty = Empty;  | 
36  | 
||
37  | 
fun is_empty Empty = true  | 
|
38  | 
| is_empty (Heap _) = false;  | 
|
39  | 
||
| 9095 | 40  | 
|
| 9413 | 41  | 
(* build heaps *)  | 
42  | 
||
43  | 
local  | 
|
44  | 
||
45  | 
fun rank Empty = 0  | 
|
46  | 
| rank (Heap (r, _, _, _)) = r;  | 
|
| 9095 | 47  | 
|
| 9413 | 48  | 
fun heap x a b =  | 
49  | 
if rank a >= rank b then Heap (rank b + 1, x, a, b)  | 
|
50  | 
else Heap (rank a + 1, x, b, a);  | 
|
51  | 
||
52  | 
in  | 
|
53  | 
||
54  | 
fun merge (h, Empty) = h  | 
|
55  | 
| merge (Empty, h) = h  | 
|
56  | 
| merge (h1 as Heap (_, x1, a1, b1), h2 as Heap (_, x2, a2, b2)) =  | 
|
57  | 
(case ord (x1, x2) of  | 
|
| 
14472
 
cba7c0a3ffb3
Removing the datatype declaration of "order" allows the standard General.order
 
paulson 
parents: 
9413 
diff
changeset
 | 
58  | 
GREATER => heap x2 a2 (merge (h1, b2))  | 
| 9413 | 59  | 
| _ => heap x1 a1 (merge (b1, h2)));  | 
60  | 
||
| 23179 | 61  | 
fun insert x h = merge (Heap (1, x, Empty, Empty), h);  | 
| 9095 | 62  | 
|
| 9413 | 63  | 
end;  | 
64  | 
||
65  | 
||
66  | 
(* minimum element *)  | 
|
67  | 
||
| 18134 | 68  | 
fun min Empty = raise List.Empty  | 
| 9413 | 69  | 
| min (Heap (_, x, _, _)) = x;  | 
70  | 
||
| 18134 | 71  | 
fun delete_min Empty = raise List.Empty  | 
| 9413 | 72  | 
| delete_min (Heap (_, _, a, b)) = merge (a, b);  | 
73  | 
||
| 28581 | 74  | 
fun min_elem h = (min h, delete_min h);  | 
75  | 
||
76  | 
||
77  | 
(* initial interval *)  | 
|
78  | 
||
79  | 
nonfix upto;  | 
|
80  | 
||
| 32784 | 81  | 
fun upto _ Empty = ([], Empty)  | 
82  | 
| upto limit (h as Heap (_, x, _, _)) =  | 
|
| 28581 | 83  | 
(case ord (x, limit) of  | 
84  | 
GREATER => ([], h)  | 
|
85  | 
| _ => upto limit (delete_min h) |>> cons x);  | 
|
86  | 
||
| 9095 | 87  | 
end;  |