author | berghofe |
Wed, 27 Nov 2002 17:16:47 +0100 | |
changeset 13728 | 8004e56204fd |
parent 13646 | 46ed3d042ba5 |
child 13801 | 6c5c5bdfae84 |
permissions | -rw-r--r-- |
12780 | 1 |
(* Title: Pure/proof_general.ML |
12778 | 2 |
ID: $Id$ |
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
|
5 |
||
6 |
Isabelle configuration for Proof General (see http://www.proofgeneral.org). |
|
7 |
*) |
|
8 |
||
9 |
signature PROOF_GENERAL = |
|
10 |
sig |
|
11 |
val setup: (theory -> theory) list |
|
12 |
val update_thy_only: string -> unit |
|
13 |
val try_update_thy_only: string -> unit |
|
13391 | 14 |
val inform_file_retracted: string -> unit |
12778 | 15 |
val inform_file_processed: string -> unit |
13728
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
16 |
val options: (string * (string * (string * (unit -> string) * (string -> unit)))) list ref |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
17 |
val process_pgip: string -> unit |
12778 | 18 |
val thms_containing: string list -> unit |
13646 | 19 |
val print_intros: unit -> unit |
12778 | 20 |
val help: unit -> unit |
21 |
val show_context: unit -> theory |
|
22 |
val kill_goal: unit -> unit |
|
23 |
val repeat_undo: int -> unit |
|
24 |
val isa_restart: unit -> unit |
|
12833 | 25 |
val full_proofs: bool -> unit |
12778 | 26 |
val init: bool -> unit |
27 |
val write_keywords: string -> unit |
|
28 |
end; |
|
29 |
||
30 |
structure ProofGeneral: PROOF_GENERAL = |
|
31 |
struct |
|
32 |
||
33 |
(* print modes *) |
|
34 |
||
35 |
val proof_generalN = "ProofGeneral"; |
|
36 |
val xsymbolsN = "xsymbols"; |
|
13538 | 37 |
val thm_depsN = "thm_deps"; |
12778 | 38 |
|
13728
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
39 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
40 |
val pgml_version_supported = "1.0"; |
12778 | 41 |
val pgmlN = "PGML"; |
42 |
fun pgml () = pgmlN mem_string ! print_mode; |
|
43 |
||
44 |
||
45 |
(* text output *) |
|
46 |
||
47 |
local |
|
48 |
||
49 |
fun xsymbols_output s = |
|
50 |
if xsymbolsN mem_string ! print_mode andalso exists_string (equal "\\") s then |
|
51 |
let val syms = Symbol.explode s |
|
52 |
in (implode (map (fn "\\" => "\\\\" | c => c) syms), real (Symbol.length syms)) end |
|
53 |
else (s, real (size s)); |
|
54 |
||
55 |
fun pgml_output (s, len) = |
|
56 |
if pgml () then (XML.text s, len) |
|
57 |
else (s, len); |
|
58 |
||
59 |
in |
|
60 |
||
61 |
fun setup_xsymbols_output () = |
|
62 |
Symbol.add_mode proof_generalN (pgml_output o xsymbols_output, Symbol.default_indent); |
|
63 |
||
64 |
end; |
|
65 |
||
66 |
||
67 |
(* token translations *) |
|
68 |
||
69 |
local |
|
70 |
||
71 |
val end_tag = oct_char "350"; |
|
72 |
val class_tag = ("class", oct_char "351"); |
|
73 |
val tfree_tag = ("tfree", oct_char "352"); |
|
74 |
val tvar_tag = ("tvar", oct_char "353"); |
|
75 |
val free_tag = ("free", oct_char "354"); |
|
76 |
val bound_tag = ("bound", oct_char "355"); |
|
77 |
val var_tag = ("var", oct_char "356"); |
|
78 |
val skolem_tag = ("skolem", oct_char "357"); |
|
79 |
||
80 |
fun xml_atom kind x = XML.element "atom" [("kind", kind)] [XML.text x]; |
|
81 |
||
82 |
fun tagit (kind, bg_tag) x = |
|
83 |
(if pgml () then xml_atom kind x else bg_tag ^ x ^ end_tag, |
|
84 |
real (Symbol.length (Symbol.explode x))); |
|
85 |
||
86 |
fun free_or_skolem x = |
|
87 |
(case try Syntax.dest_skolem x of |
|
88 |
None => tagit free_tag x |
|
89 |
| Some x' => tagit skolem_tag x'); |
|
90 |
||
91 |
fun var_or_skolem s = |
|
92 |
(case Syntax.read_var s of |
|
93 |
Var ((x, i), _) => |
|
94 |
(case try Syntax.dest_skolem x of |
|
95 |
None => tagit var_tag s |
|
96 |
| Some x' => tagit skolem_tag (Syntax.string_of_vname (x', i))) |
|
97 |
| _ => tagit var_tag s); |
|
98 |
||
99 |
val proof_general_trans = |
|
100 |
Syntax.tokentrans_mode proof_generalN |
|
101 |
[("class", tagit class_tag), |
|
102 |
("tfree", tagit tfree_tag), |
|
103 |
("tvar", tagit tvar_tag), |
|
104 |
("free", free_or_skolem), |
|
105 |
("bound", tagit bound_tag), |
|
106 |
("var", var_or_skolem)]; |
|
107 |
||
108 |
in val setup = [Theory.add_tokentrfuns proof_general_trans] end; |
|
109 |
||
110 |
||
111 |
||
112 |
(* messages and notification *) |
|
113 |
||
114 |
local |
|
115 |
||
116 |
fun decorated_output bg en prfx = |
|
117 |
writeln_default o enclose bg en o prefix_lines prfx; |
|
118 |
||
119 |
fun message kind bg en prfx s = |
|
120 |
if pgml () then writeln_default (XML.element kind [] [prefix_lines prfx s]) |
|
121 |
else decorated_output bg en prfx s; |
|
122 |
||
13526 | 123 |
in |
12778 | 124 |
|
13526 | 125 |
val notify = message "notify" (oct_char "360") (oct_char "361") ""; |
12778 | 126 |
|
127 |
fun setup_messages () = |
|
128 |
(writeln_fn := message "output" "" "" ""; |
|
129 |
priority_fn := message "information" (oct_char "360") (oct_char "361") ""; |
|
130 |
tracing_fn := message "tracing" (oct_char "360" ^ oct_char "375") (oct_char "361") ""; |
|
131 |
warning_fn := message "warning" (oct_char "362") (oct_char "363") "### "; |
|
132 |
error_fn := message "error" (oct_char "364") (oct_char "365") "*** "); |
|
133 |
||
134 |
fun tell_clear_goals () = notify "Proof General, please clear the goals buffer."; |
|
135 |
fun tell_clear_response () = notify "Proof General, please clear the response buffer."; |
|
136 |
fun tell_file msg path = notify ("Proof General, " ^ msg ^ " " ^ quote (File.sysify_path path)); |
|
137 |
||
138 |
end; |
|
139 |
||
140 |
||
141 |
(* theory / proof state output *) |
|
142 |
||
143 |
local |
|
144 |
||
145 |
fun tmp_markers f = |
|
146 |
setmp Display.current_goals_markers (oct_char "366", oct_char "367", "") f (); |
|
147 |
||
148 |
fun statedisplay prts = |
|
149 |
writeln_default (XML.element "statedisplay" [] [Pretty.string_of (Pretty.chunks prts)]); |
|
150 |
||
151 |
fun print_current_goals n m st = |
|
152 |
if pgml () then statedisplay (Display.pretty_current_goals n m st) |
|
153 |
else tmp_markers (fn () => Display.print_current_goals_default n m st); |
|
154 |
||
155 |
fun print_state b st = |
|
156 |
if pgml () then statedisplay (Toplevel.pretty_state b st) |
|
157 |
else tmp_markers (fn () => Toplevel.print_state_default b st); |
|
158 |
||
159 |
in |
|
160 |
||
161 |
fun setup_state () = |
|
162 |
(Display.print_current_goals_fn := print_current_goals; |
|
163 |
Toplevel.print_state_fn := print_state; |
|
164 |
Toplevel.prompt_state_fn := (suffix (oct_char "372") o Toplevel.prompt_state_default)); |
|
165 |
||
166 |
end; |
|
167 |
||
168 |
||
13538 | 169 |
(* theorem dependency output *) |
13526 | 170 |
|
171 |
local |
|
172 |
||
13545 | 173 |
val spaces_quote = space_implode " " o map quote; |
174 |
||
13526 | 175 |
fun tell_thm_deps ths = |
13538 | 176 |
conditional (thm_depsN mem_string ! print_mode) (fn () => |
177 |
let |
|
13545 | 178 |
val names = map Thm.name_of_thm ths \ ""; |
179 |
val deps = Symtab.keys (foldl (uncurry Proofterm.thms_of_proof) |
|
180 |
(Symtab.empty, map Thm.proof_of ths)) \ ""; |
|
13538 | 181 |
in |
13545 | 182 |
if null names orelse null deps then () |
183 |
else notify ("Proof General, theorem dependencies of " ^ spaces_quote names ^ " are " |
|
184 |
^ spaces_quote deps) |
|
13538 | 185 |
end); |
13526 | 186 |
|
187 |
in |
|
188 |
||
189 |
fun setup_present_hook () = |
|
190 |
Present.add_hook (fn _ => fn res => tell_thm_deps (flat (map #2 res))); |
|
191 |
||
192 |
end; |
|
193 |
||
194 |
||
12778 | 195 |
(* theory loader actions *) |
196 |
||
197 |
local |
|
198 |
||
199 |
fun add_master_files name files = |
|
200 |
let val masters = [ThyLoad.thy_path name, ThyLoad.ml_path name] |
|
201 |
in masters @ gen_rems (op = o pairself Path.base) (files, masters) end; |
|
202 |
||
203 |
fun trace_action action name = |
|
204 |
if action = ThyInfo.Update then |
|
205 |
seq (tell_file "this file is loaded:") (ThyInfo.loaded_files name) |
|
206 |
else if action = ThyInfo.Outdate orelse action = ThyInfo.Remove then |
|
207 |
seq (tell_file "you can unlock the file") (add_master_files name (ThyInfo.loaded_files name)) |
|
208 |
else (); |
|
209 |
||
210 |
in |
|
211 |
fun setup_thy_loader () = ThyInfo.add_hook trace_action; |
|
212 |
fun sync_thy_loader () = seq (trace_action ThyInfo.Update) (ThyInfo.names ()); |
|
213 |
end; |
|
214 |
||
215 |
||
216 |
(* prepare theory context *) |
|
217 |
||
218 |
val thy_name = Path.pack o Path.drop_ext o Path.base o Path.unpack; |
|
219 |
val update_thy_only = setmp MetaSimplifier.trace_simp false ThyInfo.update_thy_only; |
|
220 |
||
221 |
fun which_context () = |
|
222 |
(case Context.get_context () of |
|
223 |
Some thy => " Using current (dynamic!) one: " ^ |
|
224 |
(case try PureThy.get_name thy of Some name => quote name | None => "<unnamed>") |
|
225 |
| None => ""); |
|
226 |
||
227 |
fun try_update_thy_only file = |
|
228 |
ThyLoad.cond_add_path (Path.dir (Path.unpack file)) (fn () => |
|
229 |
let val name = thy_name file in |
|
230 |
if is_some (ThyLoad.check_file (ThyLoad.thy_path name)) then update_thy_only name |
|
231 |
else warning ("Unkown theory context of ML file." ^ which_context ()) |
|
232 |
end) (); |
|
233 |
||
234 |
||
235 |
(* get informed about files *) |
|
236 |
||
13391 | 237 |
val inform_file_retracted = ThyInfo.if_known_thy ThyInfo.remove_thy o thy_name; |
238 |
val inform_file_processed = ThyInfo.if_known_thy ThyInfo.touch_child_thys o thy_name; |
|
12778 | 239 |
|
240 |
fun proper_inform_file_processed file state = |
|
241 |
let val name = thy_name file in |
|
13391 | 242 |
ThyInfo.if_known_thy ThyInfo.touch_child_thys name; |
12778 | 243 |
if not (Toplevel.is_toplevel state) then |
244 |
warning ("Not at toplevel -- cannot register theory " ^ quote name) |
|
245 |
else Library.transform_error ThyInfo.pretend_use_thy_only name handle ERROR_MESSAGE msg => |
|
246 |
(warning msg; warning ("Failed to register theory " ^ quote name)) |
|
247 |
end; |
|
248 |
||
249 |
||
13728
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
250 |
(* options *) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
251 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
252 |
fun nat_option r = ("nat", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
253 |
(fn () => string_of_int (!r)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
254 |
(fn s => (case Syntax.read_nat s of |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
255 |
None => error "nat_option: illegal value" |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
256 |
| Some i => r := i))); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
257 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
258 |
fun bool_option r = ("boolean", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
259 |
(fn () => Bool.toString (!r)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
260 |
(fn "false" => r := false | "true" => r := true |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
261 |
| _ => error "bool_option: illegal value")); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
262 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
263 |
val proof_option = ("boolean", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
264 |
(fn () => Bool.toString (!proofs >= 2)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
265 |
(fn "false" => proofs := 1 | "true" => proofs := 2 |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
266 |
| _ => error "proof_option: illegal value")); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
267 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
268 |
val thm_deps_option = ("boolean", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
269 |
(fn () => Bool.toString ("thm_deps" mem !print_mode)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
270 |
(fn "false" => print_mode := Library.gen_rems (op =) (!print_mode, ["thm_deps"]) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
271 |
| "true" => print_mode := (["thm_deps"] @ !print_mode) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
272 |
| _ => error "thm_deps_option: illegal value")); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
273 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
274 |
val print_depth_option = ("nat", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
275 |
(fn () => "10"), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
276 |
(fn s => (case Syntax.read_nat s of |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
277 |
None => error "print_depth_option: illegal value" |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
278 |
| Some i => print_depth i))); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
279 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
280 |
val options = ref |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
281 |
[("show-types", ("Whether to show types in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
282 |
bool_option show_types)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
283 |
("show-sorts", ("Whether to show sorts in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
284 |
bool_option show_sorts)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
285 |
("show-consts", ("Whether to show types of consts in Isabelle goals.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
286 |
bool_option show_consts)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
287 |
("long-names", ("Whether to show fully qualified names in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
288 |
bool_option long_names)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
289 |
("eta-contract", ("Whether to print terms eta-contracted in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
290 |
bool_option Syntax.eta_contract)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
291 |
("trace-simplifier", ("Whether to trace the Simplifier in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
292 |
bool_option trace_simp)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
293 |
("trace-rules", ("Whether to trace the standard rules in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
294 |
bool_option trace_rules)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
295 |
("quick-and-dirty", ("Whether to take a few short cuts occasionally.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
296 |
bool_option quick_and_dirty)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
297 |
("full-proofs", ("Whether to record full proof objects internally.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
298 |
proof_option)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
299 |
("trace-unification", ("Whether to output error diagnostics during unification.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
300 |
bool_option Pattern.trace_unify_fail)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
301 |
("show-main-goal", ("Whether to show main goal.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
302 |
bool_option Proof.show_main_goal)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
303 |
("global-timing", ("Whether to enable timing in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
304 |
bool_option Library.timing)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
305 |
("theorem-dependencies", ("Whether to track theorem dependencies within Proof General.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
306 |
thm_deps_option)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
307 |
("goals-limit", ("Setting for maximum number of goals printed in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
308 |
nat_option goals_limit)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
309 |
("prems-limit", ("Setting for maximum number of premises printed in Isabelle/Isar.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
310 |
nat_option ProofContext.prems_limit)), |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
311 |
("print-depth", ("Setting for the ML print depth in Isabelle.", |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
312 |
print_depth_option))]; |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
313 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
314 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
315 |
(* Sending PGIP commands to the interface *) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
316 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
317 |
fun issue_pgip pgips = notify (XML.element "pgip" [] pgips); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
318 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
319 |
fun usespgml () = |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
320 |
issue_pgip [XML.element "usespgml" [("version", pgml_version_supported)] []]; |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
321 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
322 |
(* NB: the default returned here is actually the current value, so |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
323 |
repeated uses of <askprefs> will not work correctly. *) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
324 |
fun show_options () = issue_pgip (map |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
325 |
(fn (name, (descr, (ty, get, _))) => (XML.element "haspref" |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
326 |
[("type", ty), ("descr", descr), ("default", get ())] [name])) (!options)); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
327 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
328 |
fun set_option name value = (case assoc (!options, name) of |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
329 |
None => warning ("Unknown option: " ^ quote name) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
330 |
| Some (_, (_, _, set)) => set value); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
331 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
332 |
fun get_option name = (case assoc (!options, name) of |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
333 |
None => warning ("Unknown option: " ^ quote name) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
334 |
| Some (_, (_, get, _)) => |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
335 |
issue_pgip [XML.element "prefval" [("name", name)] [get ()]]); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
336 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
337 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
338 |
(* Processing PGIP commands from the interface *) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
339 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
340 |
(* FIXME: matching on attributes is a bit too strict here *) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
341 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
342 |
fun process_pgip_element pgip = (case pgip of |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
343 |
XML.Elem ("askpgml", _, []) => usespgml () |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
344 |
| XML.Elem ("askprefs", _, []) => show_options () |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
345 |
| XML.Elem ("getpref", [("name", name)], []) => get_option name |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
346 |
| XML.Elem ("setpref", [("name", name)], [XML.Text value]) => |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
347 |
set_option name value |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
348 |
| XML.Elem (e, _, _) => warning ("Unrecognized PGIP command: " ^ e) |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
349 |
| XML.Text t => warning ("Unrecognized PGIP command:\n" ^ t)); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
350 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
351 |
fun process_pgip s = (case XML.tree_of_string s of |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
352 |
XML.Elem ("pgip", _, pgips) => seq process_pgip_element pgips |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
353 |
| _ => warning ("Invalid PGIP packet received\n" ^ s)); |
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
354 |
|
8004e56204fd
Added some functions for processing PGIP (thanks to David Aspinall).
berghofe
parents:
13646
diff
changeset
|
355 |
|
12778 | 356 |
(* misc commands for ProofGeneral/isa *) |
357 |
||
358 |
fun thms_containing ss = |
|
13284 | 359 |
ProofContext.print_thms_containing (ProofContext.init (the_context ())) None ss; |
12778 | 360 |
|
13646 | 361 |
fun print_intros() = |
362 |
let val proof_state = Toplevel.proof_of (Toplevel.get_state ()) |
|
363 |
val (ctxt,(_,st)) = Proof.get_goal proof_state |
|
364 |
val prt_fact = ProofContext.pretty_fact ctxt |
|
365 |
val thy = ProofContext.theory_of ctxt |
|
366 |
val facts = map (apsnd single) (Goals.find_intros_goal thy st 1) |
|
367 |
in map prt_fact (rev facts) |> Pretty.chunks |> Pretty.writeln end |
|
368 |
||
12778 | 369 |
val welcome = priority o Session.welcome; |
370 |
val help = welcome; |
|
371 |
val show_context = Context.the_context; |
|
372 |
||
373 |
fun kill_goal () = (Goals.reset_goals (); tell_clear_goals ()); |
|
374 |
||
375 |
fun no_print_goals f = setmp Display.print_current_goals_fn (fn _ => fn _ => fn _ => ()) f; |
|
376 |
||
377 |
fun repeat_undo 0 = () |
|
378 |
| repeat_undo 1 = undo () |
|
379 |
| repeat_undo n = (no_print_goals undo (); repeat_undo (n - 1)); |
|
380 |
||
381 |
||
382 |
(* restart top-level loop (keeps most state information) *) |
|
383 |
||
384 |
local |
|
385 |
||
386 |
fun restart isar = |
|
387 |
(if isar then tell_clear_goals () else kill_goal (); |
|
388 |
tell_clear_response (); |
|
389 |
welcome ()); |
|
390 |
||
391 |
in |
|
392 |
||
393 |
fun isa_restart () = restart false; |
|
394 |
fun isar_restart () = (sync_thy_loader (); restart true; raise Toplevel.RESTART); |
|
395 |
||
396 |
end; |
|
397 |
||
398 |
||
12833 | 399 |
fun full_proofs true = proofs := 2 |
400 |
| full_proofs false = proofs := 1; |
|
401 |
||
402 |
||
12778 | 403 |
(* outer syntax *) |
404 |
||
405 |
local structure P = OuterParse and K = OuterSyntax.Keyword in |
|
406 |
||
407 |
val old_undoP = (*same name for compatibility with PG/Isabelle99*) |
|
408 |
OuterSyntax.improper_command "undo" "undo last command (no output)" K.control |
|
409 |
(Scan.succeed (Toplevel.no_timing o IsarCmd.undo)); |
|
410 |
||
411 |
val undoP = |
|
412 |
OuterSyntax.improper_command "ProofGeneral.undo" "(internal)" K.control |
|
413 |
(Scan.succeed (Toplevel.no_timing o IsarCmd.undo)); |
|
414 |
||
415 |
val context_thy_onlyP = |
|
416 |
OuterSyntax.improper_command "ProofGeneral.context_thy_only" "(internal)" K.control |
|
417 |
(P.name >> (Toplevel.no_timing oo IsarThy.init_context update_thy_only)); |
|
418 |
||
419 |
val try_context_thy_onlyP = |
|
420 |
OuterSyntax.improper_command "ProofGeneral.try_context_thy_only" "(internal)" K.control |
|
421 |
(P.name >> (Toplevel.no_timing oo |
|
422 |
(Toplevel.imperative (K ()) oo IsarThy.init_context try_update_thy_only))); |
|
423 |
||
424 |
val restartP = |
|
425 |
OuterSyntax.improper_command "ProofGeneral.restart" "(internal)" K.control |
|
426 |
(P.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative isar_restart))); |
|
427 |
||
428 |
val kill_proofP = |
|
429 |
OuterSyntax.improper_command "ProofGeneral.kill_proof" "(internal)" K.control |
|
430 |
(Scan.succeed (Toplevel.no_timing o IsarCmd.kill_proof_notify tell_clear_goals)); |
|
431 |
||
432 |
val inform_file_processedP = |
|
433 |
OuterSyntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" K.control |
|
434 |
(P.name >> (Toplevel.no_timing oo |
|
435 |
(fn name => Toplevel.keep (proper_inform_file_processed name)))); |
|
436 |
||
437 |
val inform_file_retractedP = |
|
438 |
OuterSyntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" K.control |
|
439 |
(P.name >> (Toplevel.no_timing oo |
|
440 |
(fn name => Toplevel.imperative (fn () => inform_file_retracted name)))); |
|
441 |
||
442 |
fun init_outer_syntax () = OuterSyntax.add_parsers |
|
443 |
[old_undoP, undoP, restartP, kill_proofP, context_thy_onlyP, try_context_thy_onlyP, |
|
444 |
inform_file_processedP, inform_file_retractedP]; |
|
445 |
||
446 |
end; |
|
447 |
||
448 |
||
449 |
(* init *) |
|
450 |
||
451 |
val initialized = ref false; |
|
452 |
||
453 |
fun init isar = |
|
454 |
(conditional (not (! initialized)) (fn () => |
|
455 |
(if isar then setmp warning_fn (K ()) init_outer_syntax () else (); |
|
456 |
setup_xsymbols_output (); |
|
457 |
setup_messages (); |
|
458 |
setup_state (); |
|
459 |
setup_thy_loader (); |
|
13526 | 460 |
setup_present_hook (); |
12778 | 461 |
set initialized; ())); |
462 |
sync_thy_loader (); |
|
463 |
print_mode := proof_generalN :: (! print_mode \ proof_generalN); |
|
464 |
set quick_and_dirty; |
|
465 |
ThmDeps.enable (); |
|
466 |
if isar then ml_prompts "ML> " "ML# " |
|
467 |
else ml_prompts ("> " ^ oct_char "372") ("- " ^ oct_char "373"); |
|
468 |
if isar then (welcome (); Isar.sync_main ()) else isa_restart ()); |
|
469 |
||
470 |
||
471 |
||
472 |
(** generate keyword classification file **) |
|
473 |
||
474 |
local |
|
475 |
||
476 |
val regexp_meta = explode ".*+?[]^$"; |
|
477 |
val regexp_quote = implode o map (fn c => if c mem regexp_meta then "\\\\" ^ c else c) o explode; |
|
478 |
||
479 |
fun defconst name strs = |
|
480 |
"\n(defconst isar-keywords-" ^ name ^ |
|
481 |
"\n '(" ^ space_implode "\n " (map (quote o regexp_quote) strs) ^ "))\n"; |
|
482 |
||
483 |
fun make_elisp_commands commands kind = |
|
484 |
defconst kind (mapfilter (fn (c, _, k, _) => if k = kind then Some c else None) commands); |
|
485 |
||
486 |
fun make_elisp_syntax (keywords, commands) = |
|
487 |
";;\n\ |
|
488 |
\;; Keyword classification tables for Isabelle/Isar.\n\ |
|
489 |
\;; This file was generated by " ^ Session.name () ^ " -- DO NOT EDIT!\n\ |
|
490 |
\;;\n\ |
|
491 |
\;; $" ^ "Id$\n\ |
|
492 |
\;;\n" ^ |
|
493 |
defconst "major" (map #1 commands) ^ |
|
494 |
defconst "minor" (filter Syntax.is_identifier keywords) ^ |
|
495 |
implode (map (make_elisp_commands commands) OuterSyntax.Keyword.kinds) ^ |
|
496 |
"\n(provide 'isar-keywords)\n"; |
|
497 |
||
498 |
in |
|
499 |
||
500 |
fun write_keywords s = |
|
501 |
(init_outer_syntax (); |
|
502 |
File.write (Path.unpack ("isar-keywords" ^ (if s = "" then "" else "-" ^ s) ^".el")) |
|
503 |
(make_elisp_syntax (OuterSyntax.dest_keywords (), OuterSyntax.dest_parsers ()))); |
|
504 |
||
505 |
end; |
|
506 |
||
507 |
||
508 |
end; |
|
509 |
||
510 |
(*a hack for Proof General 3.2 to avoid problems with escapes in ML commands*) |
|
511 |
infix \\\\ val op \\\\ = op \\; |