| author | haftmann | 
| Fri, 19 Jun 2009 19:45:00 +0200 | |
| changeset 31725 | f08507464b9d | 
| parent 30670 | 9bb872667af6 | 
| child 32738 | 15bb09ca0378 | 
| permissions | -rw-r--r-- | 
| 21642 | 1  | 
(* Title: Pure/ProofGeneral/proof_general_emacs.ML  | 
2  | 
Author: David Aspinall and Markus Wenzel  | 
|
3  | 
||
4  | 
Isabelle/Isar configuration for Emacs Proof General.  | 
|
| 26549 | 5  | 
See also http://proofgeneral.inf.ed.ac.uk  | 
| 21642 | 6  | 
*)  | 
7  | 
||
8  | 
signature PROOF_GENERAL =  | 
|
9  | 
sig  | 
|
| 26549 | 10  | 
val test_markupN: string  | 
| 21642 | 11  | 
val init: bool -> unit  | 
| 24874 | 12  | 
val init_outer_syntax: unit -> unit  | 
| 24289 | 13  | 
val sendback: string -> Pretty.T list -> unit  | 
| 21642 | 14  | 
end;  | 
15  | 
||
16  | 
structure ProofGeneral: PROOF_GENERAL =  | 
|
17  | 
struct  | 
|
18  | 
||
| 21945 | 19  | 
|
| 21642 | 20  | 
(* print modes *)  | 
21  | 
||
22  | 
val proof_generalN = "ProofGeneralEmacs"; (*token markup (colouring vars, etc.)*)  | 
|
23  | 
val thm_depsN = "thm_deps"; (*meta-information about theorem deps*)  | 
|
| 26549 | 24  | 
val test_markupN = "test_markup"; (*XML markup for everything*)  | 
25  | 
||
| 
29321
 
6b9ecb3a70ab
removed print mode "PGASCII" -- 7-bit ASCII communication now always enabled;
 
wenzelm 
parents: 
28426 
diff
changeset
 | 
26  | 
fun special ch = Symbol.SOH ^ ch;  | 
| 21642 | 27  | 
|
28  | 
||
| 29326 | 29  | 
(* render markup *)  | 
| 21642 | 30  | 
|
31  | 
local  | 
|
32  | 
||
| 29326 | 33  | 
fun render_trees ts = fold render_tree ts  | 
34  | 
and render_tree (XML.Text s) = Buffer.add s  | 
|
35  | 
| render_tree (XML.Elem (name, props, ts)) =  | 
|
36  | 
let  | 
|
37  | 
val (bg1, en1) =  | 
|
38  | 
if name <> Markup.promptN andalso print_mode_active test_markupN  | 
|
39  | 
then XML.output_markup (name, props)  | 
|
40  | 
else Markup.no_output;  | 
|
41  | 
val (bg2, en2) =  | 
|
| 
30292
 
a3bb22493f11
render_tree: suppress markup only for empty body (of status messages, cf. da275b7809bd) in order to recover hilite;
 
wenzelm 
parents: 
29349 
diff
changeset
 | 
