| author | haftmann | 
| Tue, 16 Oct 2007 23:12:38 +0200 | |
| changeset 25061 | 250e1da3204b | 
| parent 24688 | a5754ca5c510 | 
| child 25732 | 308315ee2b6d | 
| permissions | -rw-r--r-- | 
| 4403 | 1  | 
(* Title: Pure/ML-Systems/smlnj.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
||
| 5708 | 4  | 
Compatibility file for Standard ML of New Jersey 110 or later.  | 
| 4403 | 5  | 
*)  | 
6  | 
||
| 
24599
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
7  | 
use "ML-Systems/proper_int.ML";  | 
| 
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
8  | 
use "ML-Systems/overloading_smlnj.ML";  | 
| 
23965
 
f93e509659c1
ML-Systems/exn.ML, ML-Systems/multithreading_dummy.ML;
 
wenzelm 
parents: 
23921 
diff
changeset
 | 
9  | 
use "ML-Systems/exn.ML";  | 
| 
24688
 
a5754ca5c510
replaced interrupt_timeout by TimeLimit.timeLimit (available on SML/NJ and Poly/ML 5.1);
 
wenzelm 
parents: 
24599 
diff
changeset
 | 
10  | 
use "ML-Systems/multithreading.ML";  | 
| 
23921
 
947152add153
added compatibility file for ML systems without multithreading;
 
wenzelm 
parents: 
23826 
diff
changeset
 | 
11  | 
|
| 16542 | 12  | 
|
13  | 
(*low-level pointer equality*)  | 
|
| 16502 | 14  | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
15  | 
CM.autoload "$smlnj/init/init.cmi";  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
16  | 
val pointer_eq = InlineT.ptreql;  | 
| 16528 | 17  | 
|
| 16502 | 18  | 
|
| 4403 | 19  | 
(* restore old-style character / string functions *)  | 
20  | 
||
| 24145 | 21  | 
val ord = mk_int o SML90.ord;  | 
22  | 
val chr = SML90.chr o dest_int;  | 
|
| 10725 | 23  | 
val explode = SML90.explode;  | 
24  | 
val implode = SML90.implode;  | 
|
| 4403 | 25  | 
|
26  | 
||
27  | 
(* New Jersey ML parameters *)  | 
|
28  | 
||
29  | 
val _ =  | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
30  | 
(Control.Print.printLength := 1000;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
31  | 
Control.Print.printDepth := 350;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
32  | 
Control.Print.stringDepth := 250;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
33  | 
Control.Print.signatures := 2;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
34  | 
Control.MC.matchRedundantError := false);  | 
| 4403 | 35  | 
|
36  | 
||
37  | 
(* Poly/ML emulation *)  | 
|
38  | 
||
| 
24599
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
39  | 
val exit = exit o dest_int;  | 
| 4403 | 40  | 
fun quit () = exit 0;  | 
41  | 
||
42  | 
(*limit the printing depth -- divided by 2 for comparibility with Poly/ML*)  | 
|
| 24329 | 43  | 
local  | 
| 
24599
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
44  | 
val depth = ref (10: int);  | 
| 24329 | 45  | 
in  | 
46  | 
fun get_print_depth () = ! depth;  | 
|
47  | 
fun print_depth n =  | 
|
48  | 
(depth := n;  | 
|
49  | 
Control.Print.printDepth := dest_int n div 2;  | 
|
50  | 
Control.Print.printLength := dest_int n);  | 
|
51  | 
end;  | 
|
| 4403 | 52  | 
|
| 5816 | 53  | 
(* compiler-independent timing functions *)  | 
| 4403 | 54  | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
55  | 
fun start_timing () =  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
56  | 
let val CPUtimer = Timer.startCPUTimer();  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
57  | 
val time = Timer.checkCPUTimer(CPUtimer)  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
58  | 
in (CPUtimer,time) end;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
59  | 
|
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
60  | 
fun end_timing (CPUtimer, {sys,usr}) =
 | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
61  | 
let open Time (*...for Time.toString, Time.+ and Time.- *)  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
62  | 
      val {sys=sys2,usr=usr2} = Timer.checkCPUTimer(CPUtimer)
 | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
