| author | wenzelm | 
| Sun, 16 Nov 2008 20:03:42 +0100 | |
| changeset 28814 | 463c9e9111ae | 
| parent 28809 | 7c2e1bbf3c36 | 
| child 28817 | c8cc94a470d4 | 
| 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  | 
|
| 
28100
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
11  | 
val new_thms_deps: theory -> theory -> string list * string list  | 
| 21940 | 12  | 
val init_pgip: bool -> unit (* main PGIP loop with true; fail with false *)  | 
| 21642 | 13  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
14  | 
(* These two are just to support the semi-PGIP Emacs mode *)  | 
| 21940 | 15  | 
val init_pgip_channel: (string -> unit) -> unit  | 
16  | 
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
 | 
17  | 
|
| 
23435
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
18  | 
(* More message functions... *)  | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
19  | 
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
 | 
20  | 
val log_msg : string -> unit (* for internal log messages *)  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
21  | 
val error_with_pos : PgipTypes.displayarea -> PgipTypes.fatality -> Position.T -> string -> unit  | 
| 22159 | 22  | 
|
| 22163 | 23  | 
val get_currently_open_file : unit -> Path.T option (* interface focus *)  | 
| 28588 | 24  | 
val add_preference: string -> Preferences.preference -> unit  | 
| 21637 | 25  | 
end  | 
26  | 
||
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
27  | 
structure ProofGeneralPgip : PROOF_GENERAL_PGIP =  | 
| 21637 | 28  | 
struct  | 
29  | 
||
30  | 
open Pgip;  | 
|
31  | 
||
| 21949 | 32  | 
|
| 
23641
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
33  | 
(** print mode **)  | 
| 21637 | 34  | 
|
| 22408 | 35  | 
val proof_generalN = "ProofGeneral";  | 
36  | 
val pgmlsymbols_flag = ref true;  | 
|
| 21637 | 37  | 
|
| 
23641
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
38  | 
|
| 
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
39  | 
(* assembling and issuing PGIP packets *)  | 
| 21637 | 40  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
41  | 
val pgip_refid = ref NONE: string option ref;  | 
| 21637 | 42  | 
val pgip_refseq = ref NONE: int option ref;  | 
43  | 
||
44  | 
local  | 
|
45  | 
val pgip_class = "pg"  | 
|
46  | 
val pgip_tag = "Isabelle/Isar"  | 
|
47  | 
val pgip_id = ref ""  | 
|
48  | 
val pgip_seq = ref 0  | 
|
49  | 
fun pgip_serial () = inc pgip_seq  | 
|
50  | 
||
51  | 
fun assemble_pgips pgips =  | 
|
| 21940 | 52  | 
    Pgip { tag = SOME pgip_tag,
 | 
53  | 
class = pgip_class,  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
54  | 
seq = pgip_serial (),  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
55  | 
id = ! pgip_id,  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
56  | 
destid = ! pgip_refid,  | 
| 21940 | 57  | 
(* destid=refid since Isabelle only communicates back to sender *)  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
58  | 
refid = ! pgip_refid,  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
59  | 
refseq = ! pgip_refseq,  | 
| 21940 | 60  | 
content = pgips }  | 
| 21637 | 61  | 
in  | 
62  | 
||
63  | 
fun init_pgip_session_id () =  | 
|
64  | 
pgip_id := getenv "HOSTNAME" ^ "/" ^ getenv "USER" ^ "/" ^  | 
|
65  | 
getenv "ISABELLE_PID" ^ "/" ^ Time.toString (Time.now ())  | 
|
66  | 
||
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
67  | 
fun matching_pgip_id id = (id = ! pgip_id)  | 
| 21637 | 68  | 
|
| 22590 | 69  | 
val output_xml_fn = ref Output.writeln_default  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
70  | 
fun output_xml s = ! output_xml_fn (XML.string_of s);  | 
| 
23641
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
71  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
72  | 
val output_pgips = XML.string_of o PgipOutput.output o assemble_pgips o map PgipOutput.output;  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
73  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
74  | 
val output_pgmlterm = XML.string_of o Pgml.pgmlterm_to_xml;  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
75  | 
val output_pgmltext = XML.string_of o Pgml.pgml_to_xml;  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
76  | 
|
| 
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
77  | 
|
| 21940 | 78  | 
fun issue_pgip_rawtext str =  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
79  | 
output_xml (PgipOutput.output (assemble_pgips (YXML.parse_body str)));  | 
| 21637 | 80  | 
|
81  | 
fun issue_pgip pgipop =  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
82  | 
output_xml (PgipOutput.output (assemble_pgips [PgipOutput.output pgipop]));  | 
| 21637 | 83  | 
|
84  | 
end;  | 
|
85  | 
||
86  | 
||
| 
23641
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
87  | 
|
| 21637 | 88  | 
(** messages and notification **)  | 
89  | 
||
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
90  | 
(* PGML terms *)  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
91  | 
|
| 21637 | 92  | 
local  | 
93  | 
||
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
94  | 
fun pgml_sym s =  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
95  | 
if ! pgmlsymbols_flag then  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
96  | 
(case Symbol.decode s of  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
97  | 
      Symbol.Sym name => Pgml.Sym {name = name, content = s}
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
98  | 
| _ => Pgml.Str s)  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
99  | 
else Pgml.Str s;  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
100  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
101  | 
val pgml_syms = map pgml_sym o Symbol.explode;  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
102  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
103  | 
val token_markups =  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
104  | 
[Markup.tclassN, Markup.tfreeN, Markup.tvarN, Markup.freeN,  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
105  | 
Markup.boundN, Markup.varN, Markup.skolemN];  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
106  | 
|
| 21637 | 107  | 
in  | 
108  | 
||
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
109  | 
fun pgml_terms (XML.Elem (name, atts, body)) =  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
110  | 
if member (op =) token_markups name then  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
111  | 
let val content = pgml_syms (Buffer.content (fold XML.add_content body Buffer.empty))  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
112  | 
        in [Pgml.Atoms {kind = SOME name, content = content}] end
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
113  | 
else  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
114  | 
let val content = maps pgml_terms body in  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
115  | 
if name = Markup.blockN then  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
116  | 
            [Pgml.Box {orient = NONE, indent = Properties.get_int atts Markup.indentN, content = content}]
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
117  | 
else if name = Markup.breakN then  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
118  | 
            [Pgml.Break {mandatory = NONE, indent = Properties.get_int atts Markup.widthN}]
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
119  | 
else if name = Markup.fbreakN then  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
120  | 
            [Pgml.Break {mandatory = SOME true, indent = NONE}]
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
121  | 
else content  | 
| 21940 | 122  | 
end  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
123  | 
| pgml_terms (XML.Text text) = map (Pgml.Raw o Pgml.atom_to_xml) (pgml_syms text);  | 
| 21637 | 124  | 
|
125  | 
end;  | 
|
126  | 
||
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
127  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
128  | 
(* messages *)  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
129  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
130  | 
fun pgml area content =  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
131  | 
  Pgml.Pgml {version = NONE, systemid = NONE, area = SOME area, content = content};
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
132  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
133  | 
fun message_content default_area s =  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
134  | 
let  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
135  | 
val body = YXML.parse_body s;  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
136  | 
val area =  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
137  | 
(case body of  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
138  | 
[XML.Elem (name, _, _)] =>  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
139  | 
if name = Markup.stateN then PgipTypes.Display else default_area  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
140  | 
| _ => default_area);  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
141  | 
in Pgml.pgml_to_xml (pgml area (maps pgml_terms body)) end;  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
142  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
143  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
144  | 
fun normalmsg area s = issue_pgip  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
145  | 
  (Normalresponse {content = [message_content area s]});
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
146  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
147  | 
fun errormsg area fatality s = issue_pgip  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
148  | 
  (Errorresponse {fatality = fatality, location = NONE, content = [message_content area s]});
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
149  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
150  | 
(*error responses with useful locations*)  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
151  | 
fun error_with_pos area fatality pos s = issue_pgip  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
152  | 
  (Errorresponse {
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
153  | 
fatality = fatality,  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
154  | 
location = SOME (PgipIsabelle.location_of_position pos),  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
155  | 
content = [message_content area s]});  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
156  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
157  | 
fun panic s = (errormsg Message Panic ("## SYSTEM EXIT ##\n" ^ s); exit 1);
 | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
