| author | wenzelm | 
| Thu, 21 Aug 2008 17:42:59 +0200 | |
| changeset 27940 | 002718f9c938 | 
| parent 27004 | 616c553c7cf1 | 
| child 28488 | 18fea7e88ea1 | 
| permissions | -rw-r--r-- | 
| 22563 | 1  | 
(* Title: Pure/ML-Systems/alice.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
||
| 22837 | 4  | 
Compatibility file for Alice 1.4.  | 
| 
23826
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
5  | 
|
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
6  | 
NOTE: there is no wrapper script; may run it interactively as follows:  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
7  | 
|
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
8  | 
$ cd Isabelle/src/Pure  | 
| 27004 | 9  | 
$ env ALICE_JIT_MODE=0 ISABELLE_HOME=$(cd ../..; pwd) alice  | 
| 23835 | 10  | 
- val ml_system = "alice";  | 
| 24807 | 11  | 
- use "ML-Systems/exn.ML";  | 
| 26226 | 12  | 
- use "ML-Systems/universal.ML";  | 
| 24807 | 13  | 
- use "ML-Systems/multithreading.ML";  | 
14  | 
- use "ML-Systems/time_limit.ML";  | 
|
| 
23826
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
15  | 
- use "ML-Systems/alice.ML";  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
16  | 
- use "ROOT.ML";  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
17  | 
- Session.finish ();  | 
| 22563 | 18  | 
*)  | 
19  | 
||
| 24597 | 20  | 
val ml_system_fix_ints = false;  | 
21  | 
||
| 26474 | 22  | 
fun forget_structure _ = ();  | 
23  | 
||
| 22563 | 24  | 
fun exit 0 = (OS.Process.exit OS.Process.success): unit  | 
25  | 
| exit _ = OS.Process.exit OS.Process.failure;  | 
|
26  | 
||
27  | 
||
28  | 
(** ML system related **)  | 
|
29  | 
||
30  | 
(*low-level pointer equality*)  | 
|
31  | 
fun pointer_eq (_: 'a, _: 'a) = false;  | 
|
32  | 
||
| 
24809
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
33  | 
|
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
34  | 
(* integer compatibility -- downgraded IntInf *)  | 
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
35  | 
|
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
36  | 
structure Time =  | 
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
37  | 
struct  | 
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
38  | 
open Time;  | 
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
39  | 
val fromMilliseconds = Time.fromMilliseconds o IntInf.fromInt;  | 
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
40  | 
val fromSeconds = Time.fromSeconds o IntInf.fromInt;  | 
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
41  | 
end;  | 
| 
 
41a21f59f74d
integer compatibility: added wrapper for structure Time;
 
wenzelm 
parents: 
24807 
diff
changeset
 | 
42  | 
|
| 24807 | 43  | 
structure IntInf =  | 
44  | 
struct  | 
|
45  | 
fun divMod (x, y) = (x div y, x mod y);  | 
|
46  | 
open Int;  | 
|
47  | 
end;  | 
|
48  | 
||
| 22563 | 49  | 
|
50  | 
(* restore old-style character / string functions *)  | 
|
51  | 
||
52  | 
exception Ord;  | 
|
53  | 
fun ord "" = raise Ord  | 
|
54  | 
| ord s = Char.ord (String.sub (s, 0));  | 
|
55  | 
||
| 22575 | 56  | 
val chr = String.str o chr;  | 
57  | 
val explode = map String.str o String.explode;  | 
|
| 22563 | 58  | 
val implode = String.concat;  | 
59  | 
||
60  | 
||
61  | 
(* Poly/ML emulation *)  | 
|
62  | 
||
63  | 
fun quit () = exit 0;  | 
|
64  | 
||
| 24329 | 65  | 
fun get_print_depth () = ! Print.depth;  | 
| 22563 | 66  | 
fun print_depth n = Print.depth := n;  | 
67  | 
||
68  | 
||
69  | 
(* compiler-independent timing functions *)  | 
|
70  | 
||
71  | 
structure Timer =  | 
|
72  | 
struct  | 
|
73  | 
open Timer;  | 
|
74  | 
type cpu_timer = unit;  | 
|
75  | 
fun startCPUTimer () = ();  | 
|
76  | 
  fun checkCPUTimer () = {sys = Time.zeroTime, usr = Time.zeroTime};
 | 
