Synchronized.value does not require locking, since assigments are atomic;
removed obsolete Synchronized.peek;
--- a/src/Pure/Concurrent/future.ML Wed Sep 16 21:31:57 2009 +0200
+++ b/src/Pure/Concurrent/future.ML Wed Sep 16 22:46:10 2009 +0200
@@ -84,7 +84,7 @@
fun group_of (Future {group, ...}) = group;
fun result_of (Future {result, ...}) = result;
-fun peek x = Synchronized.peek (result_of x);
+fun peek x = Synchronized.value (result_of x);
fun is_finished x = is_some (peek x);
fun value x = Future
--- a/src/Pure/Concurrent/synchronized.ML Wed Sep 16 21:31:57 2009 +0200
+++ b/src/Pure/Concurrent/synchronized.ML Wed Sep 16 22:46:10 2009 +0200
@@ -8,7 +8,6 @@
sig
type 'a var
val var: string -> 'a -> 'a var
- val peek: 'a var -> 'a
val value: 'a var -> 'a
val timed_access: 'a var -> ('a -> Time.time option) -> ('a -> ('b * 'a) option) -> 'b option
val guarded_access: 'a var -> ('a -> ('b * 'a) option) -> 'b
@@ -33,9 +32,7 @@
cond = ConditionVar.conditionVar (),
var = ref x};
-fun peek (Var {var, ...}) = ! var; (*unsynchronized!*)
-
-fun value (Var {name, lock, cond, var}) = SimpleThread.synchronized name lock (fn () => ! var);
+fun value (Var {var, ...}) = ! var;
(* synchronized access *)