63  | 
in "User " ^ toString (usr2-usr) ^  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
64  | 
" All "^ toString (sys2-sys + usr2-usr) ^  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
65  | 
" secs"  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
66  | 
handle Time => ""  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
67  | 
end;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
68  | 
|
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
69  | 
fun check_timer timer =  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
70  | 
let  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
71  | 
    val {sys, usr} = Timer.checkCPUTimer timer;
 | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
72  | 
val gc = Timer.checkGCTime timer; (* FIXME already included in usr? *)  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
73  | 
in (sys, usr, gc) end;  | 
| 4403 | 74  | 
|
75  | 
||
| 16660 | 76  | 
(*prompts*)  | 
| 4977 | 77  | 
fun ml_prompts p1 p2 =  | 
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
78  | 
(Control.primaryPrompt := p1; Control.secondaryPrompt := p2);  | 
| 4977 | 79  | 
|
| 17511 | 80  | 
(*dummy implementation*)  | 
| 16681 | 81  | 
fun profile (n: int) f x = f x;  | 
82  | 
||
| 17511 | 83  | 
(*dummy implementation*)  | 
| 16681 | 84  | 
fun exception_trace f = f ();  | 
| 4977 | 85  | 
|
| 18384 | 86  | 
(*dummy implementation*)  | 
87  | 
fun print x = x;  | 
|
| 16681 | 88  | 
|
| 23770 | 89  | 
(*dummy implementation*)  | 
90  | 
fun makestring x = "dummy string for SML New Jersey";  | 
|
91  | 
||
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
92  | 
|
| 24290 | 93  | 
(* toplevel pretty printing (see also Pure/pure_setup.ML) *)  | 
| 4403 | 94  | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
95  | 
fun make_pp path pprint =  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
96  | 
let  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
97  | 
open PrettyPrint;  | 
| 4403 | 98  | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
99  | 
fun pp pps obj =  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
100  | 
pprint obj  | 
| 24145 | 101  | 
(string pps, openHOVBox pps o Rel o dest_int,  | 
102  | 
          fn wd => break pps {nsp=dest_int wd, offset=0}, fn () => newline pps,
 | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
103  | 
fn () => closeBox pps);  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
104  | 
in (path, pp) end;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
105  | 
|
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
106  | 
fun install_pp (path, pp) = CompilerPPTable.install_pp path pp;  | 
| 4403 | 107  | 
|
108  | 
||
109  | 
(* ML command execution *)  | 
|
110  | 
||
| 
24599
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
111  | 
fun use_text (tune: string -> string) name (print, err) verbose txt =  | 
| 5090 | 112  | 
let  | 
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
113  | 
val ref out_orig = Control.Print.out;  | 
| 5090 | 114  | 
|
115  | 
val out_buffer = ref ([]: string list);  | 
|
116  | 
    val out = {say = (fn s => out_buffer := s :: ! out_buffer), flush = (fn () => ())};
 | 
|
| 10914 | 117  | 
fun output () =  | 
| 7890 | 118  | 
let val str = implode (rev (! out_buffer))  | 
| 10914 | 119  | 
in String.substring (str, 0, Int.max (0, size str - 1)) end;  | 
| 5090 | 120  | 
in  | 
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
121  | 
Control.Print.out := out;  | 
| 
24599
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
122  | 
Backend.Interact.useStream (TextIO.openString (tune txt)) handle exn =>  | 
| 22144 | 123  | 
(Control.Print.out := out_orig;  | 
124  | 
err ((if name = "" then "" else "Error in " ^ name ^ "\n") ^ output ()); raise exn);  | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
125  | 
Control.Print.out := out_orig;  | 
| 10914 | 126  | 
if verbose then print (output ()) else ()  | 
| 5090 | 127  | 
end;  | 
| 4403 | 128  | 
|
| 
24599
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
129  | 
fun use_file tune output verbose name =  | 
| 
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
130  | 
let  | 
| 
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
131  | 
val instream = TextIO.openIn name;  | 
| 
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
132  | 
val txt = TextIO.inputAll instream before TextIO.closeIn instream;  | 
| 
 
7b0ecf9a9055
use_text/file: tune text (cf. ML_Parse.fix_ints);
 
wenzelm 
parents: 
24329 
diff
changeset
 | 
