| author | aspinall | 
| Fri, 06 Jul 2007 17:52:52 +0200 | |
| changeset 23610 | 5ade06703b07 | 
| parent 23435 | 061f28854017 | 
| child 23621 | e070a6ab1891 | 
| 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  | 
|
| 
23435
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
17  | 
(* More message functions... *)  | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
18  | 
val nonfatal_error : string -> unit (* recoverable (batch) error: carry on scripting *)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
19  | 
val log_msg : string -> unit (* for internal log messages *)  | 
| 22163 | 20  | 
val error_with_pos : PgipTypes.fatality -> Position.T -> string -> unit  | 
| 22159 | 21  | 
|
| 22163 | 22  | 
val get_currently_open_file : unit -> Path.T option (* interface focus *)  | 
| 21637 | 23  | 
end  | 
24  | 
||
25  | 
structure ProofGeneralPgip : PROOF_GENERAL_PGIP =  | 
|
26  | 
struct  | 
|
27  | 
||
28  | 
open Pgip;  | 
|
29  | 
||
| 21949 | 30  | 
|
| 22408 | 31  | 
(* print mode *)  | 
| 21637 | 32  | 
|
| 22408 | 33  | 
val proof_generalN = "ProofGeneral";  | 
34  | 
val pgmlsymbols_flag = ref true;  | 
|
| 21637 | 35  | 
|
36  | 
local  | 
|
37  | 
||
38  | 
fun xsym_output "\\" = "\\\\"  | 
|
39  | 
| xsym_output s = if Symbol.is_raw s then Symbol.decode_raw s else s;  | 
|
40  | 
||
41  | 
fun pgml_sym s =  | 
|
42  | 
(case Symbol.decode s of  | 
|
43  | 
Symbol.Char s => XML.text s  | 
|
| 22408 | 44  | 
| Symbol.Sym sn =>  | 
45  | 
let val ascii = implode (map xsym_output (Symbol.explode s))  | 
|
46  | 
    in if !pgmlsymbols_flag then XML.element "sym" [("name", sn)] [XML.text ascii]
 | 
|
47  | 
else ascii end  | 
|
48  | 
  | Symbol.Ctrl sn => XML.element "ctrl" [("name", sn)] [XML.text sn] (* FIXME: no such PGML *)
 | 
|
49  | 
| Symbol.Raw raw => raw);  | 
|
| 21637 | 50  | 
|
51  | 
fun pgml_output str =  | 
|
52  | 
let val syms = Symbol.explode str  | 
|
53  | 
in (implode (map pgml_sym syms), real (Symbol.length syms)) end;  | 
|
54  | 
||
55  | 
in  | 
|
56  | 
||
| 22408 | 57  | 
fun setup_proofgeneral_output () =  | 
58  | 
Output.add_mode proof_generalN  | 
|
| 
22397
 
ec7ecf706955
Fix idvalue output and PGML print mode raw encode/decode.
 
aspinall 
parents: 
22337 
diff
changeset
 | 
59  | 
(pgml_output, K pgml_output, Symbol.default_indent, Symbol.encode_raw);  | 
| 21637 | 60  | 
|
61  | 
end;  | 
|
62  | 
||
63  | 
||
64  | 
(* token translations *)  | 
|
65  | 
||
66  | 
local  | 
|
67  | 
||
68  | 
val class_tag = "class"  | 
|
69  | 
val tfree_tag = "tfree"  | 
|
70  | 
val tvar_tag = "tvar"  | 
|
71  | 
val free_tag = "free"  | 
|
72  | 
val bound_tag = "bound"  | 
|
73  | 
val var_tag = "var"  | 
|
74  | 
val skolem_tag = "skolem"  | 
|
75  | 
||
76  | 
fun xml_atom kind x = XML.element "atom" [("kind", kind)] [XML.text x];
 | 