42  | 
if null ts then Markup.no_output  | 
| 29326 | 43  | 
else if name = Markup.stateN then (special "O" ^ "\n", "\n" ^ special "P")  | 
44  | 
else if name = Markup.sendbackN then (special "W", special "X")  | 
|
45  | 
else if name = Markup.hiliteN then (special "0", special "1")  | 
|
46  | 
else if name = Markup.tclassN then (special "B", special "A")  | 
|
47  | 
else if name = Markup.tfreeN then (special "C", special "A")  | 
|
48  | 
else if name = Markup.tvarN then (special "D", special "A")  | 
|
49  | 
else if name = Markup.freeN then (special "E", special "A")  | 
|
50  | 
else if name = Markup.boundN then (special "F", special "A")  | 
|
51  | 
else if name = Markup.varN then (special "G", special "A")  | 
|
52  | 
else if name = Markup.skolemN then (special "H", special "A")  | 
|
53  | 
else Markup.no_output;  | 
|
54  | 
in  | 
|
55  | 
Buffer.add bg1 #>  | 
|
56  | 
Buffer.add bg2 #>  | 
|
57  | 
render_trees ts #>  | 
|
58  | 
Buffer.add en2 #>  | 
|
59  | 
Buffer.add en1  | 
|
60  | 
end;  | 
|
| 21642 | 61  | 
|
62  | 
in  | 
|
63  | 
||
| 29326 | 64  | 
fun render text =  | 
65  | 
Buffer.content (render_trees (YXML.parse_body text) Buffer.empty);  | 
|
| 21642 | 66  | 
|
67  | 
end;  | 
|
68  | 
||
69  | 
||
| 29326 | 70  | 
(* messages *)  | 
71  | 
||
72  | 
fun message bg en prfx text =  | 
|
73  | 
(case render text of  | 
|
74  | 
"" => ()  | 
|
75  | 
| s => Output.writeln_default (enclose bg en (prefix_lines prfx s)));  | 
|
| 
23641
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
76  | 
|
| 29326 | 77  | 
fun setup_messages () =  | 
78  | 
(Output.writeln_fn := message "" "" "";  | 
|
| 
30670
 
9bb872667af6
suppress status output for traditional tty modes (including Proof General);
 
wenzelm 
parents: 
30292 
diff
changeset
 | 
79  | 
Output.status_fn := (fn _ => ());  | 
| 29326 | 80  | 
Output.priority_fn := message (special "I") (special "J") "";  | 
81  | 
Output.tracing_fn := message (special "I" ^ special "V") (special "J") "";  | 
|
82  | 
Output.debug_fn := message (special "K") (special "L") "+++ ";  | 
|
83  | 
Output.warning_fn := message (special "K") (special "L") "### ";  | 
|
84  | 
Output.error_fn := message (special "M") (special "N") "*** ";  | 
|
85  | 
Output.prompt_fn := (fn s => Output.std_output (render s ^ special "S")));  | 
|
86  | 
||
87  | 
fun panic s =  | 
|
88  | 
  (message (special "M") (special "N") "!!! " ("## SYSTEM EXIT ##\n" ^ s); exit 1);
 | 
|
| 
23641
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
89  | 
|
| 
 
d6f9d3acffaa
toplevel prompt/print_state: proper markup, removed hooks;
 
wenzelm 
parents: 
23621 
diff
changeset
 | 
90  | 
|
| 29326 | 91  | 
(* notification *)  | 
| 21642 | 92  | 
|
| 29326 | 93  | 
val emacs_notify = message (special "I") (special "J") "";  | 
| 21642 | 94  | 
|
95  | 
fun tell_clear_goals () =  | 
|
| 21940 | 96  | 
emacs_notify "Proof General, please clear the goals buffer.";  | 
| 21642 | 97  | 
|
98  | 
fun tell_clear_response () =  | 
|
| 21940 | 99  | 
emacs_notify "Proof General, please clear the response buffer.";  | 
| 21642 | 100  | 
|
101  | 
fun tell_file_loaded path =  | 
|
| 21940 | 102  | 
  emacs_notify ("Proof General, this file is loaded: " ^ quote (File.platform_path path));
 | 
| 21642 | 103  | 
|
104  | 
fun tell_file_retracted path =  | 
|
| 21940 | 105  | 
  emacs_notify ("Proof General, you can unlock the file " ^ quote (File.platform_path path));
 | 
| 21642 | 106  | 
|
| 24289 | 107  | 
fun sendback heading prts =  | 
108  | 
Pretty.writeln (Pretty.big_list heading [Pretty.markup Markup.sendback prts]);  | 
|
109  | 
||
| 21642 | 110  | 
|
111  | 
(* theory loader actions *)  | 
|
112  | 
||
113  | 
local  | 
|
114  | 
||
115  | 
fun trace_action action name =  | 
|
116  | 
if action = ThyInfo.Update then  | 
|
117  | 
List.app tell_file_loaded (ThyInfo.loaded_files name)  | 
|
118  | 
else if action = ThyInfo.Outdate orelse action = ThyInfo.Remove then  | 
|
119  | 
List.app tell_file_retracted (ThyInfo.loaded_files name)  | 
|
120  | 
else ();  | 
|
121  | 
||
122  | 
in  | 
|
123  | 
fun setup_thy_loader () = ThyInfo.add_hook trace_action;  | 
|
| 26613 | 124  | 
fun sync_thy_loader () = List.app (trace_action ThyInfo.Update) (ThyInfo.get_names ());  | 
| 21642 | 125  | 
end;  | 
126  | 
||
127  | 
||
| 21948 | 128  | 
(* get informed about files *)  | 
| 21642 | 129  | 
|
| 25442 | 130  | 
(*liberal low-level version*)  | 
131  | 
val thy_name = perhaps (try (unsuffix ".thy")) o List.last o space_explode "/";  | 
|
132  | 
||
133  | 
val inform_file_retracted = ThyInfo.if_known_thy ThyInfo.remove_thy o thy_name;  | 
|
134  | 
||
| 
27577
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
135  | 
fun inform_file_processed file =  | 
| 
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
136  | 
let  | 
| 
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
137  | 
val name = thy_name file;  | 
| 
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
138  | 
    val _ = name = "" andalso error ("Bad file name: " ^ quote file);
 | 