133  | 
in use_text tune name output verbose txt end;  | 
| 21770 | 134  | 
|
135  | 
||
| 4403 | 136  | 
|
| 5816 | 137  | 
(** interrupts **)  | 
138  | 
||
| 12990 | 139  | 
exception Interrupt;  | 
140  | 
||
141  | 
fun ignore_interrupt f x =  | 
|
| 5816 | 142  | 
let  | 
| 12990 | 143  | 
val old_handler = Signals.setHandler (Signals.sigINT, Signals.IGNORE);  | 
| 
23965
 
f93e509659c1
ML-Systems/exn.ML, ML-Systems/multithreading_dummy.ML;
 
wenzelm 
parents: 
23921 
diff
changeset
 | 
144  | 
val result = Exn.capture f x;  | 
| 12990 | 145  | 
val _ = Signals.setHandler (Signals.sigINT, old_handler);  | 
| 
23965
 
f93e509659c1
ML-Systems/exn.ML, ML-Systems/multithreading_dummy.ML;
 
wenzelm 
parents: 
23921 
diff
changeset
 | 
146  | 
in Exn.release result end;  | 
| 5816 | 147  | 
|
| 12990 | 148  | 
fun raise_interrupt f x =  | 
149  | 
let  | 
|
| 5816 | 150  | 
val interrupted = ref false;  | 
| 
23965
 
f93e509659c1
ML-Systems/exn.ML, ML-Systems/multithreading_dummy.ML;
 
wenzelm 
parents: 
23921 
diff
changeset
 | 
151  | 
val result = ref (Exn.Result ());  | 
| 12990 | 152  | 
val old_handler = Signals.inqHandler Signals.sigINT;  | 
| 5816 | 153  | 
in  | 
| 12990 | 154  | 
SMLofNJ.Cont.callcc (fn cont =>  | 
155  | 
(Signals.setHandler (Signals.sigINT, Signals.HANDLER (fn _ => (interrupted := true; cont)));  | 
|
| 
23965
 
f93e509659c1
ML-Systems/exn.ML, ML-Systems/multithreading_dummy.ML;
 
wenzelm 
parents: 
23921 
diff
changeset
 | 
156  | 
result := Exn.capture f x));  | 
| 12990 | 157  | 
Signals.setHandler (Signals.sigINT, old_handler);  | 
| 
23965
 
f93e509659c1
ML-Systems/exn.ML, ML-Systems/multithreading_dummy.ML;
 
wenzelm 
parents: 
23921 
diff
changeset
 | 
158  | 
if ! interrupted then raise Interrupt else Exn.release (! result)  | 
| 12990 | 159  | 
end;  | 
| 5816 | 160  | 
|
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
161  | 
|
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
162  | 
(* basis library fixes *)  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
163  | 
|
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
164  | 
structure TextIO =  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
165  | 
struct  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
166  | 
open TextIO;  | 
| 
23139
 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 
wenzelm 
parents: 
22144 
diff
changeset
 | 
167  | 
fun inputLine is = TextIO.inputLine is  | 
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
168  | 
handle IO.Io _ => raise Interrupt;  | 
| 
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
169  | 
end;  | 
| 17511 | 170  | 
|
| 
18790
 
418131f631fc
interrupt_timeout now raises Interrupt instead of SML90.Interrupt
 
webertj 
parents: 
18760 
diff
changeset
 | 
