src/Pure/ML-Systems/multithreading_dummy.ML
author wenzelm
Tue, 24 Jul 2007 19:44:31 +0200
changeset 23960 c07ae96cbfc4
child 23973 b6ce6de5b700
permissions -rw-r--r--
Compatibility file for ML systems without multithreading.

(*  Title:      Pure/ML-Systems/multithreading_dummy.ML
    ID:         $Id$
    Author:     Makarius

Compatibility file for ML systems without multithreading.
*)

signature MULTITHREADING =
sig
  val number_of_threads: int ref
  val self_critical: unit -> bool
  val CRITICAL: (unit -> 'a) -> 'a
end;

structure Multithreading: MULTITHREADING =
struct

val number_of_threads = ref 0;

fun self_critical () = false;
fun CRITICAL e = e ();

end;

val CRITICAL = Multithreading.CRITICAL;