|
77  | 
||
| 21940 | 78  | 
fun tagit kind x =  | 
| 21949 | 79  | 
(xml_atom kind x, real (Symbol.length (Symbol.explode x)));  | 
| 21637 | 80  | 
|
81  | 
fun free_or_skolem x =  | 
|
82  | 
(case try Name.dest_skolem x of  | 
|
83  | 
NONE => tagit free_tag x  | 
|
84  | 
| SOME x' => tagit skolem_tag x');  | 
|
85  | 
||
86  | 
fun var_or_skolem s =  | 
|
87  | 
(case Syntax.read_variable s of  | 
|
88  | 
SOME (x, i) =>  | 
|
89  | 
(case try Name.dest_skolem x of  | 
|
90  | 
NONE => tagit var_tag s  | 
|
91  | 
| SOME x' => tagit skolem_tag  | 
|
| 22678 | 92  | 
(setmp show_question_marks true Term.string_of_vname (x', i)))  | 
| 21637 | 93  | 
| NONE => tagit var_tag s);  | 
94  | 
||
95  | 
val proof_general_trans =  | 
|
96  | 
Syntax.tokentrans_mode proof_generalN  | 
|
97  | 
  [("class", tagit class_tag),
 | 
|
98  | 
   ("tfree", tagit tfree_tag),
 | 
|
99  | 
   ("tvar", tagit tvar_tag),
 | 
|
100  | 
   ("free", free_or_skolem),
 | 
|
101  | 
   ("bound", tagit bound_tag),
 | 
|
102  | 
   ("var", var_or_skolem)];
 | 
|
103  | 
||
104  | 
in  | 
|
105  | 
||
106  | 
val _ = Context.add_setup (Theory.add_tokentrfuns proof_general_trans);  | 
|
107  | 
||
108  | 
end;  | 
|
109  | 
||
110  | 
||
111  | 
(** assembling and issuing PGIP packets **)  | 
|
112  | 
||
113  | 
val pgip_refid = ref NONE: string option ref;  | 
|
114  | 
val pgip_refseq = ref NONE: int option ref;  | 
|
115  | 
||
116  | 
local  | 
|
117  | 
val pgip_class = "pg"  | 
|
118  | 
val pgip_tag = "Isabelle/Isar"  | 
|
119  | 
val pgip_id = ref ""  | 
|
120  | 
val pgip_seq = ref 0  | 
|
121  | 
fun pgip_serial () = inc pgip_seq  | 
|
122  | 
||
123  | 
fun assemble_pgips pgips =  | 
|
| 21940 | 124  | 
    Pgip { tag = SOME pgip_tag,
 | 
125  | 
class = pgip_class,  | 
|
126  | 
seq = pgip_serial(),  | 
|
127  | 
id = !pgip_id,  | 
|
128  | 
destid = !pgip_refid,  | 
|
129  | 
(* destid=refid since Isabelle only communicates back to sender *)  | 
|
130  | 
refid = !pgip_refid,  | 
|
131  | 
refseq = !pgip_refseq,  | 
|
132  | 
content = pgips }  | 
|
| 21637 | 133  | 
in  | 
134  | 
||
135  | 
fun init_pgip_session_id () =  | 
|
136  | 
pgip_id := getenv "HOSTNAME" ^ "/" ^ getenv "USER" ^ "/" ^  | 
|
137  | 
getenv "ISABELLE_PID" ^ "/" ^ Time.toString (Time.now ())  | 
|
138  | 
||
139  | 
fun matching_pgip_id id = (id = !pgip_id)  | 
|
140  | 
||
| 22590 | 141  | 
val output_xml_fn = ref Output.writeln_default  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
142  | 
fun output_xml s = (!output_xml_fn) (XML.string_of_tree s); (* TODO: string buffer *)  | 
| 21637 | 143  | 
|
| 21940 | 144  | 
fun issue_pgip_rawtext str =  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
145  | 
output_xml (PgipOutput.output (assemble_pgips [XML.Rawtext str]))  | 
| 21637 | 146  | 
|
147  | 
fun issue_pgips pgipops =  | 
|
148  | 
output_xml (PgipOutput.output (assemble_pgips (map PgipOutput.output pgipops)));  | 
|
149  | 
||
150  | 
fun issue_pgip pgipop =  | 
|
151  | 
output_xml (PgipOutput.output (assemble_pgips [PgipOutput.output pgipop]));  | 
|
152  | 
||
153  | 
end;  | 
|
154  | 
||
155  | 
||
156  | 
(** messages and notification **)  | 
|
157  | 
||
158  | 
local  | 
|
159  | 
val delay_msgs = ref false (*true: accumulate messages*)  | 
|
160  | 
val delayed_msgs = ref []  | 
|
161  | 
||
162  | 
fun queue_or_issue pgip =  | 
|
| 21940 | 163  | 
if ! delay_msgs then  | 
164  | 
delayed_msgs := pgip :: ! delayed_msgs  | 
|
165  | 
else issue_pgip pgip  | 
|
| 21637 | 166  | 
in  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
167  | 
(* Normal responses are messages for the user *)  | 
| 21940 | 168  | 
fun normalmsg area cat urgent s =  | 
169  | 
let  | 
|
170  | 
            val content = XML.Elem("pgmltext",[],[XML.Rawtext s])
 | 
|
171  | 
            val pgip = Normalresponse {area=area,messagecategory=cat,
 | 
|
172  | 
urgent=urgent,content=[content] }  | 
|
173  | 
in  | 
|
174  | 
queue_or_issue pgip  | 
|
175  | 
end  | 
|
| 21637 | 176  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
177  | 
(* 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
 | 
178  | 
fun errormsg fatality s =  | 
| 21940 | 179  | 
let  | 
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
180  | 
              val content = XML.Elem("pgmltext",[],[XML.Rawtext s])
 | 
| 21940 | 181  | 
              val pgip = Errorresponse {area=NONE,fatality=fatality,
 | 
182  | 
content=[content],  | 
|
183  | 
location=NONE}  | 
|
184  | 
in  | 
|
185  | 
queue_or_issue pgip  | 
|
186  | 
end  | 
|
| 21637 | 187  | 
|
| 22159 | 188  | 
(* Error responses with useful locations *)  | 
189  | 
fun error_with_pos fatality pos s =  | 
|
190  | 
let  | 
|
191  | 
              val content = XML.Elem("pgmltext",[],[XML.Rawtext s])
 | 
|
192  | 
              val pgip = Errorresponse {area=NONE,fatality=fatality,
 | 
|
193  | 
content=[content],  | 
|
194  | 
location=SOME (PgipIsabelle.location_of_position pos)}  | 
|
195  | 
in  | 
|
196  | 
queue_or_issue pgip  | 
|
197  | 
end  | 
|
198  | 
||
| 21637 | 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  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
208  | 
add new lines explicitly in PGIP: they are left implicit. It means that PGIP messages  | 
| 
22042
 
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 () =  | 
|
| 22590 | 212  | 
(Output.writeln_fn := (fn s => normalmsg Message Normal false s);  | 
213  | 
Output.priority_fn := (fn s => normalmsg Message Normal true s);  | 
|
214  | 
Output.tracing_fn := (fn s => normalmsg Message Tracing false s);  | 
|
215  | 
Output.warning_fn := (fn s => errormsg Warning s);  | 
|
216  | 
Output.error_fn := (fn s => errormsg Fatal s);  | 
|
217  | 
Output.debug_fn := (fn s => errormsg Debug s));  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
218  | 
|
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
219  | 
fun panic s = (errormsg Panic ("## SYSTEM EXIT ##\n" ^ s); exit 1);
 | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
220  | 
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
 | 
221  | 
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
 | 
222  | 
|
| 21637 | 223  | 
|
224  | 
(* immediate messages *)  | 
|
225  | 
||
226  | 
fun tell_clear_goals ()      = issue_pgip (Cleardisplay {area=Display})
 | 
|
227  | 
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
 | 
228  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
229  | 
fun tell_file_loaded completed path =  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
230  | 
    issue_pgip (Informfileloaded {url=PgipTypes.pgipurl_of_path path,
 | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
231  | 
completed=completed})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
232  | 
fun tell_file_outdated completed path =  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
233  | 
    issue_pgip (Informfileoutdated {url=PgipTypes.pgipurl_of_path path,
 | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
234  | 
completed=completed})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
235  | 
fun tell_file_retracted completed path =  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
236  | 
    issue_pgip (Informfileretracted {url=PgipTypes.pgipurl_of_path path,
 | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
237  | 
completed=completed})  | 
| 21637 | 238  | 
|
239  | 
||
240  | 
||
241  | 
(** theory / proof state output **)  | 
|
242  | 
||
243  | 
local  | 
|
244  | 
||
245  | 
fun statedisplay prts =  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
246  | 
let  | 
| 23610 | 247  | 
val display = Pretty.chunks prts  | 
248  | 
val pgml = Pgml.Pgml  | 
|
249  | 
   { version=SOME PgipIsabelle.isabelle_pgml_version_supported, 
 | 
|
250  | 
systemid=SOME PgipIsabelle.systemid,  | 
|
251  | 
content = Pgml.Subterm  | 
|
252  | 
		   { kind=SOME "statedisplay",
 | 
|
253  | 
param=NONE,place=NONE,name=NONE,decoration=NONE,  | 
|
254  | 
action=NONE,pos=NONE,xref=NONE,  | 
|
255  | 
		     content=[Pgml.Atoms { kind = NONE,
 | 
|
256  | 
content = [Pgml.Str (Pretty.output display)] }] }  | 
|
257  | 
(* [PgmlIsabelle.pgml_of_prettyT display] *) }  | 
|
258  | 
(* TODO: PGML markup for proof state navigation *)  | 
|
| 21637 | 259  | 
in  | 
| 23610 | 260  | 
        issue_pgip (Proofstate {pgml=pgml})
 | 
| 21637 | 261  | 
end  | 
262  | 
||
263  | 
fun print_current_goals n m st =  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
264  | 
case (Display.pretty_current_goals n m st) of  | 
| 22028 | 265  | 
[] => tell_clear_goals()  | 
266  | 
| prts => statedisplay prts  | 
|
| 21637 | 267  | 
|
268  | 
fun print_state b st =  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
269  | 
case (Toplevel.pretty_state b st) of  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
270  | 
[] => tell_clear_goals()  | 
| 22028 | 271  | 
| prts => statedisplay prts  | 
| 21637 | 272  | 
|
273  | 
in  | 
|
274  | 
||
275  | 
fun setup_state () =  | 
|
276  | 
(Display.print_current_goals_fn := print_current_goals;  | 
|
277  | 
Toplevel.print_state_fn := print_state);  | 
|
278  | 
||
279  | 
end;  | 
|
280  | 
||
281  | 
||
282  | 
(* theory loader actions *)  | 
|
283  | 
||
284  | 
local  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
285  | 
(* da: TODO: PGIP has a completed flag so the prover can indicate to the  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
286  | 
interface which files are busy performing a particular action.  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
287  | 
To make use of this we need to adjust the hook in thy_info.ML  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
288  | 
(may actually be difficult to tell the interface *which* action is in  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
289  | 
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
 | 
290  | 
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
 | 
291  | 
operations).  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
292  | 
*)  | 
| 21637 | 293  | 
fun trace_action action name =  | 
294  | 
if action = ThyInfo.Update then  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
295  | 
List.app (tell_file_loaded true) (ThyInfo.loaded_files name)  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
296  | 
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
 | 
297  | 
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
 | 
298  | 
else if action = ThyInfo.Remove then  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
299  | 
List.app (tell_file_retracted true) (ThyInfo.loaded_files name)  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
300  | 
else ()  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
301  | 
|
| 21637 | 302  | 
|
303  | 
in  | 
|
304  | 
fun setup_thy_loader () = ThyInfo.add_hook trace_action;  | 
|
305  | 
fun sync_thy_loader () = List.app (trace_action ThyInfo.Update) (ThyInfo.names ());  | 
|
306  | 
end;  | 
|
307  | 
||
308  | 
||
| 21949 | 309  | 
(* get informed about files *)  | 
| 21637 | 310  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
311  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base;  | 
| 21637 | 312  | 
|
313  | 
val inform_file_retracted = ThyInfo.if_known_thy ThyInfo.remove_thy o thy_name;  | 
|
314  | 
val inform_file_processed = ThyInfo.if_known_thy ThyInfo.touch_child_thys o thy_name;  | 
|
315  | 
||
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
316  | 
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
 | 
317  | 
let val name = thy_name path in  | 
| 21637 | 318  | 
if Toplevel.is_toplevel state andalso ThyInfo.known_thy name then  | 
319  | 
(ThyInfo.touch_child_thys name;  | 
|
320  | 
ThyInfo.pretend_use_thy_only name handle ERROR msg =>  | 
|
321  | 
       (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
 | 
322  | 
tell_file_retracted true (Path.base path)))  | 
| 21637 | 323  | 
else raise Toplevel.UNDEF  | 
324  | 
end;  | 
|
325  | 
||
326  | 
||
327  | 
(* restart top-level loop (keeps most state information) *)  | 
|
328  | 
||
329  | 
val welcome = priority o Session.welcome;  | 
|
330  | 
||
331  | 
fun restart () =  | 
|
332  | 
(sync_thy_loader ();  | 
|
333  | 
tell_clear_goals ();  | 
|
334  | 
tell_clear_response ();  | 
|
335  | 
welcome ();  | 
|
336  | 
raise Toplevel.RESTART)  | 
|
337  | 
||
338  | 
||
339  | 
(* theorem dependency output *)  | 
|
| 22408 | 340  | 
|
341  | 
val show_theorem_dependencies = ref false;  | 
|
342  | 
||
| 21637 | 343  | 
local  | 
344  | 
||
345  | 
val spaces_quote = space_implode " " o map quote;  | 
|
346  | 
||
347  | 
fun thm_deps_message (thms, deps) =  | 
|
| 21940 | 348  | 
let  | 
349  | 
        val valuethms = XML.Elem("value",[("name", "thms")],[XML.Text thms])
 | 
|
350  | 
        val valuedeps = XML.Elem("value",[("name", "deps")],[XML.Text deps])
 | 
|
| 21637 | 351  | 
in  | 
| 21940 | 352  | 
        issue_pgip (Metainforesponse {attrs=[("infotype", "isabelle_theorem_dependencies")],
 | 
353  | 
content=[valuethms,valuedeps]})  | 
|
| 21637 | 354  | 
end  | 
355  | 
||
| 21969 | 356  | 
fun tell_thm_deps ths =  | 
| 22408 | 357  | 
if !show_theorem_dependencies then  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
358  | 
let  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
359  | 
val names = map PureThy.get_name_hint (filter PureThy.has_name_hint ths);  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
360  | 
val deps = (Symtab.keys (fold Proofterm.thms_of_proof'  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
361  | 
(map Thm.proof_of ths) Symtab.empty))  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
362  | 
in  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
363  | 
if null names orelse null deps then ()  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
364  | 
else thm_deps_message (spaces_quote names, spaces_quote deps)  | 
| 22225 | 365  | 
end  | 
366  | 
else ()  | 
|
| 21637 | 367  | 
|
368  | 
in  | 
|
369  | 
||
370  | 
fun setup_present_hook () =  | 
|
371  | 
Present.add_hook (fn _ => fn res => tell_thm_deps (maps #2 res));  | 
|
372  | 
||
373  | 
end;  | 
|
374  | 
||
375  | 
(** lexicalstructure element with keywords (PGIP version of elisp keywords file) **)  | 
|
376  | 
||
| 21940 | 377  | 
fun lexicalstructure_keywords () =  | 
| 21637 | 378  | 
let val commands = OuterSyntax.dest_keywords ()  | 
| 21940 | 379  | 
fun category_of k = if k mem commands then "major" else "minor"  | 
| 21637 | 380  | 
(* NB: we might filter to only include words like elisp case (OuterSyntax.is_keyword). *)  | 
| 21940 | 381  | 
fun keyword_elt (keyword,help,kind,_) =  | 
382  | 
            XML.Elem("keyword", [("word", keyword), ("category", category_of kind)],
 | 
|
383  | 
                     [XML.Elem("shorthelp", [], [XML.Text help])])
 | 
|
384  | 
in  | 
|
| 21637 | 385  | 
(* Also, note we don't call init_outer_syntax here to add interface commands,  | 
386  | 
but they should never appear in scripts anyway so it shouldn't matter *)  | 
|
387  | 
            Lexicalstructure {content = map keyword_elt (OuterSyntax.dest_parsers()) }
 | 
|
388  | 
end  | 
|
389  | 
||
390  | 
(* TODO: we can issue a lexicalstructure/keyword when the syntax gets extended dynamically;  | 
|
391  | 
hooks needed in outer_syntax.ML to do that. *)  | 
|
392  | 
||
393  | 
||
394  | 
(* Configuration: GUI config, proverinfo messages *)  | 
|
395  | 
||
396  | 
local  | 
|
397  | 
val isabellewww = "http://isabelle.in.tum.de/"  | 
|
398  | 
val staticconfig = "~~/lib/ProofGeneral/pgip_isar.xml"  | 
|
| 21940 | 399  | 
fun orenv v d = case getenv v of "" => d | s => s  | 
| 21637 | 400  | 
fun config_file() = orenv "ISABELLE_PGIPCONFIG" staticconfig  | 
401  | 
fun isabelle_www() = orenv "ISABELLE_HOMEPAGE" isabellewww  | 
|
402  | 
in  | 
|
403  | 
fun send_pgip_config () =  | 
|
404  | 
let  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21655 
diff
changeset
 | 
405  | 
val path = Path.explode (config_file())  | 
| 21940 | 406  | 
val ex = File.exists path  | 
| 21637 | 407  | 
|
| 21940 | 408  | 
val wwwpage =  | 
409  | 
(Url.explode (isabelle_www()))  | 
|
| 21969 | 410  | 
handle ERROR _ =>  | 
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
411  | 
                   (panic ("Error in URL in environment variable ISABELLE_HOMEPAGE.");
 | 
| 21940 | 412  | 
Url.explode isabellewww)  | 
413  | 
||
414  | 
val proverinfo =  | 
|
| 21637 | 415  | 
            Proverinfo { name = "Isabelle",
 | 
| 21940 | 416  | 
version = version,  | 
417  | 
instance = Session.name(),  | 
|
418  | 
descr = "The Isabelle/Isar theorem prover",  | 
|
419  | 
url = wwwpage,  | 
|
420  | 
filenameextns = ".thy;" }  | 
|
| 21637 | 421  | 
in  | 
| 21940 | 422  | 
if ex then  | 
423  | 
(issue_pgip proverinfo;  | 
|
424  | 
issue_pgip_rawtext (File.read path);  | 
|
425  | 
issue_pgip (lexicalstructure_keywords()))  | 
|
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
426  | 
        else panic ("PGIP configuration file \"" ^ config_file() ^ "\" not found")
 | 
| 21637 | 427  | 
end;  | 
428  | 
end  | 
|
429  | 
||
430  | 
||
| 
22216
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
431  | 
(* Preferences: tweak for PGIP interfaces *)  | 
| 
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
432  | 
|
| 
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
433  | 
val preferences = ref Preferences.preferences;  | 
| 
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
434  | 
|
| 
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
435  | 
fun setup_preferences_tweak() =  | 
| 
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
436  | 
preferences :=  | 
| 
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
437  | 
     (!preferences |> Preferences.set_default ("show-question-marks","false")
 | 
| 22408 | 438  | 
|> Preferences.remove "show-question-marks" (* we use markup, not ?s *)  | 
| 22590 | 439  | 
|> Preferences.remove "theorem-dependencies" (* set internally *)  | 
440  | 
|> Preferences.remove "full-proofs") (* set internally *)  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
441  | 
|
| 21637 | 442  | 
|
443  | 
||
444  | 
(* Sending commands to Isar *)  | 
|
445  | 
||
446  | 
fun isarcmd s =  | 
|
447  | 
s |> OuterSyntax.scan |> OuterSyntax.read  | 
|
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
448  | 
(*|> map (Toplevel.position (Position.name "PGIP message") o #3)*)  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
449  | 
|> map #3  | 
| 
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
450  | 
|> Toplevel.>>>;  | 
| 21637 | 451  | 
|
| 21940 | 452  | 
(* TODO:  | 
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
453  | 
- 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
 | 
454  | 
- 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
 | 
455  | 
*)  | 
| 21637 | 456  | 
|
457  | 
(* load an arbitrary file (must be .thy or .ML) *)  | 
|
458  | 
||
459  | 
fun use_thy_or_ml_file file =  | 
|
460  | 
let  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21655 
diff
changeset
 | 
461  | 
val (path,extn) = Path.split_ext (Path.explode file)  | 
| 21637 | 462  | 
in  | 
463  | 
case extn of  | 
|
| 21940 | 464  | 
            "" => isarcmd ("use_thy " ^ quote (Path.implode path))
 | 
465  | 
          | "thy" => isarcmd ("use_thy " ^ quote (Path.implode path))
 | 
|
| 21637 | 466  | 
          | "ML" => isarcmd ("use " ^ quote file)
 | 
| 22028 | 467  | 
          | other => error ("Don't know how to read a file with extension " ^ quote other)
 | 
| 21637 | 468  | 
end  | 
469  | 
||
470  | 
||
| 21867 | 471  | 
(******* PGIP actions *******)  | 
| 21637 | 472  | 
|
473  | 
||
| 21940 | 474  | 
(* Responses to each of the PGIP input commands.  | 
| 21637 | 475  | 
These are programmed uniformly for extensibility. *)  | 
476  | 
||
| 21940 | 477  | 
fun askpgip (Askpgip vs) =  | 
| 
23435
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
478  | 
(issue_pgip  | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
479  | 
         (Usespgip { version = PgipIsabelle.isabelle_pgip_version_supported,
 | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
480  | 
pgipelems = PgipIsabelle.accepted_inputs });  | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
481  | 
send_pgip_config())  | 
| 21637 | 482  | 
|
| 21940 | 483  | 
fun askpgml (Askpgml vs) =  | 
| 21637 | 484  | 
issue_pgip  | 
| 21940 | 485  | 
        (Usespgml { version = PgipIsabelle.isabelle_pgml_version_supported })
 | 
| 21637 | 486  | 
|
| 21902 | 487  | 
fun askprefs (Askprefs vs) =  | 
| 21940 | 488  | 
let  | 
489  | 
        fun preference_of {name, descr, default, pgiptype, get, set } =
 | 
|
490  | 
            { name = name, descr = SOME descr, default = SOME default,
 | 
|
491  | 
pgiptype = pgiptype }  | 
|
| 21637 | 492  | 
in  | 
| 21940 | 493  | 
List.app (fn (prefcat, prefs) =>  | 
494  | 
                     issue_pgip (Hasprefs {prefcategory=SOME prefcat,
 | 
|
495  | 
prefs=map preference_of prefs}))  | 
|
| 
22216
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
496  | 
(!preferences)  | 
| 21940 | 497  | 
end  | 
| 21637 | 498  | 
|
| 21902 | 499  | 
fun askconfig (Askconfig vs) = () (* TODO: add config response *)  | 
| 21637 | 500  | 
|
501  | 
local  | 
|
| 21940 | 502  | 
fun lookuppref pref =  | 
503  | 
case AList.lookup (op =)  | 
|
504  | 
(map (fn p => (#name p,p))  | 
|
| 
22216
 
c3f5aa951a68
Tweak preferences for PGIP interfaces. Make <askids> return theory successors instead of parents (ideally should be content elements).
 
aspinall 
parents: 
22171 
diff
changeset
 | 
505  | 
(maps snd (!preferences))) pref of  | 
| 21940 | 506  | 
            NONE => error ("Unknown prover preference: " ^ quote pref)
 | 
507  | 
| SOME p => p  | 
|
| 21637 | 508  | 
in  | 
| 21940 | 509  | 
fun setpref (Setpref vs) =  | 
510  | 
let  | 
|
511  | 
val name = #name vs  | 
|
512  | 
val value = #value vs  | 
|
513  | 
val set = #set (lookuppref name)  | 
|
| 21637 | 514  | 
in  | 
| 21940 | 515  | 
set value  | 
| 21637 | 516  | 
end  | 
517  | 
||
| 21902 | 518  | 
fun getpref (Getpref vs) =  | 
| 21940 | 519  | 
let  | 
520  | 
val name = #name vs  | 
|
521  | 
val get = #get (lookuppref name)  | 
|
522  | 
in  | 
|
| 21637 | 523  | 
        issue_pgip (Prefval {name=name, value=get ()})
 | 
524  | 
end  | 
|
525  | 
end  | 
|
526  | 
||
527  | 
fun proverinit vs = restart ()  | 
|
528  | 
||
529  | 
fun proverexit vs = isarcmd "quit"  | 
|
530  | 
||
| 22408 | 531  | 
fun set_proverflag_quiet b =  | 
532  | 
isarcmd (if b then "disable_pr" else "enable_pr")  | 
|
| 21637 | 533  | 
|
| 22408 | 534  | 
fun set_proverflag_pgmlsymbols b =  | 
535  | 
(pgmlsymbols_flag := b;  | 
|
536  | 
change print_mode  | 
|
| 22590 | 537  | 
(fn mode =>  | 
| 22408 | 538  | 
remove (op =) Symbol.xsymbolsN mode @ (if b then [Symbol.xsymbolsN] else [])))  | 
539  | 
||
540  | 
fun set_proverflag_thmdeps b =  | 
|
541  | 
(show_theorem_dependencies := b;  | 
|
542  | 
proofs := (if b then 1 else 2))  | 
|
| 21637 | 543  | 
|
| 22408 | 544  | 
fun setproverflag (Setproverflag vs) =  | 
545  | 
let  | 
|
| 22590 | 546  | 
val flagname = #flagname vs  | 
547  | 
val value = #value vs  | 
|
| 22408 | 548  | 
in  | 
| 22590 | 549  | 
(case flagname of  | 
550  | 
"quiet" => set_proverflag_quiet value  | 
|
551  | 
| "pgmlsymbols" => set_proverflag_pgmlsymbols value  | 
|
552  | 
| "metainfo:thmdeps" => set_proverflag_thmdeps value  | 
|
| 
23435
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
553  | 
           | _ => log_msg ("Unrecognised prover control flag: " ^ 
 | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
554  | 
(quote flagname) ^ " ignored."))  | 
| 22408 | 555  | 
end  | 
556  | 
||
| 21637 | 557  | 
|
| 21940 | 558  | 
fun dostep (Dostep vs) =  | 
559  | 
let  | 
|
560  | 
val text = #text vs  | 
|
561  | 
in  | 
|
562  | 
isarcmd text  | 
|
| 21637 | 563  | 
end  | 
564  | 
||
| 21902 | 565  | 
fun undostep (Undostep vs) =  | 
| 21940 | 566  | 
let  | 
567  | 
val times = #times vs  | 
|
568  | 
in  | 
|
569  | 
        isarcmd ("undos_proof " ^ Int.toString times)
 | 
|
| 21637 | 570  | 
end  | 
571  | 
||
572  | 
fun redostep vs = isarcmd "redo"  | 
|
| 21940 | 573  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
574  | 
fun abortgoal vs = isarcmd "kill" (* was: ProofGeneral.kill_proof *)  | 
| 21637 | 575  | 
|
576  | 
||
| 21867 | 577  | 
(*** PGIP identifier tables ***)  | 
578  | 
||
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
579  | 
(* TODO: these ones should be triggered by hooks after a  | 
| 22159 | 580  | 
declaration addition/removal, to be sent automatically. *)  | 
| 21867 | 581  | 
|
| 22159 | 582  | 
fun addids t  = issue_pgip (Addids {idtables = [t]})
 | 
583  | 
fun delids t  = issue_pgip (Delids {idtables = [t]})
 | 
|
| 21867 | 584  | 
|
| 21940 | 585  | 
fun askids (Askids vs) =  | 
| 21637 | 586  | 
let  | 
| 21940 | 587  | 
val url = #url vs (* ask for identifiers within a file *)  | 
588  | 
val thyname = #thyname vs (* ask for identifiers within a theory *)  | 
|
589  | 
val objtype = #objtype vs (* ask for identifiers of a particular type *)  | 
|
| 21867 | 590  | 
|
| 21940 | 591  | 
        fun idtable ty ctx ids = {objtype=ty,context=ctx,ids=ids}
 | 
| 21867 | 592  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
593  | 
        fun setids t = issue_pgip (Setids {idtables = [t]})
 | 
| 22159 | 594  | 
|
| 22225 | 595  | 
(* fake one-level nested "subtheories" by picking apart names. *)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
596  | 
val thms_of_thy =  | 
| 
22243
 
0f24888c8591
proper use of NameSpace.is_qualified (avoids compatibility issues of the SML B library);
 
wenzelm 
parents: 
22228 
diff
changeset
 | 
597  | 
map fst o NameSpace.extern_table o PureThy.theorems_of o ThyInfo.get_theory  | 
| 
 
0f24888c8591
proper use of NameSpace.is_qualified (avoids compatibility issues of the SML B library);
 
wenzelm 
parents: 
22228 
diff
changeset
 | 
598  | 
val immed_thms_of_thy = filter_out NameSpace.is_qualified o thms_of_thy  | 
| 22225 | 599  | 
fun thy_prefix s = case space_explode NameSpace.separator s of  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
600  | 
x::_::_ => SOME x (* String.find? *)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
601  | 
| _ => NONE  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
602  | 
fun subthys_of_thy s =  | 
| 23178 | 603  | 
List.foldl (fn (NONE,xs) => xs | (SOME x,xs) => insert op= x xs) []  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
604  | 
(map thy_prefix (thms_of_thy s))  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
605  | 
fun subthms_of_thy thy =  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
606  | 
(case thy_prefix thy of  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
607  | 
NONE => immed_thms_of_thy thy  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
608  | 
| SOME prf => filter (String.isPrefix (unprefix (prf ^ NameSpace.separator) thy))  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
609  | 
(thms_of_thy prf))  | 
| 22225 | 610  | 
val qualified_thms_of_thy = (* for global query with single response *)  | 
| 22590 | 611  | 
(map fst) o NameSpace.dest_table o PureThy.theorems_of o ThyInfo.get_theory;  | 
| 
22249
 
5460a5e4caa2
askids/qualified_thms_of_thy: version which returns names in domain of get_thm
 
aspinall 
parents: 
22243 
diff
changeset
 | 
612  | 
(* da: this version is equivalent to my previous, but splits up theorem sets with names  | 
| 
 
5460a5e4caa2
askids/qualified_thms_of_thy: version which returns names in domain of get_thm
 
aspinall 
parents: 
22243 
diff
changeset
 | 
613  | 
that I can't get to access later with get_thm. Anyway, would rather use sets.  | 
| 
 
5460a5e4caa2
askids/qualified_thms_of_thy: version which returns names in domain of get_thm
 
aspinall 
parents: 
22243 
diff
changeset
 | 
614  | 
Is above right way to get qualified names in that case? Filtering required again?  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
615  | 
map PureThy.get_name_hint o filter PureThy.has_name_hint o  | 
| 
22249
 
5460a5e4caa2
askids/qualified_thms_of_thy: version which returns names in domain of get_thm
 
aspinall 
parents: 
22243 
diff
changeset
 | 
616  | 
map snd o PureThy.thms_of o ThyInfo.get_theory; *)  | 
| 
 
5460a5e4caa2
askids/qualified_thms_of_thy: version which returns names in domain of get_thm
 
aspinall 
parents: 
22243 
diff
changeset
 | 
617  | 
in  | 
| 21940 | 618  | 
case (thyname,objtype) of  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
619  | 
(NONE, NONE) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
620  | 
setids (idtable ObjFile NONE (ThyInfo.names())) (*FIXME: uris*)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
621  | 
| (NONE, SOME ObjFile) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
622  | 
setids (idtable ObjFile NONE (ThyInfo.names())) (*FIXME: uris*)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
623  | 
| (SOME fi, SOME ObjFile) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
624  | 
setids (idtable ObjTheory (SOME fi) [fi]) (* TODO: check exists *)  | 
| 22225 | 625  | 
| (NONE, SOME ObjTheory) =>  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
626  | 
setids (idtable ObjTheory NONE (ThyInfo.names()))  | 
| 22225 | 627  | 
| (SOME thy, SOME ObjTheory) =>  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
628  | 
setids (idtable ObjTheory (SOME thy) (subthys_of_thy thy))  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
629  | 
| (SOME thy, SOME ObjTheorem) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
630  | 
setids (idtable ObjTheorem (SOME thy) (subthms_of_thy thy))  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
631  | 
| (NONE, SOME ObjTheorem) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
632  | 
(* A large query, but not unreasonable. ~5000 results for HOL.*)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
633  | 
(* Several setids should be allowed, but Eclipse code is currently broken:  | 
| 23226 | 634  | 
List.app (fn thy => setids (idtable ObjTheorem (SOME thy) (subthms_of_thy thy)))  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
635  | 
(ThyInfo.names()) *)  | 
| 22225 | 636  | 
setids (idtable ObjTheorem NONE (* this one gives ~7000 for HOL *)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
637  | 
(maps qualified_thms_of_thy (ThyInfo.names())))  | 
| 22225 | 638  | 
         | _ => warning ("askids: ignored argument combination")
 | 
| 21637 | 639  | 
end  | 
640  | 
||
| 22159 | 641  | 
fun askrefs (Askrefs vs) =  | 
642  | 
let  | 
|
643  | 
val url = #url vs (* ask for references of a file (i.e. immediate pre-requisites) *)  | 
|
644  | 
val thyname = #thyname vs (* ask for references of a theory (other theories) *)  | 
|
645  | 
val objtype = #objtype vs (* ask for references of a particular type... *)  | 
|
646  | 
val name = #name vs (* ... with this name *)  | 
|
647  | 
||
648  | 
        fun idtable ty ctx ids = {objtype=ty,context=ctx,ids=ids}
 | 
|
649  | 
||
650  | 
val thms_of_thy = map fst o PureThy.thms_of o ThyInfo.get_theory  | 
|
651  | 
||
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
652  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base  | 
| 22159 | 653  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
654  | 
fun filerefs f =  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
655  | 
let val thy = thy_name f  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
656  | 
val (_,filerefs) = OuterSyntax.deps_thy thy true f (* (Path.unpack f); *)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
657  | 
in  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
658  | 
                issue_pgip (Setrefs {url=url, thyname=NONE, objtype=SOME PgipTypes.ObjFile,
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
659  | 
name=NONE, idtables=[], fileurls=filerefs})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
660  | 
end  | 
| 22159 | 661  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
662  | 
fun thyrefs thy =  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
663  | 
let val ml_path = ThyLoad.ml_path thy  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
664  | 
val (thyrefs,_) = OuterSyntax.deps_thy thy true ml_path (* (Path.unpack f); *)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
665  | 
in  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
666  | 
                issue_pgip (Setrefs {url=url, thyname=thyname, objtype=SOME PgipTypes.ObjTheory,
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
667  | 
                                     name=NONE, idtables=[{context=NONE, objtype=PgipTypes.ObjTheory,
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
668  | 
ids=thyrefs}], fileurls=[]})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
669  | 
end  | 
| 22159 | 670  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
671  | 
fun thmrefs thmname =  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
672  | 
let  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
673  | 
(* TODO: interim: this is probably not right.  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
674  | 
What we want is mapping onto simple PGIP name/context model. *)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
675  | 
val ctx = Toplevel.context_of (Toplevel.get_state()) (* NB: raises UNDEF *)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
676  | 
val thy = Context.theory_of_proof ctx  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
677  | 
val ths = [PureThy.get_thm thy (PureThy.Name thmname)]  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
678  | 
val deps = filter_out (equal "")  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
679  | 
(Symtab.keys (fold Proofterm.thms_of_proof  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
680  | 
(map Thm.proof_of ths) Symtab.empty))  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
681  | 
in  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
682  | 
if null deps then ()  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
683  | 
                else issue_pgip (Setrefs {url=url, thyname=thyname, name=name,
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
684  | 
objtype=SOME PgipTypes.ObjTheorem,  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
685  | 
                                          idtables=[{context=NONE, objtype=PgipTypes.ObjTheorem,
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
686  | 
ids=deps}], fileurls=[]})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
687  | 
end  | 
| 22159 | 688  | 
in  | 
689  | 
case (url,thyname,objtype,name) of  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
690  | 
(SOME file, NONE, _, _) => filerefs file  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
691  | 
| (_,SOME thy,_,_) => thyrefs thy  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
692  | 
| (_,_,SOME PgipTypes.ObjTheorem,SOME thmname) => thmrefs thmname  | 
| 22159 | 693  | 
          | _  => error ("Unimplemented/invalid case of <askrefs>")
 | 
694  | 
end  | 
|
695  | 
||
696  | 
||
697  | 
||
| 21940 | 698  | 
fun showid (Showid vs) =  | 
| 21637 | 699  | 
let  | 
| 21940 | 700  | 
val thyname = #thyname vs  | 
701  | 
val objtype = #objtype vs  | 
|
702  | 
val name = #name vs  | 
|
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
703  | 
|
| 21940 | 704  | 
val topthy = Toplevel.theory_of o Toplevel.get_state  | 
| 21637 | 705  | 
|
| 22590 | 706  | 
fun splitthy id =  | 
707  | 
let val comps = NameSpace.explode id  | 
|
708  | 
in case comps of  | 
|
709  | 
(thy::(rest as _::_)) => (ThyInfo.get_theory thy, space_implode "." rest)  | 
|
710  | 
| [plainid] => (topthy(),plainid)  | 
|
711  | 
| _ => raise Toplevel.UNDEF (* assert false *)  | 
|
712  | 
end  | 
|
713  | 
||
| 21637 | 714  | 
|
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
715  | 
fun idvalue strings =  | 
| 
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
716  | 
            issue_pgip (Idvalue { thyname=thyname, objtype=objtype, name=name, 
 | 
| 22590 | 717  | 
                                  text=[XML.Elem("pgmltext",[],
 | 
718  | 
map XML.Rawtext strings)] })  | 
|
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
719  | 
|
| 22590 | 720  | 
fun string_of_thm th = Output.output  | 
721  | 
(Pretty.string_of  | 
|
722  | 
(Display.pretty_thm_aux  | 
|
723  | 
(Sign.pp (Thm.theory_of_thm th))  | 
|
724  | 
false (* quote *)  | 
|
725  | 
false (* show hyps *)  | 
|
726  | 
[] (* asms *)  | 
|
727  | 
th))  | 
|
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
728  | 
|
| 
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
729  | 
fun strings_of_thm (thy, name) = map string_of_thm (get_thms thy (Name name))  | 
| 
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
730  | 
|
| 22590 | 731  | 
val string_of_thy = Output.output o  | 
732  | 
Pretty.string_of o (ProofDisplay.pretty_full_theory false)  | 
|
| 21940 | 733  | 
in  | 
734  | 
case (thyname, objtype) of  | 
|
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
735  | 
(_,ObjTheory) => idvalue [string_of_thy (ThyInfo.get_theory name)]  | 
| 
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
736  | 
| (SOME thy, ObjTheorem) => idvalue (strings_of_thm (ThyInfo.get_theory thy, name))  | 
| 
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
737  | 
| (NONE, ObjTheorem) => idvalue (strings_of_thm (splitthy name))  | 
| 21940 | 738  | 
          | (_, ot) => error ("Cannot show objects of type "^(PgipTypes.name_of_objtype ot))
 | 
| 21637 | 739  | 
end  | 
740  | 
||
| 21867 | 741  | 
(*** Inspecting state ***)  | 
742  | 
||
| 21940 | 743  | 
(* The file which is currently being processed interactively.  | 
| 21637 | 744  | 
In the pre-PGIP code, this was informed to Isabelle and the theory loader  | 
745  | 
on completion, but that allows for circularity in case we read  | 
|
746  | 
ourselves. So PGIP opens the filename at the start of a script.  | 
|
747  | 
We ought to prevent problems by modifying the theory loader to know  | 
|
| 21940 | 748  | 
about this special status, but for now we just keep a local reference.  | 
749  | 
*)  | 
|
| 21637 | 750  | 
|
751  | 
val currently_open_file = ref (NONE : pgipurl option)  | 
|
752  | 
||
| 22163 | 753  | 
fun get_currently_open_file () = ! currently_open_file;  | 
754  | 
||
| 21940 | 755  | 
fun askguise vs =  | 
| 21637 | 756  | 
(* The "guise" is the PGIP abstraction of the prover's state.  | 
757  | 
The <informguise> message is merely used for consistency checking. *)  | 
|
| 21940 | 758  | 
let  | 
759  | 
val openfile = !currently_open_file  | 
|
| 21637 | 760  | 
|
| 21940 | 761  | 
val topthy = Toplevel.theory_of o Toplevel.get_state  | 
762  | 
val topthy_name = Context.theory_name o topthy  | 
|
| 21637 | 763  | 
|
| 21940 | 764  | 
val opentheory = SOME (topthy_name()) handle Toplevel.UNDEF => NONE  | 
| 21637 | 765  | 
|
| 21940 | 766  | 
fun topproofpos () = try Toplevel.proof_position_of (Isar.state ());  | 
767  | 
val openproofpos = topproofpos()  | 
|
| 21637 | 768  | 
in  | 
769  | 
        issue_pgip (Informguise { file = openfile,
 | 
|
| 21940 | 770  | 
theory = opentheory,  | 
771  | 
(* would be nice to get thm name... *)  | 
|
772  | 
theorem = NONE,  | 
|
773  | 
proofpos = openproofpos })  | 
|
| 21637 | 774  | 
end  | 
775  | 
||
| 21902 | 776  | 
fun parsescript (Parsescript vs) =  | 
| 21637 | 777  | 
let  | 
| 21940 | 778  | 
val text = #text vs  | 
779  | 
val systemdata = #systemdata vs  | 
|
780  | 
val location = #location vs (* TODO: extract position *)  | 
|
| 21637 | 781  | 
|
| 21867 | 782  | 
val _ = start_delay_msgs () (* gather parsing errs/warns *)  | 
783  | 
val doc = PgipParser.pgip_parser text  | 
|
| 21637 | 784  | 
val errs = end_delayed_msgs ()  | 
785  | 
||
| 21940 | 786  | 
val sysattrs = PgipTypes.opt_attr "systemdata" systemdata  | 
787  | 
val locattrs = PgipTypes.attrs_of_location location  | 
|
| 21637 | 788  | 
in  | 
789  | 
        issue_pgip (Parseresult { attrs= sysattrs@locattrs,
 | 
|
| 21940 | 790  | 
doc = doc,  | 
791  | 
errs = map PgipOutput.output errs })  | 
|
| 21637 | 792  | 
end  | 
793  | 
||
794  | 
fun showproofstate vs = isarcmd "pr"  | 
|
795  | 
||
| 21969 | 796  | 
fun showctxt vs = isarcmd "print_context"  | 
| 21637 | 797  | 
|
| 21902 | 798  | 
fun searchtheorems (Searchtheorems vs) =  | 
| 21940 | 799  | 
let  | 
800  | 
val arg = #arg vs  | 
|
| 21637 | 801  | 
in  | 
| 21969 | 802  | 
        isarcmd ("find_theorems " ^ arg)
 | 
| 21637 | 803  | 
end  | 
804  | 
||
| 21940 | 805  | 
fun setlinewidth (Setlinewidth vs) =  | 
806  | 
let  | 
|
807  | 
val width = #width vs  | 
|
| 21637 | 808  | 
in  | 
| 21940 | 809  | 
        isarcmd ("pretty_setmargin " ^ Int.toString width) (* FIXME: conversion back/forth! *)
 | 
| 21637 | 810  | 
end  | 
811  | 
||
| 21902 | 812  | 
fun viewdoc (Viewdoc vs) =  | 
| 21940 | 813  | 
let  | 
814  | 
val arg = #arg vs  | 
|
815  | 
in  | 
|
816  | 
        isarcmd ("print_" ^ arg)   (* FIXME: isatool doc?.  Return URLs, maybe? *)
 | 
|
| 21637 | 817  | 
end  | 
818  | 
||
| 21867 | 819  | 
(*** Theory ***)  | 
820  | 
||
| 21902 | 821  | 
fun doitem (Doitem vs) =  | 
| 21637 | 822  | 
let  | 
| 21940 | 823  | 
val text = #text vs  | 
| 21637 | 824  | 
in  | 
| 21940 | 825  | 
isarcmd text  | 
| 21637 | 826  | 
end  | 
827  | 
||
828  | 
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
 | 
829  | 
isarcmd "undo"  | 
| 21637 | 830  | 
|
831  | 
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
 | 
832  | 
isarcmd "redo"  | 
| 21637 | 833  | 
|
| 21940 | 834  | 
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
 | 
835  | 
isarcmd "kill" (* was: "init_toplevel" *)  | 
| 21637 | 836  | 
|
| 21902 | 837  | 
fun retracttheory (Retracttheory vs) =  | 
| 21940 | 838  | 
let  | 
839  | 
val thyname = #thyname vs  | 
|
| 21637 | 840  | 
in  | 
| 21940 | 841  | 
        isarcmd ("kill_thy " ^ quote thyname)
 | 
| 21637 | 842  | 
end  | 
843  | 
||
| 21867 | 844  | 
|
845  | 
(*** Files ***)  | 
|
846  | 
||
847  | 
(* Path management: we allow theory files to have dependencies in  | 
|
848  | 
their own directory, but when we change directory for a new file we  | 
|
849  | 
remove the path. Leaving it there can cause confusion with  | 
|
850  | 
difference in batch mode.  | 
|
| 21940 | 851  | 
NB: PGIP does not assume that the prover has a load path.  | 
| 21867 | 852  | 
*)  | 
853  | 
||
854  | 
local  | 
|
855  | 
val current_working_dir = ref (NONE : string option)  | 
|
856  | 
in  | 
|
| 21940 | 857  | 
fun changecwd_dir newdirpath =  | 
858  | 
let  | 
|
| 21867 | 859  | 
val newdir = File.platform_path newdirpath  | 
| 21940 | 860  | 
in  | 
| 21867 | 861  | 
(case (!current_working_dir) of  | 
862  | 
NONE => ()  | 
|
863  | 
| SOME dir => ThyLoad.del_path dir;  | 
|
864  | 
ThyLoad.add_path newdir;  | 
|
865  | 
current_working_dir := SOME newdir)  | 
|
866  | 
end  | 
|
867  | 
end  | 
|
868  | 
||
| 21940 | 869  | 
fun changecwd (Changecwd vs) =  | 
870  | 
let  | 
|
871  | 
val url = #url vs  | 
|
872  | 
val newdir = PgipTypes.path_of_pgipurl url  | 
|
| 21867 | 873  | 
in  | 
| 21940 | 874  | 
changecwd_dir url  | 
| 21867 | 875  | 
end  | 
876  | 
||
| 21902 | 877  | 
fun openfile (Openfile vs) =  | 
| 21940 | 878  | 
let  | 
| 21867 | 879  | 
val url = #url vs  | 
880  | 
val filepath = PgipTypes.path_of_pgipurl url  | 
|
881  | 
val filedir = Path.dir filepath  | 
|
882  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base  | 
|
883  | 
val openfile_retract = Output.no_warnings (ThyInfo.if_known_thy ThyInfo.remove_thy) o thy_name;  | 
|
884  | 
in  | 
|
885  | 
case !currently_open_file of  | 
|
| 22028 | 886  | 
          SOME f => raise PGIP ("<openfile> when a file is already open!\nCurrently open file: " ^
 | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
887  | 
PgipTypes.string_of_pgipurl url)  | 
| 21867 | 888  | 
| NONE => (openfile_retract filepath;  | 
| 21940 | 889  | 
changecwd_dir filedir;  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
890  | 
                   priority ("Working in file: " ^ PgipTypes.string_of_pgipurl url);
 | 
| 21940 | 891  | 
currently_open_file := SOME url)  | 
| 21867 | 892  | 
end  | 
893  | 
||
894  | 
fun closefile vs =  | 
|
895  | 
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
 | 
896  | 
SOME f => (proper_inform_file_processed f (Isar.state());  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
897  | 
                   priority ("Finished working in file: " ^ PgipTypes.string_of_pgipurl f);
 | 
| 21867 | 898  | 
currently_open_file := NONE)  | 
899  | 
      | NONE => raise PGIP ("<closefile> when no file is open!")
 | 
|
900  | 
||
| 21940 | 901  | 
fun loadfile (Loadfile vs) =  | 
902  | 
let  | 
|
903  | 
val url = #url vs  | 
|
904  | 
in  | 
|
| 
22171
 
58f554f51f0d
Let <loadfile> execute even while a file is open interactively.
 
aspinall 
parents: 
22163 
diff
changeset
 | 
905  | 
(* da: this doesn't seem to cause a problem, batch loading uses  | 
| 
 
58f554f51f0d
Let <loadfile> execute even while a file is open interactively.
 
aspinall 
parents: 
22163 
diff
changeset
 | 
906  | 
a different state context. Of course confusion is still possible,  | 
| 
 
58f554f51f0d
Let <loadfile> execute even while a file is open interactively.
 
aspinall 
parents: 
22163 
diff
changeset
 | 
907  | 
e.g. file loaded depends on open file which is not yet saved. *)  | 
| 
 
58f554f51f0d
Let <loadfile> execute even while a file is open interactively.
 
aspinall 
parents: 
22163 
diff
changeset
 | 
908  | 
(* case !currently_open_file of  | 
| 22028 | 909  | 
            SOME f => raise PGIP ("<loadfile> when a file is open!\nCurrently open file: " ^
 | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
910  | 
PgipTypes.string_of_pgipurl url)  | 
| 
22171
 
58f554f51f0d
Let <loadfile> execute even while a file is open interactively.
 
aspinall 
parents: 
22163 
diff
changeset
 | 
911  | 
| NONE => *)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
912  | 
use_thy_or_ml_file (File.platform_path url)  | 
| 21637 | 913  | 
end  | 
914  | 
||
915  | 
fun abortfile vs =  | 
|
916  | 
case !currently_open_file of  | 
|
917  | 
SOME f => (isarcmd "init_toplevel";  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
918  | 
                   priority ("Aborted working in file: " ^
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
919  | 
PgipTypes.string_of_pgipurl f);  | 
| 21940 | 920  | 
currently_open_file := NONE)  | 
| 21637 | 921  | 
      | NONE => raise PGIP ("<abortfile> when no file is open!")
 | 
922  | 
||
| 21940 | 923  | 
fun retractfile (Retractfile vs) =  | 
924  | 
let  | 
|
925  | 
val url = #url vs  | 
|
| 21637 | 926  | 
in  | 
| 21940 | 927  | 
case !currently_open_file of  | 
| 21637 | 928  | 
            SOME f => raise PGIP ("<retractfile> when a file is open!")
 | 
| 22028 | 929  | 
          | NONE => (priority ("Retracting file: " ^ PgipTypes.string_of_pgipurl url);
 | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
930  | 
(* TODO: next should be in thy loader, here just for testing *)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
931  | 
let  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
932  | 
val name = thy_name url  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
933  | 
in List.app (tell_file_retracted false) (ThyInfo.loaded_files name) end;  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
934  | 
inform_file_retracted url)  | 
| 21637 | 935  | 
end  | 
936  | 
||
937  | 
||
| 21867 | 938  | 
(*** System ***)  | 
| 21637 | 939  | 
|
| 21902 | 940  | 
fun systemcmd (Systemcmd vs) =  | 
| 21940 | 941  | 
let  | 
| 21637 | 942  | 
val arg = #arg vs  | 
943  | 
in  | 
|
944  | 
isarcmd arg  | 
|
945  | 
end  | 
|
946  | 
||
947  | 
exception PGIP_QUIT;  | 
|
948  | 
fun quitpgip vs = raise PGIP_QUIT  | 
|
949  | 
||
| 21902 | 950  | 
fun process_input inp = case inp  | 
951  | 
of Pgip.Askpgip _ => askpgip inp  | 
|
952  | 
| Pgip.Askpgml _ => askpgml inp  | 
|
| 21940 | 953  | 
| Pgip.Askprefs _ => askprefs inp  | 
| 21902 | 954  | 
| Pgip.Askconfig _ => askconfig inp  | 
955  | 
| Pgip.Getpref _ => getpref inp  | 
|
956  | 
| Pgip.Setpref _ => setpref inp  | 
|
957  | 
| Pgip.Proverinit _ => proverinit inp  | 
|
958  | 
| Pgip.Proverexit _ => proverexit inp  | 
|
| 22408 | 959  | 
| Pgip.Setproverflag _ => setproverflag inp  | 
| 21902 | 960  | 
| Pgip.Dostep _ => dostep inp  | 
961  | 
| Pgip.Undostep _ => undostep inp  | 
|
962  | 
| Pgip.Redostep _ => redostep inp  | 
|
| 
23435
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
963  | 
| Pgip.Forget _ => error "<forget> not implemented by Isabelle"  | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
964  | 
| Pgip.Restoregoal _ => error "<restoregoal> not implemented by Isabelle"  | 
| 21902 | 965  | 
| Pgip.Abortgoal _ => abortgoal inp  | 
966  | 
| Pgip.Askids _ => askids inp  | 
|
| 22159 | 967  | 
| Pgip.Askrefs _ => askrefs inp  | 
| 21902 | 968  | 
| Pgip.Showid _ => showid inp  | 
969  | 
| Pgip.Askguise _ => askguise inp  | 
|
970  | 
| Pgip.Parsescript _ => parsescript inp  | 
|
971  | 
| Pgip.Showproofstate _ => showproofstate inp  | 
|
972  | 
| Pgip.Showctxt _ => showctxt inp  | 
|
973  | 
| Pgip.Searchtheorems _ => searchtheorems inp  | 
|
974  | 
| Pgip.Setlinewidth _ => setlinewidth inp  | 
|
975  | 
| Pgip.Viewdoc _ => viewdoc inp  | 
|
976  | 
| Pgip.Doitem _ => doitem inp  | 
|
977  | 
| Pgip.Undoitem _ => undoitem inp  | 
|
978  | 
| Pgip.Redoitem _ => redoitem inp  | 
|
979  | 
| Pgip.Aborttheory _ => aborttheory inp  | 
|
980  | 
| Pgip.Retracttheory _ => retracttheory inp  | 
|
981  | 
| Pgip.Loadfile _ => loadfile inp  | 
|
982  | 
| Pgip.Openfile _ => openfile inp  | 
|
983  | 
| Pgip.Closefile _ => closefile inp  | 
|
984  | 
| Pgip.Abortfile _ => abortfile inp  | 
|
985  | 
| Pgip.Retractfile _ => retractfile inp  | 
|
986  | 
| Pgip.Changecwd _ => changecwd inp  | 
|
987  | 
| Pgip.Systemcmd _ => systemcmd inp  | 
|
988  | 
| Pgip.Quitpgip _ => quitpgip inp  | 
|
| 21637 | 989  | 
|
990  | 
||
| 21940 | 991  | 
fun process_pgip_element pgipxml =  | 
| 21637 | 992  | 
case pgipxml of  | 
| 21969 | 993  | 
xml as (XML.Elem elem) =>  | 
| 21940 | 994  | 
(case Pgip.input elem of  | 
995  | 
             NONE => warning ("Unrecognized PGIP command, ignored: \n" ^
 | 
|
996  | 
(XML.string_of_tree xml))  | 
|
997  | 
| SOME inp => (process_input inp)) (* errors later; packet discarded *)  | 
|
| 21969 | 998  | 
| XML.Text t => ignored_text_warning t  | 
999  | 
| XML.Rawtext t => ignored_text_warning t  | 
|
| 21637 | 1000  | 
and ignored_text_warning t =  | 
| 21940 | 1001  | 
if size (Symbol.strip_blanks t) > 0 then  | 
1002  | 
           warning ("Ignored text in PGIP packet: \n" ^ t)
 | 
|
| 21637 | 1003  | 
else ()  | 
1004  | 
||
1005  | 
fun process_pgip_tree xml =  | 
|
1006  | 
(pgip_refid := NONE;  | 
|
1007  | 
pgip_refseq := NONE;  | 
|
1008  | 
(case xml of  | 
|
1009  | 
          XML.Elem ("pgip", attrs, pgips) =>
 | 
|
1010  | 
(let  | 
|
1011  | 
val class = PgipTypes.get_attr "class" attrs  | 
|
1012  | 
val dest = PgipTypes.get_attr_opt "destid" attrs  | 
|
| 21940 | 1013  | 
val seq = PgipTypes.read_pgipnat (PgipTypes.get_attr "seq" attrs)  | 
| 21637 | 1014  | 
(* Respond to prover broadcasts, or messages for us. Ignore rest *)  | 
| 21940 | 1015  | 
val processit =  | 
1016  | 
case dest of  | 
|
| 21637 | 1017  | 
NONE => class = "pa"  | 
| 21940 | 1018  | 
| SOME id => matching_pgip_id id  | 
1019  | 
in if processit then  | 
|
1020  | 
(pgip_refid := PgipTypes.get_attr_opt "id" attrs;  | 
|
1021  | 
pgip_refseq := SOME seq;  | 
|
1022  | 
List.app process_pgip_element pgips;  | 
|
1023  | 
(* return true to indicate <ready/> *)  | 
|
1024  | 
true)  | 
|
1025  | 
else  | 
|
1026  | 
(* no response to ignored messages. *)  | 
|
1027  | 
false  | 
|
| 21637 | 1028  | 
end)  | 
1029  | 
| _ => raise PGIP "Invalid PGIP packet received")  | 
|
1030  | 
handle PGIP msg =>  | 
|
1031  | 
(Output.error_msg ((msg ^ "\nPGIP error occured in XML text below:\n") ^  | 
|
| 21940 | 1032  | 
(XML.string_of_tree xml));  | 
1033  | 
true))  | 
|
| 21637 | 1034  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1035  | 
(* External input *)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1036  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1037  | 
val process_pgip_plain = K () o process_pgip_tree o XML.tree_of_string  | 
| 21637 | 1038  | 
|
1039  | 
(* PGIP loop: process PGIP input only *)  | 
|
1040  | 
||
1041  | 
local  | 
|
1042  | 
||
1043  | 
exception XML_PARSE  | 
|
1044  | 
||
1045  | 
fun loop ready src =  | 
|
1046  | 
let  | 
|
1047  | 
val _ = if ready then issue_pgip (Ready ()) else ()  | 
|
| 21969 | 1048  | 
val pgipo =  | 
1049  | 
(case try Source.get_single src of  | 
|
1050  | 
SOME pgipo => pgipo  | 
|
1051  | 
| NONE => raise XML_PARSE)  | 
|
| 21637 | 1052  | 
in  | 
1053  | 
case pgipo of  | 
|
1054  | 
NONE => ()  | 
|
1055  | 
| SOME (pgip,src') =>  | 
|
1056  | 
let  | 
|
| 21940 | 1057  | 
val ready' = (process_pgip_tree pgip)  | 
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
1058  | 
handle PGIP_QUIT => raise PGIP_QUIT  | 
| 22590 | 1059  | 
| e => (handler (e,SOME src'); true)  | 
| 21637 | 1060  | 
in  | 
1061  | 
loop ready' src'  | 
|
1062  | 
end  | 
|
1063  | 
end handle e => handler (e,SOME src) (* error in XML parse or Ready issue *)  | 
|
1064  | 
||
1065  | 
and handler (e,srco) =  | 
|
1066  | 
case (e,srco) of  | 
|
1067  | 
(XML_PARSE,SOME src) =>  | 
|
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
1068  | 
panic "Invalid XML input, aborting" (* TODO: attempt recovery *)  | 
| 21940 | 1069  | 
| (Interrupt,SOME src) =>  | 
1070  | 
(Output.error_msg "Interrupt during PGIP processing"; loop true src)  | 
|
1071  | 
| (Toplevel.UNDEF,SOME src) =>  | 
|
1072  | 
(Output.error_msg "No working context defined"; loop true src)  | 
|
1073  | 
| (e,SOME src) =>  | 
|
1074  | 
(Output.error_msg (Toplevel.exn_message e); loop true src)  | 
|
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
1075  | 
| (PGIP_QUIT,_) => ()  | 
| 21637 | 1076  | 
| (_,NONE) => ()  | 
1077  | 
in  | 
|
1078  | 
(* TODO: add socket interface *)  | 
|
1079  | 
||
1080  | 
val xmlP = XML.scan_comment_whspc |-- XML.parse_elem >> single  | 
|
1081  | 
||
1082  | 
val tty_src = Source.set_prompt "" (Source.source Symbol.stopper xmlP NONE Source.tty)  | 
|
1083  | 
||
1084  | 
fun pgip_toplevel x = loop true x  | 
|
1085  | 
end  | 
|
1086  | 
||
1087  | 
||
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1088  | 
local  | 
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1089  | 
(* Extra command for embedding prover-control inside document (obscure/debug usage). *)  | 
| 21637 | 1090  | 
|
| 
22014
 
4b70cbd96007
removed Toplevel.print_exn hook -- existing error_msg_fn does the job;
 
wenzelm 
parents: 
22001 
diff
changeset
 | 
1091  | 
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
 | 
1092  | 
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
 | 
1093  | 
(OuterParse.text >> (Toplevel.no_timing oo  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1094  | 
(fn txt => Toplevel.imperative (fn () => process_pgip_plain txt))));  | 
| 21637 | 1095  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1096  | 
in  | 
| 21637 | 1097  | 
|
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1098  | 
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
 | 
1099  | 
|
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1100  | 
end  | 
| 
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1101  | 
|
| 21637 | 1102  | 
|
1103  | 
||
1104  | 
(* init *)  | 
|
1105  | 
||
1106  | 
val initialized = ref false;  | 
|
1107  | 
||
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
1108  | 
fun init_pgip false = panic "No Proof General interface support for Isabelle/classic mode."  | 
| 21969 | 1109  | 
| init_pgip true =  | 
1110  | 
(! initialized orelse  | 
|
| 22590 | 1111  | 
(Output.no_warnings init_outer_syntax ();  | 
| 
22014
 
4b70cbd96007
removed Toplevel.print_exn hook -- existing error_msg_fn does the job;
 
wenzelm 
parents: 
22001 
diff
changeset
 | 
1112  | 
PgipParser.init ();  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
1113  | 
setup_preferences_tweak ();  | 
| 22590 | 1114  | 
setup_proofgeneral_output ();  | 
| 21969 | 1115  | 
setup_messages ();  | 
1116  | 
setup_state ();  | 
|
1117  | 
setup_thy_loader ();  | 
|
1118  | 
setup_present_hook ();  | 
|
1119  | 
init_pgip_session_id ();  | 
|
1120  | 
welcome ();  | 
|
1121  | 
set initialized);  | 
|
1122  | 
sync_thy_loader ();  | 
|
1123  | 
change print_mode (cons proof_generalN o remove (op =) proof_generalN);  | 
|
1124  | 
pgip_toplevel tty_src);  | 
|
| 21637 | 1125  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1126  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1127  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1128  | 
(** Out-of-loop PGIP commands (for Emacs hybrid mode) **)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1129  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1130  | 
local  | 
| 22590 | 1131  | 
val pgip_output_channel = ref Output.writeln_default  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1132  | 
in  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1133  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1134  | 
(* Set recipient for PGIP results *)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1135  | 
fun init_pgip_channel writefn =  | 
| 21940 | 1136  | 
(init_pgip_session_id();  | 
1137  | 
pgip_output_channel := writefn)  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1138  | 
|
| 21940 | 1139  | 
(* Process a PGIP command.  | 
1140  | 
This works for preferences but not generally guaranteed  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1141  | 
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
 | 
1142  | 
fun process_pgip str =  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1143  | 
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
 | 
1144  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1145  | 
end  | 
| 21637 | 1146  | 
|
1147  | 
end;  |