171  | 
|
| 16542 | 172  | 
(** Signal handling: emulation of the Poly/ML Signal structure. Note that types  | 
| 
15702
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
173  | 
Posix.Signal.signal and Signals.signal differ **)  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
174  | 
|
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
175  | 
structure IsaSignal =  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
176  | 
struct  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
177  | 
|
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
178  | 
datatype sig_handle = SIG_DFL | SIG_IGN | SIG_HANDLE of Signals.signal -> unit;  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
179  | 
|
| 16542 | 180  | 
(*From the SML/NJ documentation: "HANDLER(f) installs a handler for a  | 
181  | 
signal. When signal is delivered to the process, the execution state  | 
|
182  | 
of the current thread will be bundled up as a continuation k, then  | 
|
183  | 
f(signal,n,k) will be called. The number n is the number of times  | 
|
184  | 
signal has been signalled since the last time f was invoked for it."*)  | 
|
| 
15702
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
185  | 
|
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
186  | 
fun toAction SIG_DFL = Signals.DEFAULT  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
187  | 
| toAction SIG_IGN = Signals.IGNORE  | 
| 16542 | 188  | 
| toAction (SIG_HANDLE iu) =  | 
| 
15702
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
189  | 
Signals.HANDLER (fn (signo,_,cont) => (iu signo; cont));  | 
| 16542 | 190  | 
|
| 
15702
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
191  | 
(*The types are correct, but I'm not sure about the semantics!*)  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
192  | 
fun fromAction Signals.DEFAULT = SIG_DFL  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
193  | 
| fromAction Signals.IGNORE = SIG_IGN  | 
| 16542 | 194  | 
| fromAction (Signals.HANDLER f) =  | 
| 
15702
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
195  | 
SIG_HANDLE (fn signo => SMLofNJ.Cont.callcc (fn k => (f (signo,0,k); ())));  | 
| 16542 | 196  | 
|
| 
15702
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
197  | 
(*Poly/ML version has type int * sig_handle -> sig_handle*)  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
198  | 
fun signal (signo, sh) = fromAction (Signals.setHandler (signo, toAction sh));  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
199  | 
|
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
200  | 
val usr1 = UnixSignals.sigUSR1  | 
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
201  | 
val usr2 = UnixSignals.sigUSR2  | 
| 17763 | 202  | 
val alrm = UnixSignals.sigALRM  | 
203  | 
val chld = UnixSignals.sigCHLD  | 
|
204  | 
val cont = UnixSignals.sigCONT  | 
|
205  | 
val int = UnixSignals.sigINT  | 
|
206  | 
val quit = UnixSignals.sigQUIT  | 
|
207  | 
val term = UnixSignals.sigTERM  | 
|
| 
15702
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
208  | 
|
| 
 
2677db44c795
new signalling primmitives for sml/nj compatibility
 
paulson 
parents: 
14981 
diff
changeset
 | 
209  | 
end;  | 
| 5816 | 210  | 
|
211  | 
||
| 4403 | 212  | 
(** OS related **)  | 
213  | 
||
| 
23826
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23770 
diff
changeset
 | 
214  | 
(* current directory *)  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23770 
diff
changeset
 | 
215  | 
|
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23770 
diff
changeset
 | 
216  | 
val cd = OS.FileSys.chDir;  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23770 
diff
changeset
 | 
217  | 
val pwd = OS.FileSys.getDir;  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23770 
diff
changeset
 | 
218  | 
|
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23770 
diff
changeset
 | 
219  | 
|
| 4403 | 220  | 
(* system command execution *)  | 
221  | 
||
222  | 
(*execute Unix command which doesn't take any input from stdin and  | 
|
223  | 
sends its output to stdout; could be done more easily by Unix.execute,  | 
|
224  | 
but that function doesn't use the PATH*)  | 
|
225  | 
fun execute command =  | 
|
226  | 
let  | 
|
227  | 
val tmp_name = OS.FileSys.tmpName ();  | 
|
228  | 
val is = (OS.Process.system (command ^ " > " ^ tmp_name); TextIO.openIn tmp_name);  | 
|
229  | 
val result = TextIO.inputAll is;  | 
|
230  | 
in  | 
|
231  | 
TextIO.closeIn is;  | 
|
232  | 
OS.FileSys.remove tmp_name;  | 
|
233  | 
result  | 
|
234  | 
end;  | 
|
235  | 
||
| 7855 | 236  | 
(*plain version; with return code*)  | 
| 24145 | 237  | 
val system = mk_int o OS.Process.system;  | 
| 7855 | 238  | 
|
| 4403 | 239  | 
|
| 17824 | 240  | 
(*Convert a process ID to a decimal string (chiefly for tracing)*)  | 
| 
21298
 
6d2306b2376d
tuned names of start_timing,/end_timing/check_timer;
 
wenzelm 
parents: 
18790 
diff
changeset
 | 
241  | 
fun string_of_pid pid =  | 
| 17824 | 242  | 
Word.fmt StringCvt.DEC (Word.fromLargeWord (Posix.Process.pidToWord pid));  | 
243  | 
||
244  | 
||
| 4403 | 245  | 
(* getenv *)  | 
246  | 
||
247  | 
fun getenv var =  | 
|
248  | 
(case OS.Process.getEnv var of  | 
|
249  | 
NONE => ""  | 
|
250  | 
| SOME txt => txt);  |