| author | blanchet | 
| Fri, 26 Aug 2011 10:25:13 +0200 | |
| changeset 44509 | 369e8c28a61a | 
| parent 44389 | a3b5fdfb04a3 | 
| child 44731 | 8f7b3a89fc15 | 
| permissions | -rw-r--r-- | 
| 
30173
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29522 
diff
changeset
 | 
1  | 
(* Title: Pure/System/isabelle_process.ML  | 
| 
25528
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
38445
 
ba9ea6b9b75c
simplified internal message format: dropped special Symbol.STX header;
 
wenzelm 
parents: 
38270 
diff
changeset
 | 
4  | 
Isabelle process wrapper, based on private fifos for maximum  | 
| 
 
ba9ea6b9b75c
simplified internal message format: dropped special Symbol.STX header;
 
wenzelm 
parents: 
38270 
diff
changeset
 | 
5  | 
robustness and performance.  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
6  | 
|
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
7  | 
Startup phases:  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
8  | 
. raw Posix process startup with uncontrolled output on stdout/stderr  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
9  | 
. stdout \002: ML running  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
10  | 
.. stdin/stdout/stderr freely available (raw ML loop)  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
11  | 
.. protocol thread initialization  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
12  | 
... switch to in_fifo/out_fifo channels (rendezvous via open)  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
13  | 
... out_fifo INIT(pid): channels ready  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
14  | 
... out_fifo STATUS(keywords)  | 
| 
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
15  | 
... out_fifo READY: main loop ready  | 
| 
25528
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
*)  | 
| 
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
signature ISABELLE_PROCESS =  | 
| 
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
sig  | 
| 42897 | 20  | 
val is_active: unit -> bool  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
21  | 
val add_command: string -> (string list -> unit) -> unit  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
22  | 
val command: string -> string list -> unit  | 
| 43684 | 23  | 
val crashes: exn list Synchronized.var  | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
24  | 
val init: string -> string -> unit  | 
| 
25528
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
end;  | 
| 
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
|
| 31797 | 27  | 
structure Isabelle_Process: ISABELLE_PROCESS =  | 
| 
25528
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
struct  | 
| 
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
|
| 42897 | 30  | 
(* print mode *)  | 
| 25554 | 31  | 
|
| 25748 | 32  | 
val isabelle_processN = "isabelle_process";  | 
33  | 
||
| 42897 | 34  | 
fun is_active () = Print_Mode.print_mode_active isabelle_processN;  | 
35  | 
||
| 26550 | 36  | 
val _ = Output.add_mode isabelle_processN Output.default_output Output.default_escape;  | 
| 28036 | 37  | 
val _ = Markup.add_mode isabelle_processN YXML.output_markup;  | 
| 25841 | 38  | 
|
39  | 
||
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
40  | 
(* commands *)  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
41  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
42  | 
local  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
43  | 
|
| 43684 | 44  | 
val commands =  | 
45  | 
Synchronized.var "Isabelle_Process.commands" (Symtab.empty: (string list -> unit) Symtab.table);  | 
|
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
46  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
47  | 
in  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
48  | 
|
| 43684 | 49  | 
fun add_command name cmd =  | 
50  | 
Synchronized.change commands (fn cmds =>  | 
|
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
51  | 
(if not (Symtab.defined cmds name) then ()  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
52  | 
    else warning ("Redefining Isabelle process command " ^ quote name);
 | 
| 43684 | 53  | 
Symtab.update (name, cmd) cmds));  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
54  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
55  | 
fun command name args =  | 
| 43684 | 56  | 
(case Symtab.lookup (Synchronized.value commands) name of  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
57  | 
    NONE => error ("Undefined Isabelle process command " ^ quote name)
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
58  | 
| SOME cmd => cmd args);  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
59  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
60  | 
end;  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
61  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
62  | 
|
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
63  | 
(* message channels *)  | 
| 25841 | 64  | 
|
65  | 
local  | 
|
| 25810 | 66  | 
|
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
67  | 
fun chunk s = [string_of_int (size s), "\n", s];  | 
| 
26574
 
