src/Pure/Concurrent/simple_thread.ML
author wenzelm
Wed, 31 Dec 2008 19:54:04 +0100
changeset 29278 e9d148a808eb
parent 28577 bd2456e0d944
child 29564 f8b933a62151
permissions -rw-r--r--
added declare_term_frees; tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28241
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Concurrent/simple_thread.ML
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     4
28577
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
     5
Simplified thread operations.
28241
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     6
*)
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     7
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     8
signature SIMPLE_THREAD =
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
     9
sig
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    10
  val fork: bool -> (unit -> unit) -> Thread.thread
28550
422e9bd169ac added fail-safe interrupt;
wenzelm
parents: 28241
diff changeset
    11
  val interrupt: Thread.thread -> unit
28577
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    12
  val synchronized: string -> Mutex.mutex -> (unit -> 'a) -> 'a
28241
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    13
end;
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    14
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    15
structure SimpleThread: SIMPLE_THREAD =
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    16
struct
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    17
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    18
fun fork interrupts body =
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    19
  Thread.fork (fn () => exception_trace (fn () => body ()),
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    20
    if interrupts then Multithreading.regular_interrupts else Multithreading.no_interrupts);
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    21
28550
422e9bd169ac added fail-safe interrupt;
wenzelm
parents: 28241
diff changeset
    22
fun interrupt thread = Thread.interrupt thread handle Thread _ => ();
422e9bd169ac added fail-safe interrupt;
wenzelm
parents: 28241
diff changeset
    23
28577
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    24
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    25
(* basic synchronization *)
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    26
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    27
fun synchronized name lock e = Exn.release (uninterruptible (fn restore_attributes => fn () =>
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    28
  let
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    29
    val _ =
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    30
      if Mutex.trylock lock then ()
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    31
      else
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    32
       (Multithreading.tracing 3 (fn () => name ^ ": locking ...");
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    33
        Mutex.lock lock;
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    34
        Multithreading.tracing 3 (fn () => name ^ ": ... locked"));
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    35
    val result = Exn.capture (restore_attributes e) ();
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    36
    val _ = Mutex.unlock lock;
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    37
  in result end) ());
bd2456e0d944 added generic combinator for synchronized evaluation (formerly in future.ML);
wenzelm
parents: 28550
diff changeset
    38
28241
de20fccf6509 Simplified thread fork interface.
wenzelm
parents:
diff changeset
    39
end;