| author | aspinall | 
| Tue, 09 Jan 2007 12:12:00 +0100 | |
| changeset 22042 | 64f8f5433f30 | 
| parent 22028 | c13f6b5bf2b8 | 
| child 22159 | 0cf0d3912239 | 
| permissions | -rw-r--r-- | 
| 21637 | 1  | 
(* Title: Pure/ProofGeneral/proof_general_pgip.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: David Aspinall and Markus Wenzel  | 
|
4  | 
||
5  | 
Isabelle configuration for Proof General using PGIP protocol.  | 
|
| 21940 | 6  | 
See http://proofgeneral.inf.ed.ac.uk/kit  | 
| 21637 | 7  | 
*)  | 
8  | 
||
9  | 
signature PROOF_GENERAL_PGIP =  | 
|
10  | 
sig  | 
|
| 21940 | 11  | 
val init_pgip: bool -> unit (* main PGIP loop with true; fail with false *)  | 
| 21642 | 12  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
13  | 
(* These two are just to support the semi-PGIP Emacs mode *)  | 
| 21940 | 14  | 
val init_pgip_channel: (string -> unit) -> unit  | 
15  | 
val process_pgip: string -> unit  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
16  | 
|
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
17  | 
(* Yet more message functions... *)  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
18  | 
val nonfatal_error : string -> unit (* recoverable error *)  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
19  | 
val log_msg : string -> unit (* for internal log messages *)  | 
| 21637 | 20  | 
end  | 
21  | 
||
22  | 
structure ProofGeneralPgip : PROOF_GENERAL_PGIP =  | 
|
23  | 
struct  | 
|
24  | 
||
25  | 
open Pgip;  | 
|
26  | 
||
| 21949 | 27  | 
|
| 21637 | 28  | 
(* print modes *)  | 
29  | 
||
30  | 
val proof_generalN = "ProofGeneral"; (*token markup (colouring vars, etc.)*)  | 
|
31  | 
val pgmlN = "PGML"; (*XML escapes and PGML symbol elements*)  | 
|
32  | 
val pgmlatomsN = "PGMLatoms"; (*variable kind decorations*)  | 
|
33  | 
val thm_depsN = "thm_deps"; (*meta-information about theorem deps*)  | 
|
34  | 
||
35  | 
||
36  | 
(* text output: print modes for xsymbol and PGML *)  | 
|
37  | 
||
38  | 
local  | 
|
39  | 
||
40  | 
fun xsym_output "\\" = "\\\\"  | 
|
41  | 
| xsym_output s = if Symbol.is_raw s then Symbol.decode_raw s else s;  | 
|
42  | 
||
43  | 
fun xsymbols_output s =  | 
|
| 21940 | 44  | 
if Output.has_mode Symbol.xsymbolsN andalso exists_string (equal "\\") s then  | 
| 21637 | 45  | 
let val syms = Symbol.explode s  | 
46  | 
in (implode (map xsym_output syms), real (Symbol.length syms)) end  | 
|
47  | 
else Symbol.default_output s;  | 
|
48  | 
||
49  | 
(* XML immediately rendered pretty printer. Take care not to double escape *)  | 
|
50  | 
fun pgml_sym s =  | 
|
51  | 
(case Symbol.decode s of  | 
|
52  | 
Symbol.Char s => XML.text s  | 
|
| 21940 | 53  | 
  | Symbol.Sym sn => XML.element "sym" [("name", sn)] [XML.text s]
 | 
| 21637 | 54  | 
  | Symbol.Ctrl sn => XML.element "ctrl" [("name", sn)] [XML.text s] (* FIXME: no such PGML! *)
 | 
55  | 
| Symbol.Raw s => s);  | 
|
56  | 
||
57  | 
fun pgml_output str =  | 
|
58  | 
let val syms = Symbol.explode str  | 
|
59  | 
in (implode (map pgml_sym syms), real (Symbol.length syms)) end;  | 
|
60  | 
||
61  | 
in  | 
|
62  | 
||
63  | 
fun setup_xsymbols_output () =  | 
|
| 21940 | 64  | 
Output.add_mode Symbol.xsymbolsN  | 
| 21637 | 65  | 
(xsymbols_output, K xsymbols_output, Symbol.default_indent, Symbol.encode_raw);  | 
66  | 
||
67  | 
fun setup_pgml_output () =  | 
|
68  | 
Output.add_mode pgmlN  | 
|
69  | 
(pgml_output, K pgml_output, Symbol.default_indent, Symbol.encode_raw);  | 
|
70  | 
||
71  | 
end;  | 
|
72  | 
||
73  | 
||
74  | 
(* token translations *)  | 
|
75  | 
||
76  | 
local  | 
|
77  | 
||
78  | 
val class_tag = "class"  | 
|
79  | 
val tfree_tag = "tfree"  | 
|
80  | 
val tvar_tag = "tvar"  | 
|
81  | 
val free_tag = "free"  | 
|
82  | 
val bound_tag = "bound"  | 
|
83  | 
val var_tag = "var"  | 
|
84  | 
val skolem_tag = "skolem"  | 
|
85  | 
||
86  | 
fun xml_atom kind x = XML.element "atom" [("kind", kind)] [XML.text x];
 | 