560d07845442
support "YXML" mode for output transfer notation;
 
wenzelm 
parents: 
26550 
diff
changeset
 | 
68  | 
|
| 
43771
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
69  | 
fun message mbox ch raw_props body =  | 
| 
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
70  | 
let  | 
| 
43772
 
c825594fd0c1
clarified YXML.embed_controls -- this is idempotent and cannot be nested;
 
wenzelm 
parents: 
43771 
diff
changeset
 | 
71  | 
val robust_props = map (pairself YXML.embed_controls) raw_props;  | 
| 
43771
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
72  | 
val header = YXML.string_of (XML.Elem ((ch, robust_props), []));  | 
| 
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
73  | 
in Mailbox.send mbox (chunk header @ chunk body) end;  | 
| 25554 | 74  | 
|
| 
44270
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
75  | 
fun standard_message mbox opt_serial ch body =  | 
| 
43771
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
76  | 
if body = "" then ()  | 
| 
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
77  | 
else  | 
| 
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
78  | 
message mbox ch  | 
| 
44270
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
79  | 
((case opt_serial of SOME i => cons (Markup.serialN, string_of_int i) | _ => I)  | 
| 
43771
 
fc524449f511
allow empty body for raw_message -- important for Invoke_Scala;
 
wenzelm 
parents: 
43746 
diff
changeset
 | 
80  | 
(Position.properties_of (Position.thread_data ()))) body;  | 
| 25554 | 81  | 
|
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
82  | 
fun message_output mbox out_stream =  | 
| 28188 | 83  | 
let  | 
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
84  | 
fun loop receive =  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
85  | 
(case receive mbox of  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
86  | 
SOME msg =>  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
87  | 
(List.app (fn s => TextIO.output (out_stream, s)) msg;  | 
| 40301 | 88  | 
loop (Mailbox.receive_timeout (seconds 0.02)))  | 
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
89  | 
| NONE => (try TextIO.flushOut out_stream; loop (SOME o Mailbox.receive)));  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
90  | 
in fn () => loop (SOME o Mailbox.receive) end;  | 
| 28188 | 91  | 
|
92  | 
in  | 
|
93  | 
||
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
94  | 
fun setup_channels in_fifo out_fifo =  | 
| 
28189
 
fbad2eb5be9e
auto_flush: uniform block buffering for all output streams;
 
wenzelm 
parents: 
28188 
diff
changeset
 | 
95  | 
let  | 
| 
39530
 
16adc476348f
Isabelle_Process: more robust rendezvous, even without proper blocking on open (Cygwin);
 
wenzelm 
parents: 
39528 
diff
changeset
 | 
96  | 
(*rendezvous*)  | 
| 
 
16adc476348f
Isabelle_Process: more robust rendezvous, even without proper blocking on open (Cygwin);
 
wenzelm 
parents: 
39528 
diff
changeset
 | 
97  | 
val in_stream = TextIO.openIn in_fifo;  | 
| 
 
16adc476348f
Isabelle_Process: more robust rendezvous, even without proper blocking on open (Cygwin);
 
wenzelm 
parents: 
39528 
diff
changeset
 | 
98  | 
val out_stream = TextIO.openOut out_fifo;  | 
| 
 
16adc476348f
Isabelle_Process: more robust rendezvous, even without proper blocking on open (Cygwin);
 
wenzelm 
parents: 
39528 
diff
changeset
 | 
99  | 
|
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
100  | 
val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdOut, IO.LINE_BUF);  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
101  | 
val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdErr, IO.LINE_BUF);  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
102  | 
|
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
103  | 
val mbox = Mailbox.create () : string list Mailbox.T;  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
104  | 
val _ = Simple_Thread.fork false (message_output mbox out_stream);  | 
| 28044 | 105  | 
in  | 
| 
44270
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
106  | 
Output.Private_Hooks.status_fn := standard_message mbox NONE "B";  | 
| 
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
107  | 
Output.Private_Hooks.report_fn := standard_message mbox NONE "C";  | 
| 
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
108  | 
Output.Private_Hooks.writeln_fn := (fn s => standard_message mbox (SOME (serial ())) "D" s);  | 
| 
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
109  | 
Output.Private_Hooks.tracing_fn := (fn s => standard_message mbox (SOME (serial ())) "E" s);  | 
| 
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
110  | 
Output.Private_Hooks.warning_fn := (fn s => standard_message mbox (SOME (serial ())) "F" s);  | 
| 
 
