| author | wenzelm | 
| Tue, 12 Aug 2014 17:28:07 +0200 | |
| changeset 57915 | 448325de6e4f | 
| parent 44387 | 0f0ba362ce50 | 
| child 59194 | b51489b75bb9 | 
| permissions | -rw-r--r-- | 
| 32817 | 1  | 
(* Title: Pure/Concurrent/lazy_sequential.ML  | 
2  | 
Author: Florian Haftmann and Makarius, TU Muenchen  | 
|
3  | 
||
| 34278 | 4  | 
Lazy evaluation with memoing of results and regular exceptions  | 
5  | 
(sequential version).  | 
|
| 32817 | 6  | 
*)  | 
7  | 
||
8  | 
structure Lazy: LAZY =  | 
|
9  | 
struct  | 
|
10  | 
||
11  | 
(* datatype *)  | 
|
12  | 
||
13  | 
datatype 'a expr =  | 
|
14  | 
Expr of unit -> 'a |  | 
|
15  | 
Result of 'a Exn.result;  | 
|
16  | 
||
17  | 
abstype 'a lazy = Lazy of 'a expr Unsynchronized.ref  | 
|
18  | 
with  | 
|
19  | 
||
20  | 
fun peek (Lazy r) =  | 
|
21  | 
(case ! r of  | 
|
22  | 
Expr _ => NONE  | 
|
| 32844 | 23  | 
| Result res => SOME res);  | 
| 32817 | 24  | 
|
25  | 
fun lazy e = Lazy (Unsynchronized.ref (Expr e));  | 
|
| 
43761
 
e72ba84ae58f
tuned signature -- corresponding to Scala version;
 
wenzelm 
parents: 
39232 
diff
changeset
 | 
26  | 
fun value a = Lazy (Unsynchronized.ref (Result (Exn.Res a)));  | 
| 32817 | 27  | 
|
| 44198 | 28  | 
fun is_finished x = is_some (peek x);  | 
29  | 
||
| 32817 | 30  | 
|
31  | 
(* force result *)  | 
|
32  | 
||
33  | 
fun force_result (Lazy r) =  | 
|
| 34278 | 34  | 
let  | 
35  | 
val result =  | 
|
36  | 
(case ! r of  | 
|
37  | 
Expr e => Exn.capture e ()  | 
|
38  | 
| Result res => res);  | 
|
| 
39232
 
69c6d3e87660
more abstract treatment of interrupts in structure Exn -- hardly ever need to mention Interrupt literally;
 
wenzelm 
parents: 
34293 
diff
changeset
 | 
39  | 
val _ = if Exn.is_interrupt_exn result then () else r := Result result;  | 
| 34278 | 40  | 
in result end;  | 
| 32817 | 41  | 
|
42  | 
fun force r = Exn.release (force_result r);  | 
|
| 
44386
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
43  | 
fun map f x = lazy (fn () => f (force x));  | 
| 
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
44  | 
|
| 
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
45  | 
|
| 
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
46  | 
(* future evaluation *)  | 
| 
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
47  | 
|
| 
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
48  | 
fun future params x =  | 
| 
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
49  | 
if is_finished x then Future.value_result (force_result x)  | 
| 
 
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
 
wenzelm 
parents: 
44198 
diff
changeset
 | 
50  | 
else (singleton o Future.forks) params (fn () => force x);  | 
| 32817 | 51  | 
|
52  | 
end;  | 
|
53  | 
end;  | 
|
54  | 
||
55  | 
type 'a lazy = 'a Lazy.lazy;  | 
|
56  |