| 
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
139  | 
val _ =  | 
| 
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
140  | 
if ThyInfo.check_known_thy name then  | 
| 28426 | 141  | 
(Isar.>> (Toplevel.commit_exit Position.none);  | 
142  | 
ThyInfo.touch_child_thys name; ThyInfo.register_thy name)  | 
|
| 
27585
 
2234ace5b538
inform_file_processed: try harder not to fail, ensure
 
wenzelm 
parents: 
27577 
diff
changeset
 | 
143  | 
handle ERROR msg =>  | 
| 27591 | 144  | 
(warning (cat_lines ["Failed to register theory: " ^ quote name, msg]);  | 
| 
27585
 
2234ace5b538
inform_file_processed: try harder not to fail, ensure
 
wenzelm 
parents: 
27577 
diff
changeset
 | 
145  | 
tell_file_retracted (ThyLoad.thy_path name))  | 
| 
27577
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
146  | 
else ();  | 
| 29349 | 147  | 
val _ = Isar.init ();  | 
| 
27577
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
148  | 
in () end;  | 
| 21642 | 149  | 
|
150  | 
||
151  | 
(* restart top-level loop (keeps most state information) *)  | 
|
152  | 
||
153  | 
val welcome = priority o Session.welcome;  | 
|
154  | 
||
155  | 
fun restart () =  | 
|
| 21940 | 156  | 
(sync_thy_loader ();  | 
157  | 
tell_clear_goals ();  | 
|
158  | 
tell_clear_response ();  | 
|
| 29349 | 159  | 
Isar.init ();  | 
| 
27577
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
160  | 
welcome ());  | 
| 21642 | 161  | 
|
162  | 
||
163  | 
(* theorem dependency output *)  | 
|
164  | 
||
165  | 
local  | 
|
166  | 
||
167  | 
val spaces_quote = space_implode " " o map quote;  | 
|
168  | 
||
169  | 
fun thm_deps_message (thms, deps) =  | 
|
| 21948 | 170  | 
  emacs_notify ("Proof General, theorem dependencies of " ^ thms ^ " are " ^ deps);
 | 
| 21642 | 171  | 
|
| 
28096
 
046418f64474
theorem dependency output: Toplevel.add_hook, ProofGeneralPgip.new_thms_deps;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
172  | 
in  | 
| 
 
046418f64474
theorem dependency output: Toplevel.add_hook, ProofGeneralPgip.new_thms_deps;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
173  | 
|
| 
28103
 
b79e61861f0f
simplified Toplevel.add_hook: cover successful transactions only;
 
wenzelm 
parents: 
28100 
diff
changeset
 | 