158  | 
fun nonfatal_error s = errormsg Message Nonfatal s;  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
159  | 
fun log_msg s = errormsg Message Log s;  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
160  | 
|
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
161  | 
(* NB: 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
 | 
162  | 
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
 | 
163  | 
can't be written without newlines. *)  | 
| 21637 | 164  | 
fun setup_messages () =  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
165  | 
(Output.writeln_fn := (fn s => normalmsg Message s);  | 
| 27604 | 166  | 
Output.status_fn := (fn _ => ());  | 
| 23840 | 167  | 
Output.priority_fn := (fn s => normalmsg Status s);  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
168  | 
Output.tracing_fn := (fn s => normalmsg Tracing s);  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
169  | 
Output.warning_fn := (fn s => errormsg Message Warning s);  | 
| 
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
170  | 
Output.error_fn := (fn s => errormsg Message Fatal s);  | 
| 
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
171  | 
Output.debug_fn := (fn s => errormsg Message Debug s));  | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
172  | 
|
| 21637 | 173  | 
|
174  | 
(* immediate messages *)  | 
|
175  | 
||
| 26706 | 176  | 
fun tell_clear_goals () =  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
177  | 
    issue_pgip (Normalresponse { content = [Pgml.pgml_to_xml (pgml Display [])] })
 | 
| 26706 | 178  | 
fun tell_clear_response () =  | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
179  | 
    issue_pgip (Normalresponse { content = [Pgml.pgml_to_xml (pgml Message [])] })
 | 
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
180  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
181  | 
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
 | 
182  | 
    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
 | 
183  | 
completed=completed})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
184  | 
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
 | 
185  | 
    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
 | 
186  | 
completed=completed})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
187  | 
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
 | 
188  | 
    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
 | 
189  | 
completed=completed})  | 
| 21637 | 190  | 
|
191  | 
||
192  | 
(* theory loader actions *)  | 
|
193  | 
||
194  | 
local  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
195  | 
(* 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
 | 
196  | 
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
 | 
197  | 
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
 | 
198  | 
(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
 | 
199  | 
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
 | 
200  | 
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
 | 
201  | 
operations).  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
202  | 
*)  | 
| 21637 | 203  | 
fun trace_action action name =  | 
204  | 
if action = ThyInfo.Update then  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
205  | 
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
 | 
206  | 
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
 | 
207  | 
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
 | 
208  | 
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
 | 
209  | 
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
 | 
210  | 
else ()  | 
| 
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
211  | 
|
| 21637 | 212  | 
|
213  | 
in  | 
|
214  | 
fun setup_thy_loader () = ThyInfo.add_hook trace_action;  | 
|
| 26613 | 215  | 
fun sync_thy_loader () = List.app (trace_action ThyInfo.Update) (ThyInfo.get_names ());  | 
| 21637 | 216  | 
end;  | 
217  | 
||
218  | 
||
| 21949 | 219  | 
(* get informed about files *)  | 
| 21637 | 220  | 
|
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
221  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base;  | 
| 21637 | 222  | 
|
223  | 
val inform_file_retracted = ThyInfo.if_known_thy ThyInfo.remove_thy o thy_name;  | 
|
224  | 
val inform_file_processed = ThyInfo.if_known_thy ThyInfo.touch_child_thys o thy_name;  | 
|
225  | 
||
| 
22042
 
64f8f5433f30
Cleanup message model: add info fatality level.  Add informfileoutdated and some use of completed flag.
 
aspinall 
parents: 
22028 
diff
changeset
 | 
226  | 
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
 | 
