| author | blanchet | 
| Tue, 22 May 2018 17:15:02 +0200 | |
| changeset 68250 | c45067867860 | 
| parent 64556 | 851ae0e7b09c | 
| child 78719 | 89038d9ef77d | 
| permissions | -rw-r--r-- | 
| 
62508
 
d0b68218ea55
discontinued RAW session: bootstrap directly from isabelle_process RAW_ML_SYSTEM;
 
wenzelm 
parents: 
62505 
diff
changeset
 | 
1  | 
(* Title: Pure/ML/ml_compiler0.ML  | 
| 62817 | 2  | 
Author: Makarius  | 
| 31312 | 3  | 
|
| 62817 | 4  | 
Runtime compilation and evaluation for bootstrap.  | 
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
5  | 
Initial ML file operations.  | 
| 31312 | 6  | 
*)  | 
7  | 
||
| 62494 | 8  | 
signature ML_COMPILER0 =  | 
9  | 
sig  | 
|
| 62821 | 10  | 
type polyml_location = PolyML.location  | 
| 62494 | 11  | 
type context =  | 
12  | 
   {name_space: ML_Name_Space.T,
 | 
|
| 
62716
 
d80b9f4990e4
explicit print_depth for the sake of Spec_Check.determine_type;
 
wenzelm 
parents: 
62668 
diff
changeset
 | 
13  | 
print_depth: int option,  | 
| 62494 | 14  | 
here: int -> string -> string,  | 
15  | 
print: string -> unit,  | 
|
16  | 
error: string -> unit}  | 
|
| 62817 | 17  | 
val make_context: ML_Name_Space.T -> context  | 
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
18  | 
  val ML: context -> {debug: bool, file: string, line: int, verbose: bool} -> string -> unit
 | 
| 
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
19  | 
  val ML_file: context -> {debug: bool, verbose: bool} -> string -> unit
 | 
| 62494 | 20  | 
val debug_option: bool option -> bool  | 
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
21  | 
val ML_file_operations: (bool option -> string -> unit) ->  | 
| 
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
22  | 
    {ML_file: string -> unit, ML_file_debug: string -> unit, ML_file_no_debug: string -> unit}
 | 
| 62494 | 23  | 
end;  | 
24  | 
||
25  | 
structure ML_Compiler0: ML_COMPILER0 =  | 
|
26  | 
struct  | 
|
27  | 
||
| 62821 | 28  | 
type polyml_location = PolyML.location;  | 
29  | 
||
30  | 
||
| 62817 | 31  | 
(* global options *)  | 
32  | 
||
33  | 
val _ = PolyML.Compiler.reportUnreferencedIds := true;  | 
|
34  | 
val _ = PolyML.Compiler.reportExhaustiveHandlers := true;  | 
|
35  | 
val _ = PolyML.Compiler.printInAlphabeticalOrder := false;  | 
|
36  | 
val _ = PolyML.Compiler.maxInlineSize := 80;  | 
|
37  | 
||
38  | 
||
39  | 
(* context *)  | 
|
40  | 
||
| 62494 | 41  | 
type context =  | 
| 
62490
 
39d01eaf5292
ML debugger support in Pure (again, see 3565c9f407ec);
 
wenzelm 
parents: 
62489 
diff
changeset
 | 
42  | 
 {name_space: ML_Name_Space.T,
 | 
| 
62716
 
d80b9f4990e4
explicit print_depth for the sake of Spec_Check.determine_type;
 
wenzelm 
parents: 
62668 
diff
changeset
 | 
43  | 
print_depth: int option,  | 
| 62493 | 44  | 
here: int -> string -> string,  | 
| 
62490
 
39d01eaf5292
ML debugger support in Pure (again, see 3565c9f407ec);
 
wenzelm 
parents: 
62489 
diff
changeset
 | 
45  | 
print: string -> unit,  | 
| 
 
39d01eaf5292
ML debugger support in Pure (again, see 3565c9f407ec);
 
wenzelm 
parents: 
62489 
diff
changeset
 | 
46  | 
error: string -> unit};  | 
| 
 