|
87  | 
||
| 21940 | 88  | 
fun tagit kind x =  | 
| 21949 | 89  | 
(xml_atom kind x, real (Symbol.length (Symbol.explode x)));  | 
| 21637 | 90  | 
|
91  | 
fun free_or_skolem x =  | 
|
92  | 
(case try Name.dest_skolem x of  | 
|
93  | 
NONE => tagit free_tag x  | 
|
94  | 
| SOME x' => tagit skolem_tag x');  | 
|
95  | 
||
96  | 
fun var_or_skolem s =  | 
|
97  | 
(case Syntax.read_variable s of  | 
|
98  | 
SOME (x, i) =>  | 
|
99  | 
(case try Name.dest_skolem x of  | 
|
100  | 
NONE => tagit var_tag s  | 
|
101  | 
| SOME x' => tagit skolem_tag  | 
|
102  | 
(setmp show_question_marks true Syntax.string_of_vname (x', i)))  | 
|
103  | 
| NONE => tagit var_tag s);  | 
|
104  | 
||
105  | 
val proof_general_trans =  | 
|
106  | 
Syntax.tokentrans_mode proof_generalN  | 
|
107  | 
  [("class", tagit class_tag),
 | 
|
108  | 
   ("tfree", tagit tfree_tag),
 | 
|
109  | 
   ("tvar", tagit tvar_tag),
 | 
|
110  | 
   ("free", free_or_skolem),
 | 
|
111  | 
   ("bound", tagit bound_tag),
 | 
|
112  | 
   ("var", var_or_skolem)];
 | 
|
113  | 
||
114  | 
in  | 
|
115  | 
||
116  | 
val _ = Context.add_setup (Theory.add_tokentrfuns proof_general_trans);  | 
|
117  | 
||
118  | 
end;  | 
|
119  | 
||
120  | 
||
121  | 
(** assembling and issuing PGIP packets **)  | 
|
122  | 
||
123  | 
val pgip_refid = ref NONE: string option ref;  | 
|
124  | 
val pgip_refseq = ref NONE: int option ref;  | 
|
125  | 
||
126  | 
local  | 
|
127  | 
val pgip_class = "pg"  | 
|
128  | 
val pgip_tag = "Isabelle/Isar"  | 
|
129  | 
val pgip_id = ref ""  | 
|
130  | 
val pgip_seq = ref 0  | 
|
131  | 
fun pgip_serial () = inc pgip_seq  | 
|
132  | 
||
133  | 
fun assemble_pgips pgips =  | 
|
| 21940 | 134  | 
    Pgip { tag = SOME pgip_tag,
 | 
135  | 
class = pgip_class,  | 
|
136  | 
seq = pgip_serial(),  | 
|
137  | 
id = !pgip_id,  | 
|
138  | 
destid = !pgip_refid,  | 
|
139  | 
(* destid=refid since Isabelle only communicates back to sender *)  | 
|
140  | 
refid = !pgip_refid,  | 
|
141  | 
refseq = !pgip_refseq,  | 
|
142  | 
content = pgips }  | 
|
| 21637 | 143  | 
in  | 
144  | 
||
145  | 
fun init_pgip_session_id () =  | 
|
146  | 
pgip_id := getenv "HOSTNAME" ^ "/" ^ getenv "USER" ^ "/" ^  | 
|
147  | 
getenv "ISABELLE_PID" ^ "/" ^ Time.toString (Time.now ())  | 
|
148  | 
||
149  | 
fun matching_pgip_id id = (id = !pgip_id)  | 
|
150  | 
||
| 21940 | 151  | 
val output_xml_fn = ref writeln_default  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
152  | 
fun output_xml s = (!output_xml_fn) (XML.string_of_tree s); (* TODO: string buffer *)  | 
| 21637 | 153  | 
|
| 21940 | 154  | 
fun issue_pgip_rawtext str =  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
155  | 
output_xml (PgipOutput.output (assemble_pgips [XML.Rawtext str]))  | 
| 21637 | 156  | 
|
157  | 
fun issue_pgips pgipops =  | 
|
158  | 
output_xml (PgipOutput.output (assemble_pgips (map PgipOutput.output pgipops)));  | 
|
159  | 
||
160  | 
fun issue_pgip pgipop =  | 
|
161  | 
output_xml (PgipOutput.output (assemble_pgips [PgipOutput.output pgipop]));  | 
|
162  | 
||
163  | 
end;  | 
|
164  | 
||
165  | 
||
166  | 
(** messages and notification **)  | 
|
167  | 
||
168  | 
local  | 
|
169  | 
val delay_msgs = ref false (*true: accumulate messages*)  | 
|
170  | 
val delayed_msgs = ref []  | 
|
171  | 
||
172  | 
fun queue_or_issue pgip =  | 
|
| 21940 | 173  | 
if ! delay_msgs then  | 
174  | 
delayed_msgs := pgip :: ! delayed_msgs  | 
|
175  | 
else issue_pgip pgip  | 
|
| 21637 | 176  | 
in  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
177  | 
(* Normal responses are messages for the user *)  | 
| 21940 | 178  | 
fun normalmsg area cat urgent s =  | 
179  | 
let  | 
|
180  | 
            val content = XML.Elem("pgmltext",[],[XML.Rawtext s])
 | 
|
181  | 
            val pgip = Normalresponse {area=area,messagecategory=cat,
 | 
|
182  | 
urgent=urgent,content=[content] }  | 
|
183  | 
in  | 
|
184  | 
queue_or_issue pgip  | 
|
185  | 
end  | 
|
| 21637 | 186  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
187  | 
(* Error responses are error messages for the user, or system-level messages *)  | 
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
188  | 
fun errormsg fatality s =  | 
| 21940 | 189  | 
let  | 
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
190  | 
              val content = XML.Elem("pgmltext",[],[XML.Rawtext s])
 | 
| 21940 | 191  | 
              val pgip = Errorresponse {area=NONE,fatality=fatality,
 | 
192  | 
content=[content],  | 
|
193  | 
(* FIXME: pass in locations *)  | 
|
194  | 
location=NONE}  | 
|
195  | 
in  | 
|
196  | 
queue_or_issue pgip  | 
|
197  | 
end  | 
|
| 21637 | 198  | 
|
199  | 
fun start_delay_msgs () = (set delay_msgs; delayed_msgs := [])  | 
|
200  | 
fun end_delayed_msgs () = (reset delay_msgs; ! delayed_msgs)  | 
|
201  | 
end;  | 
|
202  | 
||
203  | 
(* NB: all of the standard error/message functions now expect already-escaped text.  | 
|
204  | 
FIXME: this may cause us problems now we're generating trees; on the other  | 
|
| 21940 | 205  | 
hand the output functions were tuned some time ago, so it ought to be  | 
| 21637 | 206  | 
enough to use Rawtext always above. *)  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
207  | 
(* NB 2: all of standard functions print strings terminated with new lines, but we don't  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
208  | 
add new lines explicitly in PGIP: they are left implicit. It means that PGIP messages  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
209  | 
can't be written without newlines. *)  | 
| 21637 | 210  | 
|
211  | 
fun setup_messages () =  | 
|
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
212  | 
(writeln_fn := (fn s => normalmsg Message Normal false s);  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
213  | 
priority_fn := (fn s => normalmsg Message Normal true s);  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
214  | 
tracing_fn := (fn s => normalmsg Message Tracing false s);  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
215  | 
info_fn := (fn s => errormsg Info s);  | 
| 22001 | 216  | 
warning_fn := (fn s => errormsg Warning s);  | 
| 
22014
 
4b70cbd96007
removed Toplevel.print_exn hook -- existing error_msg_fn does the job;
 
wenzelm 
parents: 
22001 
diff
changeset
 | 
217  | 
error_fn := (fn s => errormsg Fatal s);  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
218  | 
panic_fn := (fn s => errormsg Panic s);  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
219  | 
debug_fn := (fn s => errormsg Debug s));  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
220  | 
|
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
221  | 
fun nonfatal_error s = errormsg Nonfatal s;  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
222  | 
fun log_msg s = errormsg Log s;  | 
| 
21983
 
9fb029d1189b
Use Isar toplevel print_exn_fn for generating error responses instead of Output.error_msg.
 
aspinall 
parents: 
21972 
diff
changeset
 | 
223  | 
|
| 21637 | 224  | 
|
225  | 
(* immediate messages *)  | 
|
226  | 
||
227  | 
fun tell_clear_goals ()      = issue_pgip (Cleardisplay {area=Display})
 | 
|
228  | 
fun tell_clear_response ()   = issue_pgip (Cleardisplay {area=Message})
 | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
229  | 
|
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
230  | 
fun tell_file_loaded completed path =  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
231  | 
    issue_pgip (Informfileloaded {url=PgipTypes.pgipurl_of_path path,
 | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
232  | 
completed=completed})  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
233  | 
fun tell_file_outdated completed path =  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
234  | 
    issue_pgip (Informfileoutdated {url=PgipTypes.pgipurl_of_path path,
 | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
235  | 
completed=completed})  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
236  | 
fun tell_file_retracted completed path =  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
237  | 
    issue_pgip (Informfileretracted {url=PgipTypes.pgipurl_of_path path,
 | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
238  | 
completed=completed})  | 
| 21637 | 239  | 
|
240  | 
||
241  | 
||
242  | 
(** theory / proof state output **)  | 
|
243  | 
||
244  | 
local  | 
|
245  | 
||
246  | 
fun statedisplay prts =  | 
|
| 22028 | 247  | 
let  | 
| 21940 | 248  | 
val display = Pretty.output (Pretty.chunks prts)  | 
| 22028 | 249  | 
(* TODO: add extra PGML markup for allow proof state navigation *)  | 
| 21940 | 250  | 
        val statedisplay = XML.Elem("statedisplay",[], [XML.Rawtext display])
 | 
| 21637 | 251  | 
in  | 
| 21940 | 252  | 
        issue_pgip (Proofstate {pgml=[XML.Elem("pgml",[],[statedisplay])]})
 | 
| 21637 | 253  | 
end  | 
254  | 
||
255  | 
fun print_current_goals n m st =  | 
|
| 22028 | 256  | 
case (Display.pretty_current_goals n m st) of  | 
257  | 
[] => tell_clear_goals()  | 
|
258  | 
| prts => statedisplay prts  | 
|
| 21637 | 259  | 
|
260  | 
fun print_state b st =  | 
|
| 22028 | 261  | 
case (Toplevel.pretty_state b st) of  | 
262  | 
[] => tell_clear_goals()  | 
|
263  | 
| prts => statedisplay prts  | 
|
| 21637 | 264  | 
|
265  | 
in  | 
|
266  | 
||
267  | 
fun setup_state () =  | 
|
268  | 
(Display.print_current_goals_fn := print_current_goals;  | 
|
269  | 
Toplevel.print_state_fn := print_state);  | 
|
270  | 
||
271  | 
end;  | 
|
272  | 
||
273  | 
||
274  | 
(* theory loader actions *)  | 
|
275  | 
||
276  | 
local  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
277  | 
(* da: TODO: PGIP has a completed flag so the prover can indicate to the  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
278  | 
interface which files are busy performing a particular action.  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
279  | 
To make use of this we need to adjust the hook in thy_info.ML  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
280  | 
(may actually be difficult to tell the interface *which* action is in  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
281  | 
progress, but we could add a generic "Lock" action which uses  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
282  | 
informfileloaded: the broker/UI should not infer too much from incomplete  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
283  | 
operations).  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
284  | 
*)  | 
| 21637 | 285  | 
fun trace_action action name =  | 
286  | 
if action = ThyInfo.Update then  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
287  | 
List.app (tell_file_loaded true) (ThyInfo.loaded_files name)  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
288  | 
else if action = ThyInfo.Outdate then  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
289  | 
List.app (tell_file_outdated true) (ThyInfo.loaded_files name)  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
290  | 
else if action = ThyInfo.Remove then  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
291  | 
List.app (tell_file_retracted true) (ThyInfo.loaded_files name)  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
292  | 
else ()  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
293  | 
|
| 21637 | 294  | 
|
295  | 
in  | 
|
296  | 
fun setup_thy_loader () = ThyInfo.add_hook trace_action;  | 
|
297  | 
fun sync_thy_loader () = List.app (trace_action ThyInfo.Update) (ThyInfo.names ());  | 
|
298  | 
end;  | 
|
299  | 
||
300  | 
||
| 21949 | 301  | 
(* get informed about files *)  | 
| 21637 | 302  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
303  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base;  | 
| 21637 | 304  | 
|
305  | 
val inform_file_retracted = ThyInfo.if_known_thy ThyInfo.remove_thy o thy_name;  | 
|
306  | 
val inform_file_processed = ThyInfo.if_known_thy ThyInfo.touch_child_thys o thy_name;  | 
|
307  | 
||
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
308  | 
fun proper_inform_file_processed path state =  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
309  | 
let val name = thy_name path in  | 
| 21637 | 310  | 
if Toplevel.is_toplevel state andalso ThyInfo.known_thy name then  | 
311  | 
(ThyInfo.touch_child_thys name;  | 
|
312  | 
ThyInfo.pretend_use_thy_only name handle ERROR msg =>  | 
|
313  | 
       (warning msg; warning ("Failed to register theory: " ^ quote name);
 | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
314  | 
tell_file_retracted true (Path.base path)))  | 
| 21637 | 315  | 
else raise Toplevel.UNDEF  | 
316  | 
end;  | 
|
317  | 
||
318  | 
||
319  | 
(* restart top-level loop (keeps most state information) *)  | 
|
320  | 
||
321  | 
val welcome = priority o Session.welcome;  | 
|
322  | 
||
323  | 
fun restart () =  | 
|
324  | 
(sync_thy_loader ();  | 
|
325  | 
tell_clear_goals ();  | 
|
326  | 
tell_clear_response ();  | 
|
327  | 
welcome ();  | 
|
328  | 
raise Toplevel.RESTART)  | 
|
329  | 
||
330  | 
||
331  | 
(* theorem dependency output *)  | 
|
332  | 
local  | 
|
333  | 
||
334  | 
val spaces_quote = space_implode " " o map quote;  | 
|
335  | 
||
336  | 
fun thm_deps_message (thms, deps) =  | 
|
| 21940 | 337  | 
let  | 
338  | 
        val valuethms = XML.Elem("value",[("name", "thms")],[XML.Text thms])
 | 
|
339  | 
        val valuedeps = XML.Elem("value",[("name", "deps")],[XML.Text deps])
 | 
|
| 21637 | 340  | 
in  | 
| 21940 | 341  | 
        issue_pgip (Metainforesponse {attrs=[("infotype", "isabelle_theorem_dependencies")],
 | 
342  | 
content=[valuethms,valuedeps]})  | 
|
| 21637 | 343  | 
end  | 
344  | 
||
345  | 
(* FIXME: check this uses non-transitive closure function here *)  | 
|
| 21969 | 346  | 
fun tell_thm_deps ths =  | 
347  | 
if Output.has_mode thm_depsN then  | 
|
348  | 
let  | 
|
349  | 
val names = filter_out (equal "") (map PureThy.get_name_hint ths);  | 
|
350  | 
val deps = filter_out (equal "")  | 
|
351  | 
(Symtab.keys (fold Proofterm.thms_of_proof  | 
|
352  | 
(map Thm.proof_of ths) Symtab.empty));  | 
|
353  | 
in  | 
|
354  | 
if null names orelse null deps then ()  | 
|
355  | 
else thm_deps_message (spaces_quote names, spaces_quote deps)  | 
|
356  | 
end  | 
|
357  | 
else ();  | 
|
| 21637 | 358  | 
|
359  | 
in  | 
|
360  | 
||
361  | 
fun setup_present_hook () =  | 
|
362  | 
Present.add_hook (fn _ => fn res => tell_thm_deps (maps #2 res));  | 
|
363  | 
||
364  | 
end;  | 
|
365  | 
||
366  | 
(** lexicalstructure element with keywords (PGIP version of elisp keywords file) **)  | 
|
367  | 
||
| 21940 | 368  | 
fun lexicalstructure_keywords () =  | 
| 21637 | 369  | 
let val commands = OuterSyntax.dest_keywords ()  | 
| 21940 | 370  | 
fun category_of k = if k mem commands then "major" else "minor"  | 
| 21637 | 371  | 
(* NB: we might filter to only include words like elisp case (OuterSyntax.is_keyword). *)  | 
| 21940 | 372  | 
fun keyword_elt (keyword,help,kind,_) =  | 
373  | 
            XML.Elem("keyword", [("word", keyword), ("category", category_of kind)],
 | 
|
374  | 
                     [XML.Elem("shorthelp", [], [XML.Text help])])
 | 
|
375  | 
in  | 
|
| 21637 | 376  | 
(* Also, note we don't call init_outer_syntax here to add interface commands,  | 
377  | 
but they should never appear in scripts anyway so it shouldn't matter *)  | 
|
378  | 
            Lexicalstructure {content = map keyword_elt (OuterSyntax.dest_parsers()) }
 | 
|
379  | 
end  | 
|
380  | 
||
381  | 
(* TODO: we can issue a lexicalstructure/keyword when the syntax gets extended dynamically;  | 
|
382  | 
hooks needed in outer_syntax.ML to do that. *)  | 
|
383  | 
||
384  | 
||
385  | 
(* Configuration: GUI config, proverinfo messages *)  | 
|
386  | 
||
387  | 
local  | 
|
388  | 
val isabellewww = "http://isabelle.in.tum.de/"  | 
|
389  | 
val staticconfig = "~~/lib/ProofGeneral/pgip_isar.xml"  | 
|
| 21940 | 390  | 
fun orenv v d = case getenv v of "" => d | s => s  | 
| 21637 | 391  | 
fun config_file() = orenv "ISABELLE_PGIPCONFIG" staticconfig  | 
392  | 
fun isabelle_www() = orenv "ISABELLE_HOMEPAGE" isabellewww  | 
|
393  | 
in  | 
|
394  | 
fun send_pgip_config () =  | 
|
395  | 
let  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21655 
diff
changeset
 | 
396  | 
val path = Path.explode (config_file())  | 
| 21940 | 397  | 
val ex = File.exists path  | 
| 21637 | 398  | 
|
| 21940 | 399  | 
val wwwpage =  | 
400  | 
(Url.explode (isabelle_www()))  | 
|
| 21969 | 401  | 
handle ERROR _ =>  | 
| 21940 | 402  | 
(Output.panic  | 
403  | 
                        ("Error in URL in environment variable ISABELLE_HOMEPAGE.");
 | 
|
404  | 
Url.explode isabellewww)  | 
|
405  | 
||
406  | 
val proverinfo =  | 
|
| 21637 | 407  | 
            Proverinfo { name = "Isabelle",
 | 
| 21940 | 408  | 
version = version,  | 
409  | 
instance = Session.name(),  | 
|
410  | 
descr = "The Isabelle/Isar theorem prover",  | 
|
411  | 
url = wwwpage,  | 
|
412  | 
filenameextns = ".thy;" }  | 
|
| 21637 | 413  | 
in  | 
| 21940 | 414  | 
if ex then  | 
415  | 
(issue_pgip proverinfo;  | 
|
416  | 
issue_pgip_rawtext (File.read path);  | 
|
417  | 
issue_pgip (lexicalstructure_keywords()))  | 
|
| 21637 | 418  | 
        else Output.panic ("PGIP configuration file \"" ^ config_file() ^ "\" not found")
 | 
419  | 
end;  | 
|
420  | 
end  | 
|
421  | 
||
422  | 
||
423  | 
||
424  | 
||
425  | 
(* Sending commands to Isar *)  | 
|
426  | 
||
427  | 
fun isarcmd s =  | 
|
428  | 
s |> OuterSyntax.scan |> OuterSyntax.read  | 
|
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
429  | 
(*|> map (Toplevel.position (Position.name "PGIP message") o #3)*)  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
430  | 
|> map #3  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
431  | 
|> Toplevel.>>>;  | 
| 21637 | 432  | 
|
| 21940 | 433  | 
(* TODO:  | 
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
434  | 
- apply a command given a transition function, e.g. IsarCmd.undo.  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
435  | 
- fix position from path of currently open file [line numbers risk garbling though].  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
436  | 
*)  | 
| 21637 | 437  | 
|
438  | 
(* load an arbitrary file (must be .thy or .ML) *)  | 
|
439  | 
||
440  | 
fun use_thy_or_ml_file file =  | 
|
441  | 
let  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21655 
diff
changeset
 | 
442  | 
val (path,extn) = Path.split_ext (Path.explode file)  | 
| 21637 | 443  | 
in  | 
444  | 
case extn of  | 
|
| 21940 | 445  | 
            "" => isarcmd ("use_thy " ^ quote (Path.implode path))
 | 
446  | 
          | "thy" => isarcmd ("use_thy " ^ quote (Path.implode path))
 | 
|
| 21637 | 447  | 
          | "ML" => isarcmd ("use " ^ quote file)
 | 
| 22028 | 448  | 
          | other => error ("Don't know how to read a file with extension " ^ quote other)
 | 
| 21637 | 449  | 
end  | 
450  | 
||
451  | 
||
| 21867 | 452  | 
(******* PGIP actions *******)  | 
| 21637 | 453  | 
|
454  | 
||
| 21940 | 455  | 
(* Responses to each of the PGIP input commands.  | 
| 21637 | 456  | 
These are programmed uniformly for extensibility. *)  | 
457  | 
||
| 21940 | 458  | 
fun askpgip (Askpgip vs) =  | 
| 21637 | 459  | 
issue_pgip  | 
460  | 
        (Usespgip { version = PgipIsabelle.isabelle_pgip_version_supported,
 | 
|
| 21940 | 461  | 
pgipelems = PgipIsabelle.accepted_inputs })  | 
| 21637 | 462  | 
|
| 21940 | 463  | 
fun askpgml (Askpgml vs) =  | 
| 21637 | 464  | 
issue_pgip  | 
| 21940 | 465  | 
        (Usespgml { version = PgipIsabelle.isabelle_pgml_version_supported })
 | 
| 21637 | 466  | 
|
| 21902 | 467  | 
fun askprefs (Askprefs vs) =  | 
| 21940 | 468  | 
let  | 
469  | 
        fun preference_of {name, descr, default, pgiptype, get, set } =
 | 
|
470  | 
            { name = name, descr = SOME descr, default = SOME default,
 | 
|
471  | 
pgiptype = pgiptype }  | 
|
| 21637 | 472  | 
in  | 
| 21940 | 473  | 
List.app (fn (prefcat, prefs) =>  | 
474  | 
                     issue_pgip (Hasprefs {prefcategory=SOME prefcat,
 | 
|
475  | 
prefs=map preference_of prefs}))  | 
|
| 21637 | 476  | 
Preferences.preferences  | 
| 21940 | 477  | 
end  | 
| 21637 | 478  | 
|
| 21902 | 479  | 
fun askconfig (Askconfig vs) = () (* TODO: add config response *)  | 
| 21637 | 480  | 
|
481  | 
local  | 
|
| 21940 | 482  | 
fun lookuppref pref =  | 
483  | 
case AList.lookup (op =)  | 
|
484  | 
(map (fn p => (#name p,p))  | 
|
485  | 
(maps snd Preferences.preferences)) pref of  | 
|
486  | 
            NONE => error ("Unknown prover preference: " ^ quote pref)
 | 
|
487  | 
| SOME p => p  | 
|
| 21637 | 488  | 
in  | 
| 21940 | 489  | 
fun setpref (Setpref vs) =  | 
490  | 
let  | 
|
491  | 
val name = #name vs  | 
|
492  | 
val value = #value vs  | 
|
493  | 
val set = #set (lookuppref name)  | 
|
| 21637 | 494  | 
in  | 
| 21940 | 495  | 
set value  | 
| 21637 | 496  | 
end  | 
497  | 
||
| 21902 | 498  | 
fun getpref (Getpref vs) =  | 
| 21940 | 499  | 
let  | 
500  | 
val name = #name vs  | 
|
501  | 
val get = #get (lookuppref name)  | 
|
502  | 
in  | 
|
| 21637 | 503  | 
        issue_pgip (Prefval {name=name, value=get ()})
 | 
504  | 
end  | 
|
505  | 
end  | 
|
506  | 
||
507  | 
fun proverinit vs = restart ()  | 
|
508  | 
||
509  | 
fun proverexit vs = isarcmd "quit"  | 
|
510  | 
||
511  | 
fun startquiet vs = isarcmd "disable_pr"  | 
|
512  | 
||
513  | 
fun stopquiet vs = isarcmd "enable_pr"  | 
|
514  | 
||
| 21940 | 515  | 
fun pgmlsymbolson vs =  | 
| 21637 | 516  | 
change print_mode (fn mode =>  | 
| 21940 | 517  | 
remove (op =) Symbol.xsymbolsN mode @ [Symbol.xsymbolsN])  | 
| 21637 | 518  | 
|
519  | 
fun pgmlsymbolsoff vs =  | 
|
520  | 
change print_mode (remove (op =) Symbol.xsymbolsN)  | 
|
521  | 
||
| 21940 | 522  | 
fun dostep (Dostep vs) =  | 
523  | 
let  | 
|
524  | 
val text = #text vs  | 
|
525  | 
in  | 
|
526  | 
isarcmd text  | 
|
| 21637 | 527  | 
end  | 
528  | 
||
| 21902 | 529  | 
fun undostep (Undostep vs) =  | 
| 21940 | 530  | 
let  | 
531  | 
val times = #times vs  | 
|
532  | 
in  | 
|
533  | 
        isarcmd ("undos_proof " ^ Int.toString times)
 | 
|
| 21637 | 534  | 
end  | 
535  | 
||
536  | 
fun redostep vs = isarcmd "redo"  | 
|
| 21940 | 537  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
538  | 
fun abortgoal vs = isarcmd "kill" (* was: ProofGeneral.kill_proof *)  | 
| 21637 | 539  | 
|
540  | 
||
| 21867 | 541  | 
(*** PGIP identifier tables ***)  | 
542  | 
||
543  | 
fun setids t = issue_pgip (Setids {idtables = [t]})
 | 
|
544  | 
fun addids t = issue_pgip (Addids {idtables = [t]})
 | 
|
545  | 
fun delids t = issue_pgip (Delids {idtables = [t]})
 | 
|
546  | 
||
547  | 
(*  | 
|
| 21940 | 548  | 
fun delallids ty =  | 
549  | 
     issue_pgip (Setids {idtables =
 | 
|
550  | 
                        [{context=NONE,objtype=ty,ids=[]}]}) *)
 | 
|
| 21867 | 551  | 
|
| 21940 | 552  | 
fun askids (Askids vs) =  | 
| 21637 | 553  | 
let  | 
| 21940 | 554  | 
val url = #url vs (* ask for identifiers within a file *)  | 
555  | 
val thyname = #thyname vs (* ask for identifiers within a theory *)  | 
|
556  | 
val objtype = #objtype vs (* ask for identifiers of a particular type *)  | 
|
| 21867 | 557  | 
|
| 21940 | 558  | 
        fun idtable ty ctx ids = {objtype=ty,context=ctx,ids=ids}
 | 
| 21867 | 559  | 
|
| 21940 | 560  | 
val thms_of_thy = map fst o PureThy.thms_of o ThyInfo.get_theory  | 
561  | 
in  | 
|
562  | 
(* case (url_attr,thyname,objtype) of  | 
|
563  | 
(NONE,NONE,NONE) =>  | 
|
564  | 
*) (* top-level: return *)  | 
|
| 21867 | 565  | 
|
| 21940 | 566  | 
(* TODO: add askids for "file" here, which returns single theory with same name *)  | 
| 21867 | 567  | 
(* FIXME: objtypes on both sides *)  | 
| 21940 | 568  | 
case (thyname,objtype) of  | 
| 21867 | 569  | 
(* only files known in top context *)  | 
| 21940 | 570  | 
(NONE, NONE) => setids (idtable ObjFile NONE (ThyInfo.names())) (*FIXME: uris?*)  | 
571  | 
| (NONE, SOME ObjFile) => setids (idtable ObjFile NONE (ThyInfo.names())) (* ditto *)  | 
|
572  | 
| (SOME fi, SOME ObjFile) => setids (idtable ObjTheory (SOME fi) [fi]) (* FIXME: lookup file *)  | 
|
573  | 
| (NONE, SOME ObjTheory) => setids (idtable ObjTheory NONE (ThyInfo.names()))  | 
|
574  | 
| (SOME thy, SOME ObjTheory) => setids (idtable ObjTheory (SOME thy) (ThyInfo.get_preds thy))  | 
|
575  | 
| (SOME thy, SOME ObjTheorem) => setids (idtable ObjTheorem (SOME thy) (thms_of_thy thy))  | 
|
576  | 
(* next case is both of above. FIXME: cleanup this *)  | 
|
577  | 
| (SOME thy, NONE) => (setids (idtable ObjTheory (SOME thy) (ThyInfo.get_preds thy));  | 
|
578  | 
setids (idtable ObjTheorem (SOME thy) (thms_of_thy thy)))  | 
|
579  | 
         | (_, SOME ot) => error ("No objects of type "^(PgipTypes.name_of_objtype ot)^" are available here.")
 | 
|
| 21637 | 580  | 
end  | 
581  | 
||
582  | 
local  | 
|
583  | 
(* accumulate printed output in a single PGIP response (inside <pgmltext>) *)  | 
|
584  | 
fun with_displaywrap pgipfn dispfn =  | 
|
585  | 
let  | 
|
| 21940 | 586  | 
val lines = ref ([]: string list);  | 
587  | 
fun wlgrablines s = lines := s :: ! lines;  | 
|
| 21637 | 588  | 
in  | 
| 21940 | 589  | 
setmp writeln_fn wlgrablines dispfn ();  | 
590  | 
issue_pgip (pgipfn (!lines))  | 
|
| 21637 | 591  | 
end;  | 
592  | 
in  | 
|
| 21940 | 593  | 
fun showid (Showid vs) =  | 
| 21637 | 594  | 
let  | 
| 21940 | 595  | 
val thyname = #thyname vs  | 
596  | 
val objtype = #objtype vs  | 
|
597  | 
val name = #name vs  | 
|
598  | 
val topthy = Toplevel.theory_of o Toplevel.get_state  | 
|
| 21637 | 599  | 
|
| 21940 | 600  | 
fun idvalue objtype name strings =  | 
601  | 
            Idvalue { name=name, objtype=objtype,
 | 
|
602  | 
                      text=[XML.Elem("pgmltext",[],map XML.Text strings)] }
 | 
|
| 21637 | 603  | 
|
| 21940 | 604  | 
fun pthm thy name = print_thm (get_thm thy (Name name))  | 
605  | 
in  | 
|
606  | 
case (thyname, objtype) of  | 
|
607  | 
(_,ObjTheory) =>  | 
|
608  | 
with_displaywrap (idvalue ObjTheory name)  | 
|
609  | 
(fn ()=>(print_theory (ThyInfo.get_theory name)))  | 
|
610  | 
| (SOME thy, ObjTheorem) =>  | 
|
611  | 
with_displaywrap (idvalue ObjTheorem name)  | 
|
612  | 
(fn ()=>(pthm (ThyInfo.get_theory thy) name))  | 
|
613  | 
| (NONE, ObjTheorem) =>  | 
|
614  | 
with_displaywrap (idvalue ObjTheorem name)  | 
|
615  | 
(fn ()=>pthm (topthy()) name)  | 
|
616  | 
          | (_, ot) => error ("Cannot show objects of type "^(PgipTypes.name_of_objtype ot))
 | 
|
| 21637 | 617  | 
end  | 
618  | 
||
| 21867 | 619  | 
(*** Inspecting state ***)  | 
620  | 
||
| 21940 | 621  | 
(* The file which is currently being processed interactively.  | 
| 21637 | 622  | 
In the pre-PGIP code, this was informed to Isabelle and the theory loader  | 
623  | 
on completion, but that allows for circularity in case we read  | 
|
624  | 
ourselves. So PGIP opens the filename at the start of a script.  | 
|
625  | 
We ought to prevent problems by modifying the theory loader to know  | 
|
| 21940 | 626  | 
about this special status, but for now we just keep a local reference.  | 
627  | 
*)  | 
|
| 21637 | 628  | 
|
629  | 
val currently_open_file = ref (NONE : pgipurl option)  | 
|
630  | 
||
| 21940 | 631  | 
fun askguise vs =  | 
| 21637 | 632  | 
(* The "guise" is the PGIP abstraction of the prover's state.  | 
633  | 
The <informguise> message is merely used for consistency checking. *)  | 
|
| 21940 | 634  | 
let  | 
635  | 
val openfile = !currently_open_file  | 
|
| 21637 | 636  | 
|
| 21940 | 637  | 
val topthy = Toplevel.theory_of o Toplevel.get_state  | 
638  | 
val topthy_name = Context.theory_name o topthy  | 
|
| 21637 | 639  | 
|
| 21940 | 640  | 
val opentheory = SOME (topthy_name()) handle Toplevel.UNDEF => NONE  | 
| 21637 | 641  | 
|
| 21940 | 642  | 
fun topproofpos () = try Toplevel.proof_position_of (Isar.state ());  | 
643  | 
val openproofpos = topproofpos()  | 
|
| 21637 | 644  | 
in  | 
645  | 
        issue_pgip (Informguise { file = openfile,
 | 
|
| 21940 | 646  | 
theory = opentheory,  | 
647  | 
(* would be nice to get thm name... *)  | 
|
648  | 
theorem = NONE,  | 
|
649  | 
proofpos = openproofpos })  | 
|
| 21637 | 650  | 
end  | 
651  | 
||
| 21902 | 652  | 
fun parsescript (Parsescript vs) =  | 
| 21637 | 653  | 
let  | 
| 21940 | 654  | 
val text = #text vs  | 
655  | 
val systemdata = #systemdata vs  | 
|
656  | 
val location = #location vs (* TODO: extract position *)  | 
|
| 21637 | 657  | 
|
| 21867 | 658  | 
val _ = start_delay_msgs () (* gather parsing errs/warns *)  | 
659  | 
val doc = PgipParser.pgip_parser text  | 
|
| 21637 | 660  | 
val errs = end_delayed_msgs ()  | 
661  | 
||
| 21940 | 662  | 
val sysattrs = PgipTypes.opt_attr "systemdata" systemdata  | 
663  | 
val locattrs = PgipTypes.attrs_of_location location  | 
|
| 21637 | 664  | 
in  | 
665  | 
        issue_pgip (Parseresult { attrs= sysattrs@locattrs,
 | 
|
| 21940 | 666  | 
doc = doc,  | 
667  | 
errs = map PgipOutput.output errs })  | 
|
| 21637 | 668  | 
end  | 
669  | 
||
670  | 
fun showproofstate vs = isarcmd "pr"  | 
|
671  | 
||
| 21969 | 672  | 
fun showctxt vs = isarcmd "print_context"  | 
| 21637 | 673  | 
|
| 21902 | 674  | 
fun searchtheorems (Searchtheorems vs) =  | 
| 21940 | 675  | 
let  | 
676  | 
val arg = #arg vs  | 
|
| 21637 | 677  | 
in  | 
| 21969 | 678  | 
        isarcmd ("find_theorems " ^ arg)
 | 
| 21637 | 679  | 
end  | 
680  | 
||
| 21940 | 681  | 
fun setlinewidth (Setlinewidth vs) =  | 
682  | 
let  | 
|
683  | 
val width = #width vs  | 
|
| 21637 | 684  | 
in  | 
| 21940 | 685  | 
        isarcmd ("pretty_setmargin " ^ Int.toString width) (* FIXME: conversion back/forth! *)
 | 
| 21637 | 686  | 
end  | 
687  | 
||
| 21902 | 688  | 
fun viewdoc (Viewdoc vs) =  | 
| 21940 | 689  | 
let  | 
690  | 
val arg = #arg vs  | 
|
691  | 
in  | 
|
692  | 
        isarcmd ("print_" ^ arg)   (* FIXME: isatool doc?.  Return URLs, maybe? *)
 | 
|
| 21637 | 693  | 
end  | 
694  | 
||
| 21867 | 695  | 
(*** Theory ***)  | 
696  | 
||
| 21902 | 697  | 
fun doitem (Doitem vs) =  | 
| 21637 | 698  | 
let  | 
| 21940 | 699  | 
val text = #text vs  | 
| 21637 | 700  | 
in  | 
| 21940 | 701  | 
isarcmd text  | 
| 21637 | 702  | 
end  | 
703  | 
||
704  | 
fun undoitem vs =  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
705  | 
isarcmd "undo"  | 
| 21637 | 706  | 
|
707  | 
fun redoitem vs =  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
708  | 
isarcmd "redo"  | 
| 21637 | 709  | 
|
| 21940 | 710  | 
fun aborttheory vs =  | 
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
711  | 
isarcmd "kill" (* was: "init_toplevel" *)  | 
| 21637 | 712  | 
|
| 21902 | 713  | 
fun retracttheory (Retracttheory vs) =  | 
| 21940 | 714  | 
let  | 
715  | 
val thyname = #thyname vs  | 
|
| 21637 | 716  | 
in  | 
| 21940 | 717  | 
        isarcmd ("kill_thy " ^ quote thyname)
 | 
| 21637 | 718  | 
end  | 
719  | 
||
| 21867 | 720  | 
|
721  | 
(*** Files ***)  | 
|
722  | 
||
723  | 
(* Path management: we allow theory files to have dependencies in  | 
|
724  | 
their own directory, but when we change directory for a new file we  | 
|
725  | 
remove the path. Leaving it there can cause confusion with  | 
|
726  | 
difference in batch mode.  | 
|
| 21940 | 727  | 
NB: PGIP does not assume that the prover has a load path.  | 
| 21867 | 728  | 
*)  | 
729  | 
||
730  | 
local  | 
|
731  | 
val current_working_dir = ref (NONE : string option)  | 
|
732  | 
in  | 
|
| 21940 | 733  | 
fun changecwd_dir newdirpath =  | 
734  | 
let  | 
|
| 21867 | 735  | 
val newdir = File.platform_path newdirpath  | 
| 21940 | 736  | 
in  | 
| 21867 | 737  | 
(case (!current_working_dir) of  | 
738  | 
NONE => ()  | 
|
739  | 
| SOME dir => ThyLoad.del_path dir;  | 
|
740  | 
ThyLoad.add_path newdir;  | 
|
741  | 
current_working_dir := SOME newdir)  | 
|
742  | 
end  | 
|
743  | 
end  | 
|
744  | 
||
| 21940 | 745  | 
fun changecwd (Changecwd vs) =  | 
746  | 
let  | 
|
747  | 
val url = #url vs  | 
|
748  | 
val newdir = PgipTypes.path_of_pgipurl url  | 
|
| 21867 | 749  | 
in  | 
| 21940 | 750  | 
changecwd_dir url  | 
| 21867 | 751  | 
end  | 
752  | 
||
| 21902 | 753  | 
fun openfile (Openfile vs) =  | 
| 21940 | 754  | 
let  | 
| 21867 | 755  | 
val url = #url vs  | 
756  | 
val filepath = PgipTypes.path_of_pgipurl url  | 
|
757  | 
val filedir = Path.dir filepath  | 
|
758  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base  | 
|
759  | 
val openfile_retract = Output.no_warnings (ThyInfo.if_known_thy ThyInfo.remove_thy) o thy_name;  | 
|
760  | 
in  | 
|
761  | 
case !currently_open_file of  | 
|
| 22028 | 762  | 
          SOME f => raise PGIP ("<openfile> when a file is already open!\nCurrently open file: " ^
 | 
763  | 
PgipTypes.string_of_pgipurl url)  | 
|
| 21867 | 764  | 
| NONE => (openfile_retract filepath;  | 
| 21940 | 765  | 
changecwd_dir filedir;  | 
| 22028 | 766  | 
		   priority ("Working in file: " ^ PgipTypes.string_of_pgipurl url);
 | 
| 21940 | 767  | 
currently_open_file := SOME url)  | 
| 21867 | 768  | 
end  | 
769  | 
||
770  | 
fun closefile vs =  | 
|
771  | 
case !currently_open_file of  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
772  | 
SOME f => (proper_inform_file_processed f (Isar.state());  | 
| 22028 | 773  | 
		   priority ("Finished working in file: " ^ PgipTypes.string_of_pgipurl f);
 | 
| 21867 | 774  | 
currently_open_file := NONE)  | 
775  | 
      | NONE => raise PGIP ("<closefile> when no file is open!")
 | 
|
776  | 
||
| 21940 | 777  | 
fun loadfile (Loadfile vs) =  | 
778  | 
let  | 
|
779  | 
val url = #url vs  | 
|
780  | 
in  | 
|
781  | 
case !currently_open_file of  | 
|
| 22028 | 782  | 
            SOME f => raise PGIP ("<loadfile> when a file is open!\nCurrently open file: " ^
 | 
783  | 
PgipTypes.string_of_pgipurl url)  | 
|
| 21637 | 784  | 
| NONE => use_thy_or_ml_file (File.platform_path url)  | 
785  | 
end  | 
|
786  | 
||
787  | 
fun abortfile vs =  | 
|
788  | 
case !currently_open_file of  | 
|
789  | 
SOME f => (isarcmd "init_toplevel";  | 
|
| 22028 | 790  | 
		   priority ("Aborted working in file: " ^ 
 | 
791  | 
PgipTypes.string_of_pgipurl f);  | 
|
| 21940 | 792  | 
currently_open_file := NONE)  | 
| 21637 | 793  | 
      | NONE => raise PGIP ("<abortfile> when no file is open!")
 | 
794  | 
||
| 21940 | 795  | 
fun retractfile (Retractfile vs) =  | 
796  | 
let  | 
|
797  | 
val url = #url vs  | 
|
| 21637 | 798  | 
in  | 
| 21940 | 799  | 
case !currently_open_file of  | 
| 21637 | 800  | 
            SOME f => raise PGIP ("<retractfile> when a file is open!")
 | 
| 22028 | 801  | 
          | NONE => (priority ("Retracting file: " ^ PgipTypes.string_of_pgipurl url);
 | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
802  | 
(* TODO: next should be in thy loader, here just for testing *)  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
803  | 
let  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
804  | 
val name = thy_name url  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
805  | 
in List.app (tell_file_retracted false) (ThyInfo.loaded_files name) end;  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
806  | 
inform_file_retracted url)  | 
| 21637 | 807  | 
end  | 
808  | 
||
809  | 
||
| 21867 | 810  | 
(*** System ***)  | 
| 21637 | 811  | 
|
| 21902 | 812  | 
fun systemcmd (Systemcmd vs) =  | 
| 21940 | 813  | 
let  | 
| 21637 | 814  | 
val arg = #arg vs  | 
815  | 
in  | 
|
816  | 
isarcmd arg  | 
|
817  | 
end  | 
|
818  | 
||
819  | 
exception PGIP_QUIT;  | 
|
820  | 
fun quitpgip vs = raise PGIP_QUIT  | 
|
821  | 
||
| 21902 | 822  | 
fun process_input inp = case inp  | 
823  | 
of Pgip.Askpgip _ => askpgip inp  | 
|
824  | 
| Pgip.Askpgml _ => askpgml inp  | 
|
| 21940 | 825  | 
| Pgip.Askprefs _ => askprefs inp  | 
| 21902 | 826  | 
| Pgip.Askconfig _ => askconfig inp  | 
827  | 
| Pgip.Getpref _ => getpref inp  | 
|
828  | 
| Pgip.Setpref _ => setpref inp  | 
|
829  | 
| Pgip.Proverinit _ => proverinit inp  | 
|
830  | 
| Pgip.Proverexit _ => proverexit inp  | 
|
831  | 
| Pgip.Startquiet _ => startquiet inp  | 
|
832  | 
| Pgip.Stopquiet _ => stopquiet inp  | 
|
833  | 
| Pgip.Pgmlsymbolson _ => pgmlsymbolson inp  | 
|
834  | 
| Pgip.Pgmlsymbolsoff _ => pgmlsymbolsoff inp  | 
|
835  | 
| Pgip.Dostep _ => dostep inp  | 
|
836  | 
| Pgip.Undostep _ => undostep inp  | 
|
837  | 
| Pgip.Redostep _ => redostep inp  | 
|
838  | 
| Pgip.Forget _ => error "<forget> not implemented"  | 
|
839  | 
| Pgip.Restoregoal _ => error "<restoregoal> not implemented"  | 
|
840  | 
| Pgip.Abortgoal _ => abortgoal inp  | 
|
841  | 
| Pgip.Askids _ => askids inp  | 
|
842  | 
| Pgip.Showid _ => showid inp  | 
|
843  | 
| Pgip.Askguise _ => askguise inp  | 
|
844  | 
| Pgip.Parsescript _ => parsescript inp  | 
|
845  | 
| Pgip.Showproofstate _ => showproofstate inp  | 
|
846  | 
| Pgip.Showctxt _ => showctxt inp  | 
|
847  | 
| Pgip.Searchtheorems _ => searchtheorems inp  | 
|
848  | 
| Pgip.Setlinewidth _ => setlinewidth inp  | 
|
849  | 
| Pgip.Viewdoc _ => viewdoc inp  | 
|
850  | 
| Pgip.Doitem _ => doitem inp  | 
|
851  | 
| Pgip.Undoitem _ => undoitem inp  | 
|
852  | 
| Pgip.Redoitem _ => redoitem inp  | 
|
853  | 
| Pgip.Aborttheory _ => aborttheory inp  | 
|
854  | 
| Pgip.Retracttheory _ => retracttheory inp  | 
|
855  | 
| Pgip.Loadfile _ => loadfile inp  | 
|
856  | 
| Pgip.Openfile _ => openfile inp  | 
|
857  | 
| Pgip.Closefile _ => closefile inp  | 
|
858  | 
| Pgip.Abortfile _ => abortfile inp  | 
|
859  | 
| Pgip.Retractfile _ => retractfile inp  | 
|
860  | 
| Pgip.Changecwd _ => changecwd inp  | 
|
861  | 
| Pgip.Systemcmd _ => systemcmd inp  | 
|
862  | 
| Pgip.Quitpgip _ => quitpgip inp  | 
|
| 21637 | 863  | 
|
864  | 
||
| 21940 | 865  | 
fun process_pgip_element pgipxml =  | 
| 21637 | 866  | 
case pgipxml of  | 
| 21969 | 867  | 
xml as (XML.Elem elem) =>  | 
| 21940 | 868  | 
(case Pgip.input elem of  | 
869  | 
             NONE => warning ("Unrecognized PGIP command, ignored: \n" ^
 | 
|
870  | 
(XML.string_of_tree xml))  | 
|
871  | 
| SOME inp => (process_input inp)) (* errors later; packet discarded *)  | 
|
| 21969 | 872  | 
| XML.Text t => ignored_text_warning t  | 
873  | 
| XML.Rawtext t => ignored_text_warning t  | 
|
| 21637 | 874  | 
and ignored_text_warning t =  | 
| 21940 | 875  | 
if size (Symbol.strip_blanks t) > 0 then  | 
876  | 
           warning ("Ignored text in PGIP packet: \n" ^ t)
 | 
|
| 21637 | 877  | 
else ()  | 
878  | 
||
879  | 
fun process_pgip_tree xml =  | 
|
880  | 
(pgip_refid := NONE;  | 
|
881  | 
pgip_refseq := NONE;  | 
|
882  | 
(case xml of  | 
|
883  | 
          XML.Elem ("pgip", attrs, pgips) =>
 | 
|
884  | 
(let  | 
|
885  | 
val class = PgipTypes.get_attr "class" attrs  | 
|
886  | 
val dest = PgipTypes.get_attr_opt "destid" attrs  | 
|
| 21940 | 887  | 
val seq = PgipTypes.read_pgipnat (PgipTypes.get_attr "seq" attrs)  | 
| 21637 | 888  | 
(* Respond to prover broadcasts, or messages for us. Ignore rest *)  | 
| 21940 | 889  | 
val processit =  | 
890  | 
case dest of  | 
|
| 21637 | 891  | 
NONE => class = "pa"  | 
| 21940 | 892  | 
| SOME id => matching_pgip_id id  | 
893  | 
in if processit then  | 
|
894  | 
(pgip_refid := PgipTypes.get_attr_opt "id" attrs;  | 
|
895  | 
pgip_refseq := SOME seq;  | 
|
896  | 
List.app process_pgip_element pgips;  | 
|
897  | 
(* return true to indicate <ready/> *)  | 
|
898  | 
true)  | 
|
899  | 
else  | 
|
900  | 
(* no response to ignored messages. *)  | 
|
901  | 
false  | 
|
| 21637 | 902  | 
end)  | 
903  | 
| _ => raise PGIP "Invalid PGIP packet received")  | 
|
904  | 
handle PGIP msg =>  | 
|
905  | 
(Output.error_msg ((msg ^ "\nPGIP error occured in XML text below:\n") ^  | 
|
| 21940 | 906  | 
(XML.string_of_tree xml));  | 
907  | 
true))  | 
|
| 21637 | 908  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
909  | 
(* External input *)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
910  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
911  | 
val process_pgip_plain = K () o process_pgip_tree o XML.tree_of_string  | 
| 21637 | 912  | 
|
913  | 
end  | 
|
914  | 
||
915  | 
||
916  | 
(* PGIP loop: process PGIP input only *)  | 
|
917  | 
||
918  | 
local  | 
|
919  | 
||
920  | 
exception XML_PARSE  | 
|
921  | 
||
922  | 
fun loop ready src =  | 
|
923  | 
let  | 
|
924  | 
val _ = if ready then issue_pgip (Ready ()) else ()  | 
|
| 21969 | 925  | 
val pgipo =  | 
926  | 
(case try Source.get_single src of  | 
|
927  | 
SOME pgipo => pgipo  | 
|
928  | 
| NONE => raise XML_PARSE)  | 
|
| 21637 | 929  | 
in  | 
930  | 
case pgipo of  | 
|
931  | 
NONE => ()  | 
|
932  | 
| SOME (pgip,src') =>  | 
|
933  | 
let  | 
|
| 21940 | 934  | 
val ready' = (process_pgip_tree pgip)  | 
935  | 
handle e => (handler (e,SOME src'); true)  | 
|
| 21637 | 936  | 
in  | 
937  | 
loop ready' src'  | 
|
938  | 
end  | 
|
939  | 
end handle e => handler (e,SOME src) (* error in XML parse or Ready issue *)  | 
|
940  | 
||
941  | 
and handler (e,srco) =  | 
|
942  | 
case (e,srco) of  | 
|
943  | 
(XML_PARSE,SOME src) =>  | 
|
944  | 
Output.panic "Invalid XML input, aborting"  | 
|
945  | 
| (PGIP_QUIT,_) => ()  | 
|
| 21940 | 946  | 
| (Interrupt,SOME src) =>  | 
947  | 
(Output.error_msg "Interrupt during PGIP processing"; loop true src)  | 
|
948  | 
| (Toplevel.UNDEF,SOME src) =>  | 
|
949  | 
(Output.error_msg "No working context defined"; loop true src)  | 
|
950  | 
| (e,SOME src) =>  | 
|
951  | 
(Output.error_msg (Toplevel.exn_message e); loop true src)  | 
|
| 21637 | 952  | 
| (_,NONE) => ()  | 
953  | 
in  | 
|
954  | 
(* TODO: add socket interface *)  | 
|
955  | 
||
956  | 
val xmlP = XML.scan_comment_whspc |-- XML.parse_elem >> single  | 
|
957  | 
||
958  | 
val tty_src = Source.set_prompt "" (Source.source Symbol.stopper xmlP NONE Source.tty)  | 
|
959  | 
||
960  | 
fun pgip_toplevel x = loop true x  | 
|
961  | 
end  | 
|
962  | 
||
963  | 
||
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
964  | 
local  | 
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
965  | 
(* Extra command for embedding prover-control inside document (obscure/debug usage). *)  | 
| 21637 | 966  | 
|
| 
22014
 
4b70cbd96007
removed Toplevel.print_exn hook -- existing error_msg_fn does the job;
 
wenzelm 
parents: 
22001 
diff
changeset
 | 
967  | 
val process_pgipP =  | 
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
968  | 
OuterSyntax.improper_command "ProofGeneral.process_pgip" "(internal)" OuterKeyword.control  | 
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
969  | 
(OuterParse.text >> (Toplevel.no_timing oo  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
970  | 
(fn txt => Toplevel.imperative (fn () => process_pgip_plain txt))));  | 
| 21637 | 971  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
972  | 
in  | 
| 21637 | 973  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
974  | 
fun init_outer_syntax () = OuterSyntax.add_parsers [process_pgipP];  | 
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
975  | 
|
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
976  | 
end  | 
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
977  | 
|
| 21637 | 978  | 
|
979  | 
||
980  | 
(* init *)  | 
|
981  | 
||
982  | 
val initialized = ref false;  | 
|
983  | 
||
| 21969 | 984  | 
fun init_pgip false =  | 
985  | 
Output.panic "No Proof General interface support for Isabelle/classic mode."  | 
|
986  | 
| init_pgip true =  | 
|
987  | 
(! initialized orelse  | 
|
988  | 
(setmp warning_fn (K ()) init_outer_syntax ();  | 
|
| 
22014
 
4b70cbd96007
removed Toplevel.print_exn hook -- existing error_msg_fn does the job;
 
wenzelm 
parents: 
22001 
diff
changeset
 | 
989  | 
PgipParser.init ();  | 
| 21969 | 990  | 
setup_xsymbols_output ();  | 
991  | 
setup_pgml_output ();  | 
|
992  | 
setup_messages ();  | 
|
993  | 
setup_state ();  | 
|
994  | 
setup_thy_loader ();  | 
|
995  | 
setup_present_hook ();  | 
|
996  | 
init_pgip_session_id ();  | 
|
997  | 
welcome ();  | 
|
998  | 
set initialized);  | 
|
999  | 
sync_thy_loader ();  | 
|
1000  | 
change print_mode (cons proof_generalN o remove (op =) proof_generalN);  | 
|
1001  | 
change print_mode (append [pgmlN, pgmlatomsN] o subtract (op =) [pgmlN, pgmlatomsN]);  | 
|
1002  | 
pgip_toplevel tty_src);  | 
|
| 21637 | 1003  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1004  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1005  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1006  | 
(** Out-of-loop PGIP commands (for Emacs hybrid mode) **)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1007  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1008  | 
local  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1009  | 
val pgip_output_channel = ref writeln_default  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1010  | 
in  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1011  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1012  | 
(* Set recipient for PGIP results *)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1013  | 
fun init_pgip_channel writefn =  | 
| 21940 | 1014  | 
(init_pgip_session_id();  | 
1015  | 
pgip_output_channel := writefn)  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1016  | 
|
| 21940 | 1017  | 
(* Process a PGIP command.  | 
1018  | 
This works for preferences but not generally guaranteed  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1019  | 
because we haven't done full setup here (e.g., no pgml mode) *)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1020  | 
fun process_pgip str =  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1021  | 
setmp output_xml_fn (!pgip_output_channel) process_pgip_plain str  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1022  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1023  | 
end  | 
| 21637 | 1024  | 
|
1025  | 
end;  | 
|
1026  |