3eaad39e520c
more careful treatment of exception serial numbers, with propagation to message channel;
 
wenzelm 
parents: 
43772 
diff
changeset
 | 
111  | 
Output.Private_Hooks.error_fn := (fn (i, s) => standard_message mbox (SOME i) "G" s);  | 
| 
43746
 
a41f618c641d
some support for raw messages, which bypass standard Symbol/YXML decoding;
 
wenzelm 
parents: 
43684 
diff
changeset
 | 
112  | 
Output.Private_Hooks.raw_message_fn := message mbox "H";  | 
| 
40133
 
b61d52de66f0
more explicit indication of Output.Private_Hooks (still public to accomodate clones of Proof General for now);
 
wenzelm 
parents: 
40132 
diff
changeset
 | 
113  | 
Output.Private_Hooks.urgent_message_fn := ! Output.Private_Hooks.writeln_fn;  | 
| 
 
b61d52de66f0
more explicit indication of Output.Private_Hooks (still public to accomodate clones of Proof General for now);
 
wenzelm 
parents: 
40132 
diff
changeset
 | 
114  | 
Output.Private_Hooks.prompt_fn := ignore;  | 
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
115  | 
message mbox "A" [] (Session.welcome ());  | 
| 
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
116  | 
in_stream  | 
| 28044 | 117  | 
end;  | 
| 25841 | 118  | 
|
| 28188 | 119  | 
end;  | 
120  | 
||
| 25841 | 121  | 
|
| 
39234
 
d76a2fd129b5
main command loops are supposed to be uninterruptible -- no special treatment here;
 
wenzelm 
parents: 
38871 
diff
changeset
 | 
122  | 
(* protocol loop -- uninterruptible *)  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
123  | 
|
| 43684 | 124  | 
val crashes = Synchronized.var "Isabelle_Process.crashes" ([]: exn list);  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
125  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
126  | 
local  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
127  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
128  | 
fun recover crash =  | 
| 43684 | 129  | 
(Synchronized.change crashes (cons crash);  | 
| 
38270
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
130  | 
warning "Recovering from Isabelle process crash -- see also Isabelle_Process.crashes");  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
131  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
132  | 
fun read_chunk stream len =  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
133  | 
let  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
134  | 
val n =  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
135  | 
(case Int.fromString len of  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
136  | 
SOME n => n  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
137  | 
      | NONE => error ("Isabelle process: malformed chunk header " ^ quote len));
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
138  | 
val chunk = TextIO.inputN (stream, n);  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
139  | 
val m = size chunk;  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
140  | 
in  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
141  | 
if m = n then chunk  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
142  | 
    else error ("Isabelle process: bad chunk (" ^ string_of_int m ^ " vs. " ^ string_of_int n ^ ")")
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
143  | 
end;  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
144  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
145  | 
fun read_command stream =  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
146  | 
(case TextIO.inputLine stream of  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
147  | 
NONE => raise Runtime.TERMINATE  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
148  | 
| SOME line => map (read_chunk stream) (space_explode "," line));  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
149  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
150  | 
fun run_command name args =  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
151  | 
Runtime.debugging (command name) args  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
152  | 
handle exn =>  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
153  | 
      error ("Isabelle process command failure: " ^ name ^ "\n" ^ ML_Compiler.exn_message exn);
 | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
154  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
155  | 
in  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
156  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
157  | 
fun loop stream =  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
158  | 
let val continue =  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
159  | 
(case read_command stream of  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
160  | 
[] => (Output.error_msg "Isabelle process: no input"; true)  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
161  | 
| name :: args => (run_command name args; true))  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
162  | 
handle Runtime.TERMINATE => false  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
163  | 
| exn => (Output.error_msg (ML_Compiler.exn_message exn) handle crash => recover crash; true);  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
164  | 
in if continue then loop stream else () end;  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
165  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
166  | 
end;  | 
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
167  | 
|
| 
 
