src/Tools/Metis/src/Portable.sig
author blanchet
Wed, 15 Sep 2010 19:22:34 +0200
changeset 39429 126b879df319
parent 39353 7f11d833d65b
child 39443 e330437cd22a
permissions -rw-r--r--
move "CRITICAL" to "PortableXxx", where it belongs and used to be; now that it is no longer used in "Random.sml"

(* ========================================================================= *)
(* ML SPECIFIC FUNCTIONS                                                     *)
(* Copyright (c) 2001-2007 Joe Hurd, distributed under the BSD License       *)
(* ========================================================================= *)

signature Portable =
sig

(* ------------------------------------------------------------------------- *)
(* The ML implementation.                                                    *)
(* ------------------------------------------------------------------------- *)

val ml : string

(* ------------------------------------------------------------------------- *)
(* Pointer equality using the run-time system.                               *)
(* ------------------------------------------------------------------------- *)

val pointerEqual : 'a * 'a -> bool

(* ------------------------------------------------------------------------- *)
(* Timing function applications.                                             *)
(* ------------------------------------------------------------------------- *)

val time : ('a -> 'b) -> 'a -> 'b

(* ------------------------------------------------------------------------- *)
(* Critical section markup (multiprocessing)                                 *)
(* ------------------------------------------------------------------------- *)

(* MODIFIED by Jasmin Blanchette *)
val CRITICAL: (unit -> 'a) -> 'a

(* ------------------------------------------------------------------------- *)
(* Generating random values.                                                 *)
(* ------------------------------------------------------------------------- *)

val randomBool : unit -> bool

val randomInt : int -> int  (* n |-> [0,n) *)

val randomReal : unit -> real  (* () |-> [0,1] *)

val randomWord : unit -> Word.word

end