|
77  | 
fun checkGCTime () = Time.zeroTime;  | 
|
78  | 
end;  | 
|
79  | 
||
80  | 
fun start_timing () =  | 
|
81  | 
let val CPUtimer = Timer.startCPUTimer();  | 
|
82  | 
val time = Timer.checkCPUTimer(CPUtimer)  | 
|
83  | 
in (CPUtimer,time) end;  | 
|
84  | 
||
85  | 
fun end_timing (CPUtimer, {sys,usr}) =
 | 
|
86  | 
let open Time (*...for Time.toString, Time.+ and Time.- *)  | 
|
87  | 
      val {sys=sys2,usr=usr2} = Timer.checkCPUTimer(CPUtimer)
 | 
|
88  | 
in "User " ^ toString (usr2-usr) ^  | 
|
89  | 
" All "^ toString (sys2-sys + usr2-usr) ^  | 
|
90  | 
" secs"  | 
|
91  | 
handle Time => ""  | 
|
92  | 
end;  | 
|
93  | 
||
94  | 
fun check_timer timer =  | 
|
95  | 
let  | 
|
96  | 
    val {sys, usr} = Timer.checkCPUTimer timer;
 | 
|
97  | 
val gc = Timer.checkGCTime timer; (* FIXME already included in usr? *)  | 
|
98  | 
in (sys, usr, gc) end;  | 
|
99  | 
||
100  | 
||
101  | 
(*prompts*)  | 
|
102  | 
fun ml_prompts p1 p2 = ();  | 
|
103  | 
||
104  | 
(*dummy implementation*)  | 
|
105  | 
fun profile (n: int) f x = f x;  | 
|
106  | 
||
107  | 
(*dummy implementation*)  | 
|
108  | 
fun exception_trace f = f ();  | 
|
109  | 
||
110  | 
(*dummy implementation*)  | 
|
111  | 
fun print x = x;  | 
|
112  | 
||
113  | 
||
| 24290 | 114  | 
(* toplevel pretty printing (see also Pure/pure_setup.ML) *)  | 
| 22563 | 115  | 
|
116  | 
fun make_pp path pprint = (path, pprint);  | 
|
117  | 
fun install_pp (path, pp) = ();  | 
|
118  | 
||
119  | 
||
120  | 
(* ML command execution *)  | 
|
121  | 
||
| 26884 | 122  | 
fun use_text _ _ _ _ _ txt = (Compiler.eval txt; ());  | 
123  | 
fun use_file _ _ _ _ name = use name;  | 
|
| 22563 | 124  | 
|
125  | 
||
126  | 
||
127  | 
(** interrupts **)  | 
|
128  | 
||
129  | 
exception Interrupt;  | 
|
130  | 
||
| 
26084
 
a7475459c740
replaced ignore/raise_interrupt by more flexible (un)interruptible combinators;
 
wenzelm 
parents: 
24809 
diff
changeset
 | 
131  | 
fun interruptible f x = f x;  | 
| 
 
a7475459c740
replaced ignore/raise_interrupt by more flexible (un)interruptible combinators;
 
wenzelm 
parents: 
24809 
diff
changeset
 | 
132  | 
fun uninterruptible f x = f (fn (g: 'c -> 'd) => g) x;  | 
| 22563 | 133  | 
|
134  | 
||
135  | 
(* basis library fixes *)  | 
|
136  | 
||
137  | 
structure TextIO =  | 
|
138  | 
struct  | 
|
139  | 
open TextIO;  | 
|
| 
23139
 
aa899bce7c3b
TextIO.inputLine: use present SML B library version;
 
wenzelm 
parents: 
22837 
diff
changeset
 | 
140  | 
fun inputLine is = TextIO.inputLine is  | 
| 22563 | 141  | 
handle IO.Io _ => raise Interrupt;  | 
142  | 
end;  | 
|
143  | 
||
144  | 
||
145  | 
||
146  | 
(** OS related **)  | 
|
147  | 
||
| 27004 | 148  | 
structure OS =  | 
149  | 
struct  | 
|
150  | 
open OS;  | 
|
151  | 
structure FileSys =  | 
|
152  | 
struct  | 
|
153  | 
open FileSys;  | 
|
154  | 
fun tmpName () =  | 
|
155  | 
let val name = FileSys.tmpName () in  | 
|
156  | 
if String.isSuffix "\000" name  | 
|
157  | 
then String.substring (name, 0, size name - 1)  | 
|
158  | 
else name  | 
|
159  | 
end;  | 
|
160  | 
end;  | 
|
161  | 
end;  | 
|
162  | 
||
| 
23826
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
163  | 
val cd = OS.FileSys.chDir;  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
164  | 
val pwd = OS.FileSys.getDir;  | 
| 
 
