author | wenzelm |
Fri, 23 Jun 2017 14:38:32 +0200 | |
changeset 66176 | b51a40281016 |
parent 66168 | fcd09fc36d7f |
child 66904 | d9783ea1160c |
permissions | -rw-r--r-- |
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
1 |
(* Title: Pure/Concurrent/lazy.ML |
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
2 |
Author: Makarius |
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
3 |
|
34278 | 4 |
Lazy evaluation with memoing of results and regular exceptions. |
40389 | 5 |
Parallel version based on (passive) futures, to avoid critical or |
6 |
multiple evaluation (unless interrupted). |
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
7 |
*) |
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
8 |
|
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
9 |
signature LAZY = |
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
10 |
sig |
28971 | 11 |
type 'a lazy |
66167 | 12 |
val lazy_name: string -> (unit -> 'a) -> 'a lazy |
28971 | 13 |
val lazy: (unit -> 'a) -> 'a lazy |
14 |
val value: 'a -> 'a lazy |
|
59193
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
15 |
val peek: 'a lazy -> 'a Exn.result option |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
16 |
val is_running: 'a lazy -> bool |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
17 |
val is_finished: 'a lazy -> bool |
29422 | 18 |
val force_result: 'a lazy -> 'a Exn.result |
28971 | 19 |
val force: 'a lazy -> 'a |
66168 | 20 |
val force_list: 'a lazy list -> 'a list |
44386
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
21 |
val map: ('a -> 'b) -> 'a lazy -> 'b lazy |
44427 | 22 |
val future: Future.params -> 'a lazy -> 'a future |
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
23 |
end; |
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
24 |
|
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
25 |
structure Lazy: LAZY = |
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
26 |
struct |
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
27 |
|
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
28 |
(* datatype *) |
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
29 |
|
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
30 |
datatype 'a expr = |
66167 | 31 |
Expr of string * (unit -> 'a) | |
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
32 |
Result of 'a future; |
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
33 |
|
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
34 |
abstype 'a lazy = Lazy of 'a expr Synchronized.var |
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
35 |
with |
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
36 |
|
66167 | 37 |
fun lazy_name name e = Lazy (Synchronized.var "lazy" (Expr (name, e))); |
38 |
fun lazy e = lazy_name "lazy" e; |
|
59193
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
39 |
fun value a = Lazy (Synchronized.var "lazy" (Result (Future.value a))); |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
40 |
|
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
41 |
fun peek (Lazy var) = |
59195 | 42 |
(case Synchronized.value var of |
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
43 |
Expr _ => NONE |
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
44 |
| Result res => Future.peek res); |
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
45 |
|
66168 | 46 |
fun pending (Lazy var) = |
47 |
(case Synchronized.value var of |
|
48 |
Expr _ => true |
|
49 |
| Result _ => false); |
|
50 |
||
51 |
||
59193
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
52 |
|
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
53 |
(* status *) |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
54 |
|
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
55 |
fun is_future pred (Lazy var) = |
59191 | 56 |
(case Synchronized.value var of |
57 |
Expr _ => false |
|
59193
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
58 |
| Result res => pred res); |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
59 |
|
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
60 |
fun is_running x = is_future (not o Future.is_finished) x; |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
61 |
|
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
62 |
fun is_finished x = |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
63 |
is_future (fn res => |
59f1591a11cb
eliminated Document.execution frontier (again, see 627fb639a2d9): just run into older execution, potentially stalling worker thread, but without global delay due to long-running tasks (notably sledgehammer);
wenzelm
parents:
59191
diff
changeset
|
64 |
Future.is_finished res andalso not (Exn.is_interrupt_exn (Future.join_result res))) x; |
59191 | 65 |
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
66 |
|
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
67 |
(* force result *) |
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
68 |
|
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
69 |
fun force_result (Lazy var) = |
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
70 |
(case peek (Lazy var) of |
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
71 |
SOME res => res |
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
72 |
| NONE => |
62923 | 73 |
Thread_Attributes.uninterruptible (fn restore_attributes => fn () => |
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
74 |
let |
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
75 |
val (expr, x) = |
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
76 |
Synchronized.change_result var |
66167 | 77 |
(fn Expr (name, e) => |
78 |
let val x = Future.promise_name name I |
|
79 |
in ((SOME (name, e), x), Result x) end |
|
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
80 |
| Result x => ((NONE, x), Result x)); |
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
81 |
in |
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
82 |
(case expr of |
66167 | 83 |
SOME (name, e) => |
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
84 |
let |
41701
430493d65cd2
Lazy.force_result: more standard treatment of interruptibility, potentially addressing races of exceptions vs. interrupts;
wenzelm
parents:
40482
diff
changeset
|
85 |
val res0 = Exn.capture (restore_attributes e) (); |
47430
b254fdaf1973
partial revert of 8a179a0493e3 -- expose failure status of result (potentially via group) instead of isolated interrupt;
wenzelm
parents:
47423
diff
changeset
|
86 |
val _ = Exn.capture (fn () => Future.fulfill_result x res0) (); |
b254fdaf1973
partial revert of 8a179a0493e3 -- expose failure status of result (potentially via group) instead of isolated interrupt;
wenzelm
parents:
47423
diff
changeset
|
87 |
val res = Future.join_result x; |
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
88 |
(*semantic race: some other threads might see the same |
40389 | 89 |
interrupt, until there is a fresh start*) |
47430
b254fdaf1973
partial revert of 8a179a0493e3 -- expose failure status of result (potentially via group) instead of isolated interrupt;
wenzelm
parents:
47423
diff
changeset
|
90 |
val _ = |
b254fdaf1973
partial revert of 8a179a0493e3 -- expose failure status of result (potentially via group) instead of isolated interrupt;
wenzelm
parents:
47423
diff
changeset
|
91 |
if Exn.is_interrupt_exn res then |
66167 | 92 |
Synchronized.change var (fn _ => Expr (name, e)) |
47430
b254fdaf1973
partial revert of 8a179a0493e3 -- expose failure status of result (potentially via group) instead of isolated interrupt;
wenzelm
parents:
47423
diff
changeset
|
93 |
else (); |
b254fdaf1973
partial revert of 8a179a0493e3 -- expose failure status of result (potentially via group) instead of isolated interrupt;
wenzelm
parents:
47423
diff
changeset
|
94 |
in res end |
41701
430493d65cd2
Lazy.force_result: more standard treatment of interruptibility, potentially addressing races of exceptions vs. interrupts;
wenzelm
parents:
40482
diff
changeset
|
95 |
| NONE => Exn.capture (restore_attributes (fn () => Future.join x)) ()) |
37183
dae865999db5
force_result within the current execution context -- avoids overhead of potential thread context switch and robustifies Interrupt handling;
wenzelm
parents:
34278
diff
changeset
|
96 |
end) ()); |
29422 | 97 |
|
28673
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
98 |
|
d746a8c12c43
renamed structure Susp to Lazy, and Susp.delay to Lazy.lazy;
wenzelm
parents:
diff
changeset
|
99 |
end; |
44386
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
100 |
|
66168 | 101 |
fun force x = Exn.release (force_result x); |
102 |
||
103 |
fun force_list xs = |
|
104 |
(List.app (fn x => if pending x then ignore (force_result x) else ()) xs; |
|
105 |
map force xs); |
|
106 |
||
66167 | 107 |
fun map f x = lazy_name "Lazy.map" (fn () => f (force x)); |
44386
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
108 |
|
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
109 |
|
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
110 |
(* future evaluation *) |
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
111 |
|
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
112 |
fun future params x = |
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
113 |
if is_finished x then Future.value_result (force_result x) |
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
114 |
else (singleton o Future.forks) params (fn () => force x); |
4048ca2658b7
some support for toplevel printing wrt. editor perspective (still inactive);
wenzelm
parents:
44298
diff
changeset
|
115 |
|
62663 | 116 |
|
117 |
(* toplevel pretty printing *) |
|
118 |
||
119 |
val _ = |
|
62819
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
wenzelm
parents:
62663
diff
changeset
|
120 |
ML_system_pp (fn depth => fn pretty => fn x => |
62663 | 121 |
(case peek x of |
62823 | 122 |
NONE => PolyML_Pretty.PrettyString "<lazy>" |
123 |
| SOME (Exn.Exn _) => PolyML_Pretty.PrettyString "<failed>" |
|
62663 | 124 |
| SOME (Exn.Res y) => pretty (y, depth))); |
125 |
||
32815
1a5e364584ae
separate concurrent/sequential versions of lazy evaluation;
wenzelm
parents:
32738
diff
changeset
|
126 |
end; |
28971 | 127 |
|
128 |
type 'a lazy = 'a Lazy.lazy; |