39d01eaf5292
ML debugger support in Pure (again, see 3565c9f407ec);
 
wenzelm 
parents: 
62489 
diff
changeset
 | 
47  | 
|
| 62817 | 48  | 
fun make_context name_space : context =  | 
49  | 
 {name_space = name_space,
 | 
|
50  | 
print_depth = NONE,  | 
|
51  | 
here = fn line => fn file => " (line " ^ Int.toString line ^ " of \"" ^ file ^ "\")",  | 
|
52  | 
print = fn s => (TextIO.output (TextIO.stdOut, s ^ "\n"); TextIO.flushOut TextIO.stdOut),  | 
|
53  | 
error = fn s => error s};  | 
|
54  | 
||
55  | 
||
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
56  | 
(* ML file operations *)  | 
| 62817 | 57  | 
|
58  | 
local  | 
|
59  | 
||
| 31312 | 60  | 
fun drop_newline s =  | 
61  | 
if String.isSuffix "\n" s then String.substring (s, 0, size s - 1)  | 
|
62  | 
else s;  | 
|
63  | 
||
| 
62529
 
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
 
wenzelm 
parents: 
62508 
diff
changeset
 | 
64  | 
fun ml_input start_line name txt =  | 
| 62504 | 65  | 
let  | 
| 64556 | 66  | 
fun input line (#"\\" :: #"<" :: #"^" :: #"h" :: #"e" :: #"r" :: #"e" :: #">" :: cs) res =  | 
| 62504 | 67  | 
let val s = "(Position.line_file_only " ^ Int.toString line ^ " \"" ^ name ^ "\")"  | 
| 
62662
 
291cc01f56f5
@{make_string} is available during Pure bootstrap;
 
wenzelm 
parents: 
62529 
diff
changeset
 | 
68  | 
in input line cs (s :: res) end  | 
| 
 
291cc01f56f5
@{make_string} is available during Pure bootstrap;
 
wenzelm 
parents: 
62529 
diff
changeset
 | 
69  | 
      | input line (#"@" :: #"{" :: #"m" :: #"a" :: #"k" :: #"e" :: #"_" ::
 | 
| 
 
291cc01f56f5
@{make_string} is available during Pure bootstrap;
 
wenzelm 
parents: 
62529 
diff
changeset
 | 
70  | 
#"s" :: #"t" :: #"r" :: #"i" :: #"n" :: #"g" :: #"}" :: cs) res =  | 
| 
62900
 
c641bf9402fd
simplified default print_depth: context is usually available, in contrast to 0d295e339f52;
 
wenzelm 
parents: 
62899 
diff
changeset
 | 
71  | 
input line cs (ML_Pretty.make_string_fn :: res)  | 
| 
62877
 
741560a5283b
avoid malformed Isabelle symbols during bootstrap;
 
wenzelm 
parents: 
62821 
diff
changeset
 | 
72  | 
      | input line (#"\\" :: #"<" :: cs) res = input line cs ("\092\092<" :: res)
 | 
| 
62662
 
291cc01f56f5
@{make_string} is available during Pure bootstrap;
 
wenzelm 
parents: 
62529 
diff
changeset
 | 
73  | 
      | input line (#"\n" :: cs) res = input (line + 1) cs ("\n" :: res)
 | 
| 
 
291cc01f56f5
@{make_string} is available during Pure bootstrap;
 
wenzelm 
parents: 
62529 
diff
changeset
 | 
74  | 
| input line (c :: cs) res = input line cs (str c :: res)  | 
| 
 
291cc01f56f5
@{make_string} is available during Pure bootstrap;
 
wenzelm 
parents: 
62529 
diff
changeset
 | 
75  | 
| input _ [] res = rev res;  | 
| 
 
291cc01f56f5
@{make_string} is available during Pure bootstrap;
 
wenzelm 
parents: 
62529 
diff
changeset
 | 
76  | 
in String.concat (input start_line (String.explode txt) []) end;  | 
| 62504 | 77  | 
|
| 62817 | 78  | 
in  | 
79  | 
||
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
80  | 
fun ML {name_space, print_depth, here, print, error, ...} {line, file, verbose, debug} text =
 | 
| 31312 | 81  | 
let  | 
| 60955 | 82  | 
val current_line = Unsynchronized.ref line;  | 
| 
62529
 
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
 
wenzelm 
parents: 
62508 
diff
changeset
 | 
83  | 
val in_buffer = Unsynchronized.ref (String.explode (ml_input line file text));  | 
| 32738 | 84  | 
val out_buffer = Unsynchronized.ref ([]: string list);  | 
| 31312 | 85  | 
fun output () = drop_newline (implode (rev (! out_buffer)));  | 
86  | 
||
87  | 
fun get () =  | 
|
88  | 
(case ! in_buffer of  | 
|
89  | 
[] => NONE  | 
|
| 60955 | 90  | 
| c :: cs =>  | 
91  | 
(in_buffer := cs; if c = #"\n" then current_line := ! current_line + 1 else (); SOME c));  | 
|
| 31312 | 92  | 
fun put s = out_buffer := s :: ! out_buffer;  | 
| 60955 | 93  | 
    fun put_message {message = msg1, hard, location = {startLine = message_line, ...}, context} =
 | 
| 31312 | 94  | 
(put (if hard then "Error: " else "Warning: ");  | 
95  | 
PolyML.prettyPrint (put, 76) msg1;  | 
|
96  | 
(case context of NONE => () | SOME msg2 => PolyML.prettyPrint (put, 76) msg2);  | 
|
| 62493 | 97  | 
      put ("At" ^ here (FixedInt.toInt message_line) file ^ "\n"));
 | 
| 31312 | 98  | 
|
99  | 
val parameters =  | 
|
100  | 
[PolyML.Compiler.CPOutStream put,  | 
|
| 31471 | 101  | 
PolyML.Compiler.CPNameSpace name_space,  | 
| 31312 | 102  | 
PolyML.Compiler.CPErrorMessageProc put_message,  | 
| 60955 | 103  | 
PolyML.Compiler.CPLineNo (fn () => ! current_line),  | 
104  | 
PolyML.Compiler.CPFileName file,  | 
|
| 62501 | 105  | 
PolyML.Compiler.CPPrintInAlphabeticalOrder false,  | 
| 
62716
 
d80b9f4990e4
explicit print_depth for the sake of Spec_Check.determine_type;
 
wenzelm 
parents: 
62668 
diff
changeset
 | 
106  | 
PolyML.Compiler.CPDebug debug] @  | 
| 
 
d80b9f4990e4
explicit print_depth for the sake of Spec_Check.determine_type;
 
wenzelm 
parents: 
62668 
diff
changeset
 | 
107  | 
(case print_depth of NONE => [] | SOME d => [PolyML.Compiler.CPPrintDepth (fn () => d)]);  | 
| 31312 | 108  | 
val _ =  | 
109  | 
(while not (List.null (! in_buffer)) do  | 
|
110  | 
PolyML.compiler (get, parameters) ())  | 
|
111  | 
handle exn =>  | 
|
| 62505 | 112  | 
if Exn.is_interrupt exn then Exn.reraise exn  | 
| 
39242
 
28d3809ff91f
primitive use_text: let interrupts pass unhindered;
 
wenzelm 
parents: 
38470 
diff
changeset
 | 
113  | 
else  | 
| 
 
28d3809ff91f
primitive use_text: let interrupts pass unhindered;
 
wenzelm 
parents: 
38470 
diff
changeset
 | 
114  | 
         (put ("Exception- " ^ General.exnMessage exn ^ " raised");
 | 
| 62505 | 115  | 
error (output ()); Exn.reraise exn);  | 
| 31312 | 116  | 
in if verbose then print (output ()) else () end;  | 
117  | 
||
| 62817 | 118  | 
end;  | 
119  | 
||
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
120  | 
fun ML_file context {verbose, debug} file =
 | 
| 31312 | 121  | 
let  | 
| 60955 | 122  | 
val instream = TextIO.openIn file;  | 
| 60956 | 123  | 
val text = Exn.release (Exn.capture TextIO.inputAll instream before TextIO.closeIn instream);  | 
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
124  | 
  in ML context {line = 1, file = file, verbose = verbose, debug = debug} text end;
 | 
| 31312 | 125  | 
|
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
126  | 
fun ML_file_operations (f: bool option -> string -> unit) =  | 
| 
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
127  | 
  {ML_file = f NONE, ML_file_debug = f (SOME true), ML_file_no_debug = f (SOME false)};
 | 
| 62817 | 128  | 
|
| 
62490
 
39d01eaf5292
ML debugger support in Pure (again, see 3565c9f407ec);
 
wenzelm 
parents: 
62489 
diff
changeset
 | 
129  | 
|
| 62494 | 130  | 
fun debug_option NONE = OS.Process.getEnv "ISABELLE_ML_DEBUGGER" = SOME "true"  | 
131  | 
| debug_option (SOME debug) = debug;  | 
|
| 
62490
 
39d01eaf5292
ML debugger support in Pure (again, see 3565c9f407ec);
 
wenzelm 
parents: 
62489 
diff
changeset
 | 
132  | 
|
| 62817 | 133  | 
end;  | 
| 62489 | 134  | 
|
| 62817 | 135  | 
|
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
136  | 
(* initial ML_file operations *)  | 
| 62494 | 137  | 
|
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
138  | 
val {ML_file, ML_file_debug, ML_file_no_debug} =
 | 
| 62817 | 139  | 
let val context = ML_Compiler0.make_context ML_Name_Space.global in  | 
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
140  | 
ML_Compiler0.ML_file_operations (fn opt_debug => fn file =>  | 
| 
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
141  | 
ML_Compiler0.ML_file context  | 
| 62494 | 142  | 
        {verbose = true, debug = ML_Compiler0.debug_option opt_debug} file
 | 
143  | 
handle ERROR msg => (#print context msg; raise error "ML error"))  | 
|
144  | 
end;  | 
|
| 
62819
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62817 
diff
changeset
 | 
145  | 
|
| 
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62817 
diff
changeset
 | 
146  | 
|
| 
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62817 
diff
changeset
 | 
147  | 
(* export type-dependent functions *)  | 
| 
 
d3ff367a16a0
careful export of type-dependent functions, without losing their special status;
 
wenzelm 
parents: 
62817 
diff
changeset
 | 
148  | 
|
| 62894 | 149  | 
if Thread_Data.is_virtual then ()  | 
150  | 
else  | 
|
| 
62902
 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 
wenzelm 
parents: 
62900 
diff
changeset
 | 
151  | 
ML_Compiler0.ML  | 
| 62894 | 152  | 
(ML_Compiler0.make_context  | 
153  | 
(ML_Name_Space.global_values  | 
|
154  | 
        [("prettyRepresentation", "ML_system_pretty"),
 | 
|
155  | 
         ("addPrettyPrinter", "ML_system_pp"),
 | 
|
156  | 
         ("addOverload", "ML_system_overload")]))
 | 
|
157  | 
    {debug = false, file = "", line = 0, verbose = false}
 | 
|
158  | 
"open PolyML RunCall";  | 
|
| 62918 | 159  | 
|
160  | 
ML_system_pp (fn depth => fn pretty => fn var => pretty (Synchronized.value var, depth));  |