463903573934
moved cd/pwd to ML compatibility layer (simplifies bootstrapping with Alice);
 
wenzelm 
parents: 
23139 
diff
changeset
 | 
165  | 
|
| 26226 | 166  | 
local  | 
167  | 
||
168  | 
fun read_file name =  | 
|
169  | 
let val is = TextIO.openIn name  | 
|
| 26504 | 170  | 
in Exn.release (Exn.capture TextIO.inputAll is before TextIO.closeIn is) end;  | 
| 26226 | 171  | 
|
172  | 
fun write_file name txt =  | 
|
173  | 
let val os = TextIO.openOut name  | 
|
| 26504 | 174  | 
in Exn.release (Exn.capture TextIO.output (os, txt) before TextIO.closeOut os) end;  | 
| 26226 | 175  | 
|
176  | 
in  | 
|
177  | 
||
178  | 
fun system_out script =  | 
|
179  | 
let  | 
|
180  | 
val script_name = OS.FileSys.tmpName ();  | 
|
181  | 
val _ = write_file script_name script;  | 
|
182  | 
||
183  | 
val output_name = OS.FileSys.tmpName ();  | 
|
184  | 
||
185  | 
val status =  | 
|
186  | 
      OS.Process.system ("perl -w \"$ISABELLE_HOME/lib/scripts/system.pl\" nogroup " ^
 | 
|
187  | 
script_name ^ " /dev/null " ^ output_name);  | 
|
188  | 
val rc = if OS.Process.isSuccess status then 0 else 1;  | 
|
189  | 
||
190  | 
val output = read_file output_name handle IO.Io _ => "";  | 
|
191  | 
val _ = OS.FileSys.remove script_name handle OS.SysErr _ => ();  | 
|
192  | 
val _ = OS.FileSys.remove output_name handle OS.SysErr _ => ();  | 
|
193  | 
in (output, rc) end;  | 
|
194  | 
||
195  | 
end;  | 
|
| 22563 | 196  | 
|
197  | 
structure OS =  | 
|
198  | 
struct  | 
|
199  | 
open OS;  | 
|
200  | 
structure FileSys =  | 
|
201  | 
struct  | 
|
202  | 
fun fileId name =  | 
|
| 26226 | 203  | 
      (case system_out ("perl -e '@_ = stat(q:" ^ name ^ ":); print $_[1]'") of
 | 
204  | 
        ("", _) => raise Fail "OS.FileSys.fileId"   (* FIXME IO.Io!? *)
 | 
|
205  | 
| (s, _) => (case Int.fromString s of NONE => raise Fail "OS.FileSys.fileId" | SOME i => i));  | 
|
| 22563 | 206  | 
val compare = Int.compare;  | 
| 27004 | 207  | 
fun fullPath name =  | 
208  | 
      (case system_out ("FILE='" ^ name ^
 | 
|
209  | 
"' && cd \"$(dirname \"$FILE\")\" && echo -n \"$(pwd -P)/$(basename \"$FILE\")\"") of  | 
|
210  | 
        ("", _) => raise SysErr ("Bad file", NONE)
 | 
|
211  | 
| (s, _) => s);  | 
|
| 22563 | 212  | 
open FileSys;  | 
213  | 
end;  | 
|
214  | 
end;  | 
|
215  | 
||
| 27004 | 216  | 
structure Posix =  | 
217  | 
struct  | 
|
218  | 
structure ProcEnv =  | 
|
219  | 
struct  | 
|
220  | 
fun getpid () = raise Fail "Posix.ProcEnv.getpoid undefined";  | 
|
221  | 
end;  | 
|
222  | 
end;  | 
|
223  | 
||
| 26227 | 224  | 
fun string_of_pid _ = raise Fail "string_of_pid undefined";  | 
225  | 
||
| 22563 | 226  | 
fun getenv var =  | 
227  | 
(case OS.Process.getEnv var of  | 
|
228  | 
NONE => ""  | 
|
229  | 
| SOME txt => txt);  |