227  | 
let val name = thy_name path in  | 
| 21637 | 228  | 
if Toplevel.is_toplevel state andalso ThyInfo.known_thy name then  | 
229  | 
(ThyInfo.touch_child_thys name;  | 
|
| 24079 | 230  | 
ThyInfo.register_thy name handle ERROR msg =>  | 
| 23913 | 231  | 
(warning (cat_lines [msg, "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
 | 
232  | 
tell_file_retracted true (Path.base path)))  | 
| 21637 | 233  | 
else raise Toplevel.UNDEF  | 
234  | 
end;  | 
|
235  | 
||
236  | 
||
237  | 
(* restart top-level loop (keeps most state information) *)  | 
|
238  | 
||
239  | 
val welcome = priority o Session.welcome;  | 
|
240  | 
||
241  | 
fun restart () =  | 
|
242  | 
(sync_thy_loader ();  | 
|
243  | 
tell_clear_goals ();  | 
|
244  | 
tell_clear_response ();  | 
|
| 27578 | 245  | 
Isar.init_point ();  | 
246  | 
welcome ());  | 
|
| 21637 | 247  | 
|
248  | 
||
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
249  | 
(* theorem dependencies *)  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
250  | 
|
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
251  | 
local  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
252  | 
|
| 28809 | 253  | 
fun add_proof_body (PBody {thms, ...}) =
 | 
254  | 
thms |> fold (fn (_, ((name, _, _), _)) => name <> "" ? Symtab.update (name, ()));  | 
|
255  | 
||
256  | 
fun add_thm th =  | 
|
| 28814 | 257  | 
(case Thm.proof_body_of th of  | 
| 28809 | 258  | 
    PBody {proof = PThm (_, ((name, _, _), body)), ...} =>
 | 
259  | 
if Thm.has_name_hint th andalso Thm.get_name_hint th = name  | 
|
260  | 
then add_proof_body (Lazy.force body)  | 
|
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
261  | 
else I  | 
| 28809 | 262  | 
| body => add_proof_body body);  | 
263  | 
||
264  | 
in  | 
|
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
265  | 
|
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
266  | 
fun thms_deps ths =  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
267  | 
let  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
268  | 
(* FIXME proper derivation names!? *)  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
269  | 
val names = map Thm.get_name_hint (filter Thm.has_name_hint ths);  | 
| 28809 | 270  | 
val deps = Symtab.keys (fold add_thm ths Symtab.empty);  | 
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
271  | 
in (names, deps) end;  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
272  | 
|
| 
28100
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
273  | 
fun new_thms_deps thy thy' =  | 
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
274  | 
let  | 
| 
28100
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
275  | 
val prev_facts = PureThy.facts_of thy;  | 
| 
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
276  | 
val facts = PureThy.facts_of thy';  | 
| 
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
277  | 
in thms_deps (maps #2 (Facts.dest_static [prev_facts] facts)) end;  | 
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
278  | 
|
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
279  | 
end;  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
280  | 
|
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
281  | 
|
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
282  | 
(* theorem dependeny output *)  | 
| 22408 | 283  | 
|
284  | 
val show_theorem_dependencies = ref false;  | 
|
285  | 
||
| 21637 | 286  | 
local  | 
287  | 
||
288  | 
val spaces_quote = space_implode " " o map quote;  | 
|
289  | 
||
290  | 
fun thm_deps_message (thms, deps) =  | 
|
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
291  | 
let  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
292  | 
    val valuethms = XML.Elem ("value", [("name", "thms")], [XML.Text thms]);
 | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
293  | 
    val valuedeps = XML.Elem ("value", [("name", "deps")], [XML.Text deps]);
 | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
294  | 
in  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
295  | 
issue_pgip (Metainforesponse  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
296  | 
      {attrs = [("infotype", "isabelle_theorem_dependencies")],
 | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
297  | 
content = [valuethms, valuedeps]})  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
298  | 
end;  | 
| 21637 | 299  | 
|
300  | 
in  | 
|
301  | 
||
| 
28103
 
b79e61861f0f
simplified Toplevel.add_hook: cover successful transactions only;
 
wenzelm 
parents: 
28100 
diff
changeset
 | 
302  | 
fun setup_present_hook () = Toplevel.add_hook (fn _ => fn state => fn state' =>  | 
| 
28100
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
303  | 
if ! show_theorem_dependencies andalso  | 
| 
28103
 
b79e61861f0f
simplified Toplevel.add_hook: cover successful transactions only;
 
wenzelm 
parents: 
28100 
diff
changeset
 | 
304  | 
can Toplevel.theory_of state andalso Toplevel.is_theory state'  | 
| 
28100
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
305  | 
then  | 
| 
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28097 
diff
changeset
 | 
306  | 
let val (names, deps) = new_thms_deps (Toplevel.theory_of state) (Toplevel.theory_of state') in  | 
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
307  | 
if null names orelse null deps then ()  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
308  | 
else thm_deps_message (spaces_quote names, spaces_quote deps)  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
309  | 
end  | 
| 
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
310  | 
else ());  | 
| 21637 | 311  | 
|
312  | 
end;  | 
|
313  | 
||
| 
28097
 
003dff7410c1
added new_thms_deps (operates on global facts, some name_hint approximation);
 
wenzelm 
parents: 
28066 
diff
changeset
 | 
314  | 
|
| 21637 | 315  | 
(** lexicalstructure element with keywords (PGIP version of elisp keywords file) **)  | 
316  | 
||
| 21940 | 317  | 
fun lexicalstructure_keywords () =  | 
| 
27353
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
27177 
diff
changeset
 | 
318  | 
let val keywords = OuterKeyword.dest_keywords ()  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
27177 
diff
changeset
 | 
319  | 
val commands = OuterKeyword.dest_commands ()  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
27177 
diff
changeset
 | 
320  | 
fun keyword_elt kind keyword =  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
27177 
diff
changeset
 | 
321  | 
            XML.Elem("keyword", [("word", keyword), ("category", kind)], [])
 | 
| 21940 | 322  | 
in  | 
| 21637 | 323  | 
(* Also, note we don't call init_outer_syntax here to add interface commands,  | 
324  | 
but they should never appear in scripts anyway so it shouldn't matter *)  | 
|
| 
27353
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
27177 
diff
changeset
 | 
325  | 
Lexicalstructure  | 
| 
 
71c4dd53d4cb
moved global keywords from OuterSyntax to OuterKeyword, tuned interfaces;
 
wenzelm 
parents: 
27177 
diff
changeset
 | 
326  | 
              {content = map (keyword_elt "minor") keywords @ map (keyword_elt "major") commands}
 | 
| 21637 | 327  | 
end  | 
328  | 
||
329  | 
(* TODO: we can issue a lexicalstructure/keyword when the syntax gets extended dynamically;  | 
|
330  | 
hooks needed in outer_syntax.ML to do that. *)  | 
|
331  | 
||
332  | 
||
333  | 
(* Configuration: GUI config, proverinfo messages *)  | 
|
334  | 
||
335  | 
local  | 
|
336  | 
val isabellewww = "http://isabelle.in.tum.de/"  | 
|
337  | 
val staticconfig = "~~/lib/ProofGeneral/pgip_isar.xml"  | 
|
| 21940 | 338  | 
fun orenv v d = case getenv v of "" => d | s => s  | 
| 21637 | 339  | 
fun config_file() = orenv "ISABELLE_PGIPCONFIG" staticconfig  | 
340  | 
fun isabelle_www() = orenv "ISABELLE_HOMEPAGE" isabellewww  | 
|
341  | 
in  | 
|
342  | 
fun send_pgip_config () =  | 
|
343  | 
let  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21655 
diff
changeset
 | 
344  | 
val path = Path.explode (config_file())  | 
| 21940 | 345  | 
val ex = File.exists path  | 
| 21637 | 346  | 
|
| 21940 | 347  | 
val wwwpage =  | 
348  | 
(Url.explode (isabelle_www()))  | 
|
| 21969 | 349  | 
handle ERROR _ =>  | 
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
350  | 
                   (panic ("Error in URL in environment variable ISABELLE_HOMEPAGE.");
 | 
| 21940 | 351  | 
Url.explode isabellewww)  | 
352  | 
||
353  | 
val proverinfo =  | 
|
| 21637 | 354  | 
            Proverinfo { name = "Isabelle",
 | 
| 
26109
 
c69c3559355b
more elaborate structure Distribution (filled-in by makedist);
 
wenzelm 
parents: 
25847 
diff
changeset
 | 
355  | 
version = Distribution.version,  | 
| 21940 | 356  | 
instance = Session.name(),  | 
357  | 
descr = "The Isabelle/Isar theorem prover",  | 
|
358  | 
url = wwwpage,  | 
|
359  | 
filenameextns = ".thy;" }  | 
|
| 21637 | 360  | 
in  | 
| 21940 | 361  | 
if ex then  | 
362  | 
(issue_pgip proverinfo;  | 
|
363  | 
issue_pgip_rawtext (File.read path);  | 
|
364  | 
issue_pgip (lexicalstructure_keywords()))  | 
|
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
365  | 
        else panic ("PGIP configuration file \"" ^ config_file() ^ "\" not found")
 | 
| 21637 | 366  | 
end;  | 
367  | 
end  | 
|
368  | 
||
369  | 
||
| 
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
 | 
370  | 
(* 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
 | 
371  | 
|
| 28588 | 372  | 
val preferences = ref Preferences.pure_preferences;  | 
| 
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
 | 
373  | 
|
| 
28066
 
611e504c1191
extended interface to preferences to allow adding new ones
 
nipkow 
parents: 
28037 
diff
changeset
 | 
374  | 
fun add_preference cat pref =  | 
| 28588 | 375  | 
CRITICAL (fn () => change preferences (Preferences.add cat pref));  | 
| 
28066
 
611e504c1191
extended interface to preferences to allow adding new ones
 
nipkow 
parents: 
28037 
diff
changeset
 | 
376  | 
|
| 28588 | 377  | 
fun setup_preferences_tweak () =  | 
378  | 
CRITICAL (fn () => change preferences  | 
|
379  | 
   (Preferences.set_default ("show-question-marks", "false") #>
 | 
|
380  | 
Preferences.remove "show-question-marks" #> (* we use markup, not ?s *)  | 
|
381  | 
Preferences.remove "theorem-dependencies" #> (* set internally *)  | 
|
382  | 
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
 | 
383  | 
|
| 21637 | 384  | 
|
385  | 
||
386  | 
(* Sending commands to Isar *)  | 
|
387  | 
||
| 26622 | 388  | 
fun isarcmd s = Isar.>>> (OuterSyntax.parse Position.none s);  | 
| 21637 | 389  | 
|
| 21940 | 390  | 
(* TODO:  | 
| 27565 | 391  | 
- apply a command given a transition function;  | 
| 
21885
 
5a11263bd8cf
Remove obsolete prefixes from error and warning messages.
 
aspinall 
parents: 
21867 
diff
changeset
 | 
392  | 
- 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
 | 
393  | 
*)  | 
| 21637 | 394  | 
|
395  | 
(* load an arbitrary file (must be .thy or .ML) *)  | 
|
396  | 
||
397  | 
fun use_thy_or_ml_file file =  | 
|
398  | 
let  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21655 
diff
changeset
 | 
399  | 
val (path,extn) = Path.split_ext (Path.explode file)  | 
| 21637 | 400  | 
in  | 
401  | 
case extn of  | 
|
| 21940 | 402  | 
            "" => isarcmd ("use_thy " ^ quote (Path.implode path))
 | 
403  | 
          | "thy" => isarcmd ("use_thy " ^ quote (Path.implode path))
 | 
|
| 21637 | 404  | 
          | "ML" => isarcmd ("use " ^ quote file)
 | 
| 22028 | 405  | 
          | other => error ("Don't know how to read a file with extension " ^ quote other)
 | 
| 21637 | 406  | 
end  | 
407  | 
||
408  | 
||
| 21867 | 409  | 
(******* PGIP actions *******)  | 
| 21637 | 410  | 
|
411  | 
||
| 21940 | 412  | 
(* Responses to each of the PGIP input commands.  | 
| 21637 | 413  | 
These are programmed uniformly for extensibility. *)  | 
414  | 
||
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
415  | 
fun askpgip (Askpgip _) =  | 
| 
23435
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
416  | 
(issue_pgip  | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
417  | 
         (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
 | 
418  | 
pgipelems = PgipIsabelle.accepted_inputs });  | 
| 
 
061f28854017
Remove dedicated flag setting elements in favour of setproverflag.  Restore displayconfig response in askpgip.
 
aspinall 
parents: 
23226 
diff
changeset
 | 
419  | 
send_pgip_config())  | 
| 21637 | 420  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
421  | 
fun askpgml (Askpgml _) =  | 
| 21637 | 422  | 
issue_pgip  | 
| 21940 | 423  | 
        (Usespgml { version = PgipIsabelle.isabelle_pgml_version_supported })
 | 
| 21637 | 424  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
425  | 
fun askprefs (Askprefs _) =  | 
| 21940 | 426  | 
let  | 
427  | 
        fun preference_of {name, descr, default, pgiptype, get, set } =
 | 
|
428  | 
            { name = name, descr = SOME descr, default = SOME default,
 | 
|
429  | 
pgiptype = pgiptype }  | 
|
| 21637 | 430  | 
in  | 
| 21940 | 431  | 
List.app (fn (prefcat, prefs) =>  | 
432  | 
                     issue_pgip (Hasprefs {prefcategory=SOME prefcat,
 | 
|
433  | 
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
 | 
434  | 
(!preferences)  | 
| 21940 | 435  | 
end  | 
| 21637 | 436  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
437  | 
fun askconfig (Askconfig _) = () (* TODO: add config response *)  | 
| 21637 | 438  | 
|
439  | 
local  | 
|
| 21940 | 440  | 
fun lookuppref pref =  | 
441  | 
case AList.lookup (op =)  | 
|
442  | 
(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
 | 
443  | 
(maps snd (!preferences))) pref of  | 
| 21940 | 444  | 
            NONE => error ("Unknown prover preference: " ^ quote pref)
 | 
445  | 
| SOME p => p  | 
|
| 21637 | 446  | 
in  | 
| 21940 | 447  | 
fun setpref (Setpref vs) =  | 
448  | 
let  | 
|
449  | 
val name = #name vs  | 
|
450  | 
val value = #value vs  | 
|
451  | 
val set = #set (lookuppref name)  | 
|
| 21637 | 452  | 
in  | 
| 21940 | 453  | 
set value  | 
| 21637 | 454  | 
end  | 
455  | 
||
| 21902 | 456  | 
fun getpref (Getpref vs) =  | 
| 21940 | 457  | 
let  | 
458  | 
val name = #name vs  | 
|
459  | 
val get = #get (lookuppref name)  | 
|
460  | 
in  | 
|
| 21637 | 461  | 
        issue_pgip (Prefval {name=name, value=get ()})
 | 
462  | 
end  | 
|
463  | 
end  | 
|
464  | 
||
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
465  | 
fun proverinit _ = restart ()  | 
| 21637 | 466  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
467  | 
fun proverexit _ = isarcmd "quit"  | 
| 21637 | 468  | 
|
| 26706 | 469  | 
fun set_proverflag_quiet b =  | 
| 22408 | 470  | 
isarcmd (if b then "disable_pr" else "enable_pr")  | 
| 21637 | 471  | 
|
| 22408 | 472  | 
fun set_proverflag_pgmlsymbols b =  | 
473  | 
(pgmlsymbols_flag := b;  | 
|
| 24614 | 474  | 
NAMED_CRITICAL "print_mode" (fn () =>  | 
| 26706 | 475  | 
change print_mode  | 
| 22590 | 476  | 
(fn mode =>  | 
| 24614 | 477  | 
remove (op =) Symbol.xsymbolsN mode @ (if b then [Symbol.xsymbolsN] else []))))  | 
| 22408 | 478  | 
|
479  | 
fun set_proverflag_thmdeps b =  | 
|
480  | 
(show_theorem_dependencies := b;  | 
|
| 25223 | 481  | 
Proofterm.proofs := (if b then 1 else 2))  | 
| 21637 | 482  | 
|
| 22408 | 483  | 
fun setproverflag (Setproverflag vs) =  | 
| 26706 | 484  | 
let  | 
| 22590 | 485  | 
val flagname = #flagname vs  | 
486  | 
val value = #value vs  | 
|
| 22408 | 487  | 
in  | 
| 22590 | 488  | 
(case flagname of  | 
489  | 
"quiet" => set_proverflag_quiet value  | 
|
490  | 
| "pgmlsymbols" => set_proverflag_pgmlsymbols value  | 
|
| 26706 | 491  | 
| "metainfo:thmdeps" => set_proverflag_thmdeps value  | 
492  | 
           | _ => log_msg ("Unrecognised prover control flag: " ^
 | 
|
| 23801 | 493  | 
(quote flagname) ^ " ignored."))  | 
| 26706 | 494  | 
end  | 
| 22408 | 495  | 
|
| 21637 | 496  | 
|
| 21940 | 497  | 
fun dostep (Dostep vs) =  | 
498  | 
let  | 
|
499  | 
val text = #text vs  | 
|
500  | 
in  | 
|
501  | 
isarcmd text  | 
|
| 21637 | 502  | 
end  | 
503  | 
||
| 21902 | 504  | 
fun undostep (Undostep vs) =  | 
| 21940 | 505  | 
let  | 
506  | 
val times = #times vs  | 
|
507  | 
in  | 
|
508  | 
        isarcmd ("undos_proof " ^ Int.toString times)
 | 
|
| 21637 | 509  | 
end  | 
510  | 
||
| 27565 | 511  | 
fun redostep _ = sys_error "redo unavailable"  | 
| 21940 | 512  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
513  | 
fun abortgoal _ = isarcmd "kill" (* was: ProofGeneral.kill_proof *)  | 
| 21637 | 514  | 
|
515  | 
||
| 21867 | 516  | 
(*** PGIP identifier tables ***)  | 
517  | 
||
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
518  | 
(* TODO: these ones should be triggered by hooks after a  | 
| 22159 | 519  | 
declaration addition/removal, to be sent automatically. *)  | 
| 21867 | 520  | 
|
| 22159 | 521  | 
fun addids t  = issue_pgip (Addids {idtables = [t]})
 | 
522  | 
fun delids t  = issue_pgip (Delids {idtables = [t]})
 | 
|
| 21867 | 523  | 
|
| 27177 | 524  | 
|
525  | 
local  | 
|
526  | 
||
527  | 
fun theory_facts name =  | 
|
528  | 
let val thy = ThyInfo.get_theory name  | 
|
529  | 
in (map PureThy.facts_of (Theory.parents_of thy), PureThy.facts_of thy) end;  | 
|
530  | 
||
531  | 
fun thms_of_thy name = map fst (theory_facts name |-> Facts.extern_static);  | 
|
532  | 
fun qualified_thms_of_thy name = map fst (theory_facts name |-> Facts.dest_static);  | 
|
533  | 
||
534  | 
in  | 
|
535  | 
||
| 21940 | 536  | 
fun askids (Askids vs) =  | 
| 21637 | 537  | 
let  | 
| 21940 | 538  | 
val url = #url vs (* ask for identifiers within a file *)  | 
539  | 
val thyname = #thyname vs (* ask for identifiers within a theory *)  | 
|
540  | 
val objtype = #objtype vs (* ask for identifiers of a particular type *)  | 
|
| 21867 | 541  | 
|
| 21940 | 542  | 
        fun idtable ty ctx ids = {objtype=ty,context=ctx,ids=ids}
 | 
| 21867 | 543  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
544  | 
        fun setids t = issue_pgip (Setids {idtables = [t]})
 | 
| 22159 | 545  | 
|
| 22225 | 546  | 
(* fake one-level nested "subtheories" by picking apart names. *)  | 
| 
22243
 
0f24888c8591
proper use of NameSpace.is_qualified (avoids compatibility issues of the SML B library);
 
wenzelm 
parents: 
22228 
diff
changeset
 | 
547  | 
val immed_thms_of_thy = filter_out NameSpace.is_qualified o thms_of_thy  | 
| 22225 | 548  | 
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
 | 
549  | 
x::_::_ => SOME x (* String.find? *)  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
550  | 
| _ => NONE  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
551  | 
fun subthys_of_thy s =  | 
| 23178 | 552  | 
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
 | 
553  | 
(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
 | 
554  | 
fun subthms_of_thy thy =  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
555  | 
(case thy_prefix thy of  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
556  | 
NONE => immed_thms_of_thy thy  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
557  | 
| 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
 | 
558  | 
(thms_of_thy prf))  | 
| 26706 | 559  | 
in  | 
| 21940 | 560  | 
case (thyname,objtype) of  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
561  | 
(NONE, NONE) =>  | 
| 26613 | 562  | 
setids (idtable ObjFile NONE (ThyInfo.get_names())) (*FIXME: uris*)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
563  | 
| (NONE, SOME ObjFile) =>  | 
| 26613 | 564  | 
setids (idtable ObjFile NONE (ThyInfo.get_names())) (*FIXME: uris*)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
565  | 
| (SOME fi, SOME ObjFile) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
566  | 
setids (idtable ObjTheory (SOME fi) [fi]) (* TODO: check exists *)  | 
| 22225 | 567  | 
| (NONE, SOME ObjTheory) =>  | 
| 26613 | 568  | 
setids (idtable ObjTheory NONE (ThyInfo.get_names()))  | 
| 22225 | 569  | 
| (SOME thy, SOME ObjTheory) =>  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
570  | 
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
 | 
571  | 
| (SOME thy, SOME ObjTheorem) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
572  | 
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
 | 
573  | 
| (NONE, SOME ObjTheorem) =>  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
574  | 
(* 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
 | 
575  | 
(* Several setids should be allowed, but Eclipse code is currently broken:  | 
| 23226 | 576  | 
List.app (fn thy => setids (idtable ObjTheorem (SOME thy) (subthms_of_thy thy)))  | 
| 26613 | 577  | 
(ThyInfo.get_names()) *)  | 
| 22225 | 578  | 
setids (idtable ObjTheorem NONE (* this one gives ~7000 for HOL *)  | 
| 26613 | 579  | 
(maps qualified_thms_of_thy (ThyInfo.get_names())))  | 
| 22225 | 580  | 
         | _ => warning ("askids: ignored argument combination")
 | 
| 21637 | 581  | 
end  | 
582  | 
||
| 27177 | 583  | 
end;  | 
584  | 
||
| 22159 | 585  | 
fun askrefs (Askrefs vs) =  | 
586  | 
let  | 
|
587  | 
val url = #url vs (* ask for references of a file (i.e. immediate pre-requisites) *)  | 
|
588  | 
val thyname = #thyname vs (* ask for references of a theory (other theories) *)  | 
|
589  | 
val objtype = #objtype vs (* ask for references of a particular type... *)  | 
|
590  | 
val name = #name vs (* ... with this name *)  | 
|
591  | 
||
592  | 
        fun idtable ty ctx ids = {objtype=ty,context=ctx,ids=ids}
 | 
|
593  | 
||
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
594  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base  | 
| 22159 | 595  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
596  | 
fun filerefs f =  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
597  | 
let val thy = thy_name f  | 
| 
24189
 
1fa9852643a3
simplified ThyLoad.deps_thy etc.: discontinued attached ML files;
 
wenzelm 
parents: 
24079 
diff
changeset
 | 
598  | 
val filerefs = #uses (ThyLoad.deps_thy (Path.dir f) thy)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
599  | 
in  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
600  | 
                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
 | 
601  | 
name=NONE, idtables=[], fileurls=filerefs})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
602  | 
end  | 
| 22159 | 603  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
604  | 
fun thyrefs thy =  | 
| 
24189
 
1fa9852643a3
simplified ThyLoad.deps_thy etc.: discontinued attached ML files;
 
wenzelm 
parents: 
24079 
diff
changeset
 | 
605  | 
let val thyrefs = #imports (ThyLoad.deps_thy Path.current thy)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
606  | 
in  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
607  | 
                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
 | 
608  | 
                                     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
 | 
609  | 
ids=thyrefs}], fileurls=[]})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
610  | 
end  | 
| 22159 | 611  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
612  | 
fun thmrefs thmname =  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
613  | 
let  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
614  | 
(* 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
 | 
615  | 
What we want is mapping onto simple PGIP name/context model. *)  | 
| 26603 | 616  | 
val ctx = Toplevel.context_of (Isar.state ()) (* NB: raises UNDEF *)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
617  | 
val thy = Context.theory_of_proof ctx  | 
| 
26343
 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
618  | 
val ths = [PureThy.get_thm thy thmname]  | 
| 28809 | 619  | 
val deps = #2 (thms_deps ths);  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
620  | 
in  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
621  | 
if null deps then ()  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
622  | 
                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
 | 
623  | 
objtype=SOME PgipTypes.ObjTheorem,  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
624  | 
                                          idtables=[{context=NONE, objtype=PgipTypes.ObjTheorem,
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
625  | 
ids=deps}], fileurls=[]})  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
626  | 
end  | 
| 22159 | 627  | 
in  | 
628  | 
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
 | 
629  | 
(SOME file, NONE, _, _) => filerefs file  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
630  | 
| (_,SOME thy,_,_) => thyrefs thy  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
631  | 
| (_,_,SOME PgipTypes.ObjTheorem,SOME thmname) => thmrefs thmname  | 
| 22159 | 632  | 
          | _  => error ("Unimplemented/invalid case of <askrefs>")
 | 
633  | 
end  | 
|
634  | 
||
635  | 
||
636  | 
||
| 21940 | 637  | 
fun showid (Showid vs) =  | 
| 21637 | 638  | 
let  | 
| 21940 | 639  | 
val thyname = #thyname vs  | 
640  | 
val objtype = #objtype vs  | 
|
641  | 
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
 | 
642  | 
|
| 26603 | 643  | 
val topthy = Toplevel.theory_of o Isar.state  | 
| 21637 | 644  | 
|
| 22590 | 645  | 
fun splitthy id =  | 
646  | 
let val comps = NameSpace.explode id  | 
|
647  | 
in case comps of  | 
|
648  | 
(thy::(rest as _::_)) => (ThyInfo.get_theory thy, space_implode "." rest)  | 
|
649  | 
| [plainid] => (topthy(),plainid)  | 
|
650  | 
| _ => raise Toplevel.UNDEF (* assert false *)  | 
|
| 26706 | 651  | 
end  | 
652  | 
||
| 21637 | 653  | 
|
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
654  | 
fun idvalue strings =  | 
| 26706 | 655  | 
            issue_pgip (Idvalue { thyname=thyname, objtype=objtype, name=name,
 | 
| 
23759
 
90bccef65004
Fix nested PGIP messages.  Update for schema simplifications.
 
aspinall 
parents: 
23723 
diff
changeset
 | 
656  | 
                                  text=[XML.Elem("pgml",[],
 | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
657  | 
maps YXML.parse_body 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
 | 
658  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
659  | 
fun string_of_thm th = Pretty.string_of  | 
| 22590 | 660  | 
(Display.pretty_thm_aux  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26723 
diff
changeset
 | 
661  | 
(Syntax.pp_global (Thm.theory_of_thm th))  | 
| 22590 | 662  | 
false (* quote *)  | 
663  | 
false (* show hyps *)  | 
|
664  | 
[] (* asms *)  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
665  | 
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
 | 
666  | 
|
| 
26343
 
0dd2eab7b296
simplified get_thm(s): back to plain name argument;
 
wenzelm 
parents: 
26336 
diff
changeset
 | 
667  | 
fun strings_of_thm (thy, name) = map string_of_thm (PureThy.get_thms thy name)  | 
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
668  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
669  | 
val string_of_thy = Pretty.string_of o ProofDisplay.pretty_full_theory false  | 
| 21940 | 670  | 
in  | 
671  | 
case (thyname, objtype) of  | 
|
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
672  | 
(_, ObjTheory) => idvalue [string_of_thy (ThyInfo.get_theory name)]  | 
| 
22337
 
d4599c206446
Fix <pgipquit>.  Remove unused pgmlatomsN.  Make showid match askids.  Use string functions for showid, avoiding printwrap kludge.
 
aspinall 
parents: 
22249 
diff
changeset
 | 
673  | 
| (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
 | 
674  | 
| (NONE, ObjTheorem) => idvalue (strings_of_thm (splitthy name))  | 
| 21940 | 675  | 
          | (_, ot) => error ("Cannot show objects of type "^(PgipTypes.name_of_objtype ot))
 | 
| 21637 | 676  | 
end  | 
677  | 
||
| 21867 | 678  | 
(*** Inspecting state ***)  | 
679  | 
||
| 21940 | 680  | 
(* The file which is currently being processed interactively.  | 
| 21637 | 681  | 
In the pre-PGIP code, this was informed to Isabelle and the theory loader  | 
682  | 
on completion, but that allows for circularity in case we read  | 
|
683  | 
ourselves. So PGIP opens the filename at the start of a script.  | 
|
684  | 
We ought to prevent problems by modifying the theory loader to know  | 
|
| 21940 | 685  | 
about this special status, but for now we just keep a local reference.  | 
686  | 
*)  | 
|
| 21637 | 687  | 
|
688  | 
val currently_open_file = ref (NONE : pgipurl option)  | 
|
689  | 
||
| 22163 | 690  | 
fun get_currently_open_file () = ! currently_open_file;  | 
691  | 
||
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
692  | 
fun askguise _ =  | 
| 21637 | 693  | 
(* The "guise" is the PGIP abstraction of the prover's state.  | 
694  | 
The <informguise> message is merely used for consistency checking. *)  | 
|
| 21940 | 695  | 
let  | 
696  | 
val openfile = !currently_open_file  | 
|
| 21637 | 697  | 
|
| 26603 | 698  | 
val topthy = Toplevel.theory_of o Isar.state  | 
| 21940 | 699  | 
val topthy_name = Context.theory_name o topthy  | 
| 21637 | 700  | 
|
| 21940 | 701  | 
val opentheory = SOME (topthy_name()) handle Toplevel.UNDEF => NONE  | 
| 21637 | 702  | 
|
| 21940 | 703  | 
fun topproofpos () = try Toplevel.proof_position_of (Isar.state ());  | 
704  | 
val openproofpos = topproofpos()  | 
|
| 21637 | 705  | 
in  | 
706  | 
        issue_pgip (Informguise { file = openfile,
 | 
|
| 21940 | 707  | 
theory = opentheory,  | 
708  | 
(* would be nice to get thm name... *)  | 
|
709  | 
theorem = NONE,  | 
|
710  | 
proofpos = openproofpos })  | 
|
| 21637 | 711  | 
end  | 
712  | 
||
| 21902 | 713  | 
fun parsescript (Parsescript vs) =  | 
| 21637 | 714  | 
let  | 
| 21940 | 715  | 
val text = #text vs  | 
716  | 
val systemdata = #systemdata vs  | 
|
717  | 
val location = #location vs (* TODO: extract position *)  | 
|
| 21637 | 718  | 
|
| 28020 | 719  | 
val doc = PgipParser.pgip_parser Position.none text  | 
| 21637 | 720  | 
|
| 21940 | 721  | 
val sysattrs = PgipTypes.opt_attr "systemdata" systemdata  | 
722  | 
val locattrs = PgipTypes.attrs_of_location location  | 
|
| 21637 | 723  | 
in  | 
724  | 
        issue_pgip (Parseresult { attrs= sysattrs@locattrs,
 | 
|
| 21940 | 725  | 
doc = doc,  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
726  | 
errs = [] })  | 
| 21637 | 727  | 
end  | 
728  | 
||
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
729  | 
fun showproofstate _ = isarcmd "pr"  | 
| 21637 | 730  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
731  | 
fun showctxt _ = isarcmd "print_context"  | 
| 21637 | 732  | 
|
| 21902 | 733  | 
fun searchtheorems (Searchtheorems vs) =  | 
| 21940 | 734  | 
let  | 
735  | 
val arg = #arg vs  | 
|
| 21637 | 736  | 
in  | 
| 21969 | 737  | 
        isarcmd ("find_theorems " ^ arg)
 | 
| 21637 | 738  | 
end  | 
739  | 
||
| 21940 | 740  | 
fun setlinewidth (Setlinewidth vs) =  | 
741  | 
let  | 
|
742  | 
val width = #width vs  | 
|
| 21637 | 743  | 
in  | 
| 21940 | 744  | 
        isarcmd ("pretty_setmargin " ^ Int.toString width) (* FIXME: conversion back/forth! *)
 | 
| 21637 | 745  | 
end  | 
746  | 
||
| 21902 | 747  | 
fun viewdoc (Viewdoc vs) =  | 
| 21940 | 748  | 
let  | 
749  | 
val arg = #arg vs  | 
|
750  | 
in  | 
|
| 
28504
 
7ad7d7d6df47
simplified main Isabelle executables: removed Isabelle and isabelle (replaced by isabelle-process), renamed isatool to isabelle;
 
wenzelm 
parents: 
28443 
diff
changeset
 | 
751  | 
        isarcmd ("print_" ^ arg)   (* FIXME: isabelle doc?.  Return URLs, maybe? *)
 | 
| 21637 | 752  | 
end  | 
753  | 
||
| 21867 | 754  | 
(*** Theory ***)  | 
755  | 
||
| 21902 | 756  | 
fun doitem (Doitem vs) =  | 
| 21637 | 757  | 
let  | 
| 21940 | 758  | 
val text = #text vs  | 
| 21637 | 759  | 
in  | 
| 21940 | 760  | 
isarcmd text  | 
| 21637 | 761  | 
end  | 
762  | 
||
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
763  | 
fun undoitem _ =  | 
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
764  | 
isarcmd "undo"  | 
| 21637 | 765  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
766  | 
fun redoitem _ =  | 
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
767  | 
isarcmd "redo"  | 
| 21637 | 768  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
769  | 
fun aborttheory _ =  | 
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
770  | 
isarcmd "kill" (* was: "init_toplevel" *)  | 
| 21637 | 771  | 
|
| 21902 | 772  | 
fun retracttheory (Retracttheory vs) =  | 
| 21940 | 773  | 
let  | 
774  | 
val thyname = #thyname vs  | 
|
| 21637 | 775  | 
in  | 
| 21940 | 776  | 
        isarcmd ("kill_thy " ^ quote thyname)
 | 
| 21637 | 777  | 
end  | 
778  | 
||
| 21867 | 779  | 
|
780  | 
(*** Files ***)  | 
|
781  | 
||
782  | 
(* Path management: we allow theory files to have dependencies in  | 
|
783  | 
their own directory, but when we change directory for a new file we  | 
|
784  | 
remove the path. Leaving it there can cause confusion with  | 
|
785  | 
difference in batch mode.  | 
|
| 21940 | 786  | 
NB: PGIP does not assume that the prover has a load path.  | 
| 21867 | 787  | 
*)  | 
788  | 
||
789  | 
local  | 
|
790  | 
val current_working_dir = ref (NONE : string option)  | 
|
791  | 
in  | 
|
| 21940 | 792  | 
fun changecwd_dir newdirpath =  | 
793  | 
let  | 
|
| 21867 | 794  | 
val newdir = File.platform_path newdirpath  | 
| 21940 | 795  | 
in  | 
| 21867 | 796  | 
(case (!current_working_dir) of  | 
797  | 
NONE => ()  | 
|
798  | 
| SOME dir => ThyLoad.del_path dir;  | 
|
799  | 
ThyLoad.add_path newdir;  | 
|
800  | 
current_working_dir := SOME newdir)  | 
|
801  | 
end  | 
|
802  | 
end  | 
|
803  | 
||
| 21940 | 804  | 
fun changecwd (Changecwd vs) =  | 
805  | 
let  | 
|
806  | 
val url = #url vs  | 
|
807  | 
val newdir = PgipTypes.path_of_pgipurl url  | 
|
| 21867 | 808  | 
in  | 
| 21940 | 809  | 
changecwd_dir url  | 
| 21867 | 810  | 
end  | 
811  | 
||
| 21902 | 812  | 
fun openfile (Openfile vs) =  | 
| 21940 | 813  | 
let  | 
| 21867 | 814  | 
val url = #url vs  | 
815  | 
val filepath = PgipTypes.path_of_pgipurl url  | 
|
816  | 
val filedir = Path.dir filepath  | 
|
817  | 
val thy_name = Path.implode o #1 o Path.split_ext o Path.base  | 
|
818  | 
val openfile_retract = Output.no_warnings (ThyInfo.if_known_thy ThyInfo.remove_thy) o thy_name;  | 
|
819  | 
in  | 
|
820  | 
case !currently_open_file of  | 
|
| 22028 | 821  | 
          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
 | 
822  | 
PgipTypes.string_of_pgipurl url)  | 
| 21867 | 823  | 
| NONE => (openfile_retract filepath;  | 
| 21940 | 824  | 
changecwd_dir filedir;  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
825  | 
                   priority ("Working in file: " ^ PgipTypes.string_of_pgipurl url);
 | 
| 21940 | 826  | 
currently_open_file := SOME url)  | 
| 21867 | 827  | 
end  | 
828  | 
||
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
829  | 
fun closefile _ =  | 
| 21867 | 830  | 
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
 | 
831  | 
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
 | 
832  | 
                   priority ("Finished working in file: " ^ PgipTypes.string_of_pgipurl f);
 | 
| 21867 | 833  | 
currently_open_file := NONE)  | 
834  | 
      | NONE => raise PGIP ("<closefile> when no file is open!")
 | 
|
835  | 
||
| 21940 | 836  | 
fun loadfile (Loadfile vs) =  | 
837  | 
let  | 
|
838  | 
val url = #url vs  | 
|
839  | 
in  | 
|
| 
22171
 
58f554f51f0d
Let <loadfile> execute even while a file is open interactively.
 
aspinall 
parents: 
22163 
diff
changeset
 | 
840  | 
(* 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
 | 
841  | 
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
 | 
842  | 
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
 | 
843  | 
(* case !currently_open_file of  | 
| 22028 | 844  | 
            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
 | 
845  | 
PgipTypes.string_of_pgipurl url)  | 
| 
22171
 
58f554f51f0d
Let <loadfile> execute even while a file is open interactively.
 
aspinall 
parents: 
22163 
diff
changeset
 | 
846  | 
| NONE => *)  | 
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
847  | 
use_thy_or_ml_file (File.platform_path url)  | 
| 21637 | 848  | 
end  | 
849  | 
||
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
850  | 
fun abortfile _ =  | 
| 21637 | 851  | 
case !currently_open_file of  | 
852  | 
SOME f => (isarcmd "init_toplevel";  | 
|
| 
22228
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
853  | 
                   priority ("Aborted working in file: " ^
 | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
854  | 
PgipTypes.string_of_pgipurl f);  | 
| 21940 | 855  | 
currently_open_file := NONE)  | 
| 21637 | 856  | 
      | NONE => raise PGIP ("<abortfile> when no file is open!")
 | 
857  | 
||
| 21940 | 858  | 
fun retractfile (Retractfile vs) =  | 
859  | 
let  | 
|
860  | 
val url = #url vs  | 
|
| 21637 | 861  | 
in  | 
| 21940 | 862  | 
case !currently_open_file of  | 
| 21637 | 863  | 
            SOME f => raise PGIP ("<retractfile> when a file is open!")
 | 
| 22028 | 864  | 
          | 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
 | 
865  | 
(* 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
 | 
866  | 
let  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
867  | 
val name = thy_name url  | 
| 
 
7c27195a4afc
proper use of PureThy.has_name_hint instead of hard-wired string'';
 
wenzelm 
parents: 
22225 
diff
changeset
 | 
868  | 
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
 | 
869  | 
inform_file_retracted url)  | 
| 21637 | 870  | 
end  | 
871  | 
||
872  | 
||
| 21867 | 873  | 
(*** System ***)  | 
| 21637 | 874  | 
|
| 21902 | 875  | 
fun systemcmd (Systemcmd vs) =  | 
| 21940 | 876  | 
let  | 
| 21637 | 877  | 
val arg = #arg vs  | 
878  | 
in  | 
|
879  | 
isarcmd arg  | 
|
880  | 
end  | 
|
881  | 
||
882  | 
exception PGIP_QUIT;  | 
|
| 
23834
 
ad6ad61332fa
avoid redundant variables in patterns (which made Alice vomit);
 
wenzelm 
parents: 
23801 
diff
changeset
 | 
883  | 
fun quitpgip _ = raise PGIP_QUIT  | 
| 21637 | 884  | 
|
| 21902 | 885  | 
fun process_input inp = case inp  | 
886  | 
of Pgip.Askpgip _ => askpgip inp  | 
|
887  | 
| Pgip.Askpgml _ => askpgml inp  | 
|
| 21940 | 888  | 
| Pgip.Askprefs _ => askprefs inp  | 
| 21902 | 889  | 
| Pgip.Askconfig _ => askconfig inp  | 
890  | 
| Pgip.Getpref _ => getpref inp  | 
|
891  | 
| Pgip.Setpref _ => setpref inp  | 
|
892  | 
| Pgip.Proverinit _ => proverinit inp  | 
|
893  | 
| Pgip.Proverexit _ => proverexit inp  | 
|
| 22408 | 894  | 
| Pgip.Setproverflag _ => setproverflag inp  | 
| 21902 | 895  | 
| Pgip.Dostep _ => dostep inp  | 
896  | 
| Pgip.Undostep _ => undostep inp  | 
|
897  | 
| 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
 | 
898  | 
| 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
 | 
899  | 
| Pgip.Restoregoal _ => error "<restoregoal> not implemented by Isabelle"  | 
| 21902 | 900  | 
| Pgip.Abortgoal _ => abortgoal inp  | 
901  | 
| Pgip.Askids _ => askids inp  | 
|
| 22159 | 902  | 
| Pgip.Askrefs _ => askrefs inp  | 
| 21902 | 903  | 
| Pgip.Showid _ => showid inp  | 
904  | 
| Pgip.Askguise _ => askguise inp  | 
|
905  | 
| Pgip.Parsescript _ => parsescript inp  | 
|
906  | 
| Pgip.Showproofstate _ => showproofstate inp  | 
|
907  | 
| Pgip.Showctxt _ => showctxt inp  | 
|
908  | 
| Pgip.Searchtheorems _ => searchtheorems inp  | 
|
909  | 
| Pgip.Setlinewidth _ => setlinewidth inp  | 
|
910  | 
| Pgip.Viewdoc _ => viewdoc inp  | 
|
911  | 
| Pgip.Doitem _ => doitem inp  | 
|
912  | 
| Pgip.Undoitem _ => undoitem inp  | 
|
913  | 
| Pgip.Redoitem _ => redoitem inp  | 
|
914  | 
| Pgip.Aborttheory _ => aborttheory inp  | 
|
915  | 
| Pgip.Retracttheory _ => retracttheory inp  | 
|
916  | 
| Pgip.Loadfile _ => loadfile inp  | 
|
917  | 
| Pgip.Openfile _ => openfile inp  | 
|
918  | 
| Pgip.Closefile _ => closefile inp  | 
|
919  | 
| Pgip.Abortfile _ => abortfile inp  | 
|
920  | 
| Pgip.Retractfile _ => retractfile inp  | 
|
921  | 
| Pgip.Changecwd _ => changecwd inp  | 
|
922  | 
| Pgip.Systemcmd _ => systemcmd inp  | 
|
923  | 
| Pgip.Quitpgip _ => quitpgip inp  | 
|
| 21637 | 924  | 
|
925  | 
||
| 21940 | 926  | 
fun process_pgip_element pgipxml =  | 
| 21637 | 927  | 
case pgipxml of  | 
| 21969 | 928  | 
xml as (XML.Elem elem) =>  | 
| 21940 | 929  | 
(case Pgip.input elem of  | 
930  | 
             NONE => warning ("Unrecognized PGIP command, ignored: \n" ^
 | 
|
| 26541 | 931  | 
(XML.string_of xml))  | 
| 21940 | 932  | 
| SOME inp => (process_input inp)) (* errors later; packet discarded *)  | 
| 21969 | 933  | 
| XML.Text t => ignored_text_warning t  | 
| 21637 | 934  | 
and ignored_text_warning t =  | 
| 21940 | 935  | 
if size (Symbol.strip_blanks t) > 0 then  | 
936  | 
           warning ("Ignored text in PGIP packet: \n" ^ t)
 | 
|
| 21637 | 937  | 
else ()  | 
938  | 
||
939  | 
fun process_pgip_tree xml =  | 
|
940  | 
(pgip_refid := NONE;  | 
|
941  | 
pgip_refseq := NONE;  | 
|
942  | 
(case xml of  | 
|
943  | 
          XML.Elem ("pgip", attrs, pgips) =>
 | 
|
944  | 
(let  | 
|
945  | 
val class = PgipTypes.get_attr "class" attrs  | 
|
946  | 
val dest = PgipTypes.get_attr_opt "destid" attrs  | 
|
| 21940 | 947  | 
val seq = PgipTypes.read_pgipnat (PgipTypes.get_attr "seq" attrs)  | 
| 21637 | 948  | 
(* Respond to prover broadcasts, or messages for us. Ignore rest *)  | 
| 21940 | 949  | 
val processit =  | 
950  | 
case dest of  | 
|
| 21637 | 951  | 
NONE => class = "pa"  | 
| 21940 | 952  | 
| SOME id => matching_pgip_id id  | 
953  | 
in if processit then  | 
|
954  | 
(pgip_refid := PgipTypes.get_attr_opt "id" attrs;  | 
|
955  | 
pgip_refseq := SOME seq;  | 
|
956  | 
List.app process_pgip_element pgips;  | 
|
957  | 
(* return true to indicate <ready/> *)  | 
|
958  | 
true)  | 
|
959  | 
else  | 
|
960  | 
(* no response to ignored messages. *)  | 
|
961  | 
false  | 
|
| 21637 | 962  | 
end)  | 
963  | 
| _ => raise PGIP "Invalid PGIP packet received")  | 
|
964  | 
handle PGIP msg =>  | 
|
965  | 
(Output.error_msg ((msg ^ "\nPGIP error occured in XML text below:\n") ^  | 
|
| 26541 | 966  | 
(XML.string_of xml));  | 
| 21940 | 967  | 
true))  | 
| 21637 | 968  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
969  | 
(* External input *)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
970  | 
|
| 26541 | 971  | 
val process_pgip_plain = K () o process_pgip_tree o XML.parse  | 
| 21637 | 972  | 
|
973  | 
(* PGIP loop: process PGIP input only *)  | 
|
974  | 
||
975  | 
local  | 
|
976  | 
||
977  | 
exception XML_PARSE  | 
|
978  | 
||
979  | 
fun loop ready src =  | 
|
980  | 
let  | 
|
981  | 
val _ = if ready then issue_pgip (Ready ()) else ()  | 
|
| 21969 | 982  | 
val pgipo =  | 
983  | 
(case try Source.get_single src of  | 
|
984  | 
SOME pgipo => pgipo  | 
|
985  | 
| NONE => raise XML_PARSE)  | 
|
| 21637 | 986  | 
in  | 
987  | 
case pgipo of  | 
|
988  | 
NONE => ()  | 
|
989  | 
| SOME (pgip,src') =>  | 
|
990  | 
let  | 
|
| 21940 | 991  | 
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
 | 
992  | 
handle PGIP_QUIT => raise PGIP_QUIT  | 
| 22590 | 993  | 
| e => (handler (e,SOME src'); true)  | 
| 21637 | 994  | 
in  | 
995  | 
loop ready' src'  | 
|
996  | 
end  | 
|
997  | 
end handle e => handler (e,SOME src) (* error in XML parse or Ready issue *)  | 
|
998  | 
||
999  | 
and handler (e,srco) =  | 
|
1000  | 
case (e,srco) of  | 
|
1001  | 
(XML_PARSE,SOME src) =>  | 
|
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
1002  | 
panic "Invalid XML input, aborting" (* TODO: attempt recovery *)  | 
| 28443 | 1003  | 
| (Exn.Interrupt,SOME src) =>  | 
| 21940 | 1004  | 
(Output.error_msg "Interrupt during PGIP processing"; loop true src)  | 
1005  | 
| (Toplevel.UNDEF,SOME src) =>  | 
|
1006  | 
(Output.error_msg "No working context defined"; loop true src)  | 
|
1007  | 
| (e,SOME src) =>  | 
|
1008  | 
(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
 | 
1009  | 
| (PGIP_QUIT,_) => ()  | 
| 21637 | 1010  | 
| (_,NONE) => ()  | 
1011  | 
in  | 
|
1012  | 
(* TODO: add socket interface *)  | 
|
1013  | 
||
| 
26552
 
5677b4faf295
renamed XML.parse_comment_whspc to XML.parse_comments;
 
wenzelm 
parents: 
26548 
diff
changeset
 | 
1014  | 
val xmlP = XML.parse_comments |-- XML.parse_element >> single  | 
| 21637 | 1015  | 
|
1016  | 
val tty_src = Source.set_prompt "" (Source.source Symbol.stopper xmlP NONE Source.tty)  | 
|
1017  | 
||
1018  | 
fun pgip_toplevel x = loop true x  | 
|
1019  | 
end  | 
|
1020  | 
||
1021  | 
||
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1022  | 
(* Extra command for embedding prover-control inside document (obscure/debug usage). *)  | 
| 21637 | 1023  | 
|
| 24867 | 1024  | 
fun init_outer_syntax () =  | 
| 
21972
 
1b68312c4cf0
Initialise parser at startup.  Remove some obsolete ProofGeneral.XXX outer syntax, mapping PGIP commands directly to Isar.
 
aspinall 
parents: 
21970 
diff
changeset
 | 
1025  | 
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
 | 
1026  | 
(OuterParse.text >> (Toplevel.no_timing oo  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1027  | 
(fn txt => Toplevel.imperative (fn () => process_pgip_plain txt))));  | 
| 21637 | 1028  | 
|
1029  | 
||
1030  | 
(* init *)  | 
|
1031  | 
||
1032  | 
val initialized = ref false;  | 
|
1033  | 
||
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
1034  | 
fun init_pgip false = panic "No Proof General interface support for Isabelle/classic mode."  | 
| 21969 | 1035  | 
| init_pgip true =  | 
1036  | 
(! initialized orelse  | 
|
| 
25445
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1037  | 
(setup_preferences_tweak ();  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
1038  | 
Output.add_mode proof_generalN Output.default_output Output.default_escape;  | 
| 
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
1039  | 
Markup.add_mode proof_generalN YXML.output_markup;  | 
| 
25445
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1040  | 
setup_messages ();  | 
| 
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1041  | 
Output.no_warnings init_outer_syntax ();  | 
| 
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1042  | 
setup_thy_loader ();  | 
| 
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1043  | 
setup_present_hook ();  | 
| 
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1044  | 
init_pgip_session_id ();  | 
| 
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1045  | 
welcome ();  | 
| 
 
01f3686f4304
Init outer syntax after message setup to avoid spurious output.
 
aspinall 
parents: 
25275 
diff
changeset
 | 
1046  | 
set initialized);  | 
| 21969 | 1047  | 
sync_thy_loader ();  | 
| 
28037
 
915b9a777441
changed Markup print mode to YXML -- explicitly decode messages before being issued;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
1048  | 
change print_mode (update (op =) proof_generalN);  | 
| 21969 | 1049  | 
pgip_toplevel tty_src);  | 
| 21637 | 1050  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1051  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1052  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1053  | 
(** Out-of-loop PGIP commands (for Emacs hybrid mode) **)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1054  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1055  | 
local  | 
| 22590 | 1056  | 
val pgip_output_channel = ref Output.writeln_default  | 
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1057  | 
in  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1058  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1059  | 
(* Set recipient for PGIP results *)  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1060  | 
fun init_pgip_channel writefn =  | 
| 21940 | 1061  | 
(init_pgip_session_id();  | 
1062  | 
pgip_output_channel := writefn)  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1063  | 
|
| 21940 | 1064  | 
(* Process a PGIP command.  | 
1065  | 
This works for preferences but not generally guaranteed  | 
|
| 
21649
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1066  | 
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
 | 
1067  | 
fun process_pgip str =  | 
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1068  | 
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
 | 
1069  | 
|
| 
 
40e6fdd26f82
Support PGIP communication for preferences in Emacs mode.
 
aspinall 
parents: 
21646 
diff
changeset
 | 
1070  | 
end  | 
| 21637 | 1071  | 
|
1072  | 
end;  |