71bb3c273dd1
native Isabelle_Process commands, based on efficient byte channel protocol for string lists;
 
wenzelm 
parents: 
38265 
diff
changeset
 | 
168  | 
|
| 25554 | 169  | 
(* init *)  | 
170  | 
||
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
171  | 
fun init in_fifo out_fifo = ignore (Simple_Thread.fork false (fn () =>  | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
172  | 
let  | 
| 40301 | 173  | 
val _ = OS.Process.sleep (seconds 0.5); (*yield to raw ML toplevel*)  | 
| 
44389
 
a3b5fdfb04a3
tuned signature -- contrast physical output primitives versus Output.raw_message;
 
wenzelm 
parents: 
44270 
diff
changeset
 | 
174  | 
val _ = Output.physical_stdout Symbol.STX;  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
175  | 
|
| 
40521
 
8896bd93488e
back to quick_and_dirty, which is still practically important since the scheduler does not jump over subproofs;
 
wenzelm 
parents: 
40518 
diff
changeset
 | 
176  | 
val _ = quick_and_dirty := true;  | 
| 
40518
 
035a27279705
defensive defaults for more robust experience for new users;
 
wenzelm 
parents: 
40301 
diff
changeset
 | 
177  | 
val _ = Goal.parallel_proofs := 0;  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
178  | 
val _ = Context.set_thread_data NONE;  | 
| 
43671
 
a250b092ac66
hard-wired print mode "xsymbols" increases chance that "iff" in HOL will print symbolic arrow;
 
wenzelm 
parents: 
42897 
diff
changeset
 | 
179  | 
val _ =  | 
| 
 
a250b092ac66
hard-wired print mode "xsymbols" increases chance that "iff" in HOL will print symbolic arrow;
 
wenzelm 
parents: 
42897 
diff
changeset
 | 
180  | 
Unsynchronized.change print_mode  | 
| 
 
a250b092ac66
hard-wired print mode "xsymbols" increases chance that "iff" in HOL will print symbolic arrow;
 
wenzelm 
parents: 
42897 
diff
changeset
 | 
181  | 
(fold (update op =)  | 
| 
 
a250b092ac66
hard-wired print mode "xsymbols" increases chance that "iff" in HOL will print symbolic arrow;
 
wenzelm 
parents: 
42897 
diff
changeset
 | 
182  | 
[Symbol.xsymbolsN, isabelle_processN, Keyword.keyword_statusN, Pretty.symbolicN]);  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
183  | 
|
| 
40134
 
8baded087d34
explicit message_output thread, with flushing after timeout, ensure atomic user-operations without the danger of IO Interrupt;
 
wenzelm 
parents: 
40133 
diff
changeset
 | 
184  | 
val in_stream = setup_channels in_fifo out_fifo;  | 
| 
38253
 
3d4e521014f7
Isabelle_Process: separate input fifo for commands (still using the old tty protocol);
 
wenzelm 
parents: 
38236 
diff
changeset
 | 
185  | 
val _ = Keyword.status ();  | 
| 
43673
 
29eb1cd29961
Thy_Info.dependencies: ignore already loaded theories, according to initial prover session status;
 
wenzelm 
parents: 
43671 
diff
changeset
 | 
186  | 
val _ = Thy_Info.status ();  | 
| 
39626
 
a5d0bcfb95a3
manage persistent syslog via Session, not Isabelle_Process;
 
wenzelm 
parents: 
39591 
diff
changeset
 | 
187  | 
val _ = Output.status (Markup.markup Markup.ready "process ready");  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
188  | 
in loop in_stream end));  | 
| 
25528
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
|
| 
 
e67230c2b952
Isabelle process wrapper -- interaction via external program.
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
end;  | 
| 
39528
 
c01d89d18ff0
refined Isabelle_Process startup: emit \002 before rendezvous on fifos, more robust treatment of startup failure with timeout, do not quit() after main loop;
 
wenzelm 
parents: 
39513 
diff
changeset
 | 
191  |