174  | 
fun setup_present_hook () = Toplevel.add_hook (fn _ => fn state => fn state' =>  | 
| 28106 | 175  | 
if print_mode_active thm_depsN andalso  | 
176  | 
can Toplevel.theory_of state andalso Toplevel.is_theory state'  | 
|
177  | 
then  | 
|
| 
28100
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28096 
diff
changeset
 | 
178  | 
let val (names, deps) =  | 
| 
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28096 
diff
changeset
 | 
179  | 
ProofGeneralPgip.new_thms_deps (Toplevel.theory_of state) (Toplevel.theory_of state')  | 
| 
 
7650d5e0f8fb
refined theorem dependency output: previous state needs to contain a theory (not empty toplevel);
 
wenzelm 
parents: 
28096 
diff
changeset
 | 
180  | 
in  | 
| 21968 | 181  | 
if null names orelse null deps then ()  | 
182  | 
else thm_deps_message (spaces_quote names, spaces_quote deps)  | 
|
183  | 
end  | 
|
| 
28096
 
046418f64474
theorem dependency output: Toplevel.add_hook, ProofGeneralPgip.new_thms_deps;
 
wenzelm 
parents: 
28020 
diff
changeset
 | 
184  | 
else ());  | 
| 21642 | 185  | 
|
186  | 
end;  | 
|
187  | 
||
188  | 
||
189  | 
(* additional outer syntax for Isar *)  | 
|
190  | 
||
191  | 
local structure P = OuterParse and K = OuterKeyword in  | 
|
192  | 
||
| 
25192
 
b568f8c5d5ca
made command 'undo' silent ('ProofGeneral.undo' becomes a historical relic);
 
wenzelm 
parents: 
24882 
diff
changeset
 | 
193  | 
fun undoP () = (*undo without output -- historical*)  | 
| 21642 | 194  | 
OuterSyntax.improper_command "ProofGeneral.undo" "(internal)" K.control  | 
| 27535 | 195  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Isar.undo 1)));  | 
| 21642 | 196  | 
|
| 24867 | 197  | 
fun restartP () =  | 
| 21642 | 198  | 
OuterSyntax.improper_command "ProofGeneral.restart" "(internal)" K.control  | 
199  | 
(P.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative restart)));  | 
|
200  | 
||
| 24867 | 201  | 
fun kill_proofP () =  | 
| 21642 | 202  | 
OuterSyntax.improper_command "ProofGeneral.kill_proof" "(internal)" K.control  | 
| 27532 | 203  | 
(Scan.succeed (Toplevel.no_timing o  | 
204  | 
Toplevel.imperative (fn () => (Isar.kill_proof (); tell_clear_goals ()))));  | 
|
205  | 
||
| 24867 | 206  | 
fun inform_file_processedP () =  | 
| 21642 | 207  | 
OuterSyntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" K.control  | 
| 
27577
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
208  | 
(P.name >> (fn file =>  | 
| 
 
7c7a9a343ca5
proper_inform_file_processed: new implementation based on global Isar state (cf. isar.ML);
 
wenzelm 
parents: 
27537 
diff
changeset
 | 
209  | 
Toplevel.no_timing o Toplevel.imperative (fn () => inform_file_processed file)));  | 
| 21642 | 210  | 
|
| 24867 | 211  | 
fun inform_file_retractedP () =  | 
| 21642 | 212  | 
OuterSyntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" K.control  | 
213  | 
(P.name >> (Toplevel.no_timing oo  | 
|
214  | 
(fn file => Toplevel.imperative (fn () => inform_file_retracted file))));  | 
|
215  | 
||
| 24867 | 216  | 
fun process_pgipP () =  | 
| 21642 | 217  | 
OuterSyntax.improper_command "ProofGeneral.process_pgip" "(internal)" K.control  | 
218  | 
(P.text >> (Toplevel.no_timing oo  | 
|
219  | 
(fn txt => Toplevel.imperative (fn () => ProofGeneralPgip.process_pgip txt))));  | 
|
220  | 
||
| 24867 | 221  | 
fun init_outer_syntax () = List.app (fn f => f ())  | 
| 27535 | 222  | 
[undoP, restartP, kill_proofP, inform_file_processedP, inform_file_retractedP, process_pgipP];  | 
| 21642 | 223  | 
|
224  | 
end;  | 
|
225  | 
||
226  | 
||
227  | 
(* init *)  | 
|
228  | 
||
229  | 
val initialized = ref false;  | 
|
230  | 
||
| 
22699
 
938c1011ac94
removed unused Output.panic hook -- internal to PG wrapper;
 
wenzelm 
parents: 
22678 
diff
changeset
 | 
231  | 
fun init false = panic "No Proof General interface support for Isabelle/classic mode."  | 
| 21940 | 232  | 
| init true =  | 
| 21968 | 233  | 
(! initialized orelse  | 
| 22590 | 234  | 
(Output.no_warnings init_outer_syntax ();  | 
| 29326 | 235  | 
Output.add_mode Symbol.xsymbolsN Symbol.output Output.default_escape;  | 
236  | 
Output.add_mode proof_generalN Output.default_output Output.default_escape;  | 
|
237  | 
Markup.add_mode proof_generalN YXML.output_markup;  | 
|
| 21968 | 238  | 
setup_messages ();  | 
| 22590 | 239  | 
ProofGeneralPgip.init_pgip_channel (! Output.priority_fn);  | 
| 21968 | 240  | 
setup_thy_loader ();  | 
241  | 
setup_present_hook ();  | 
|
242  | 
set initialized);  | 
|
243  | 
sync_thy_loader ();  | 
|
| 25749 | 244  | 
change print_mode (update (op =) proof_generalN);  | 
| 26643 | 245  | 
       Isar.toplevel_loop {init = true, welcome = true, sync = true, secure = Secure.is_secure ()});
 | 
| 21642 | 246  | 
|
247  | 
end;  |