| 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
 | 
|  |     14 |   val inform_file_processed: string -> unit
 | 
|  |     15 |   val inform_file_retracted: string -> unit
 | 
|  |     16 |   val thms_containing: string list -> unit
 | 
|  |     17 |   val help: unit -> unit
 | 
|  |     18 |   val show_context: unit -> theory
 | 
|  |     19 |   val kill_goal: unit -> unit
 | 
|  |     20 |   val repeat_undo: int -> unit
 | 
|  |     21 |   val isa_restart: unit -> unit
 | 
| 12833 |     22 |   val full_proofs: bool -> unit
 | 
| 12778 |     23 |   val init: bool -> unit
 | 
|  |     24 |   val write_keywords: string -> unit
 | 
|  |     25 | end;
 | 
|  |     26 | 
 | 
|  |     27 | structure ProofGeneral: PROOF_GENERAL =
 | 
|  |     28 | struct
 | 
|  |     29 | 
 | 
|  |     30 | (* print modes *)
 | 
|  |     31 | 
 | 
|  |     32 | val proof_generalN = "ProofGeneral";
 | 
|  |     33 | val xsymbolsN = "xsymbols";
 | 
|  |     34 | 
 | 
|  |     35 | val pgmlN = "PGML";
 | 
|  |     36 | fun pgml () = pgmlN mem_string ! print_mode;
 | 
|  |     37 | 
 | 
|  |     38 | 
 | 
|  |     39 | (* text output *)
 | 
|  |     40 | 
 | 
|  |     41 | local
 | 
|  |     42 | 
 | 
|  |     43 | fun xsymbols_output s =
 | 
|  |     44 |   if xsymbolsN mem_string ! print_mode andalso exists_string (equal "\\") s then
 | 
|  |     45 |     let val syms = Symbol.explode s
 | 
|  |     46 |     in (implode (map (fn "\\" => "\\\\" | c => c) syms), real (Symbol.length syms)) end
 | 
|  |     47 |   else (s, real (size s));
 | 
|  |     48 | 
 | 
|  |     49 | fun pgml_output (s, len) =
 | 
|  |     50 |   if pgml () then (XML.text s, len)
 | 
|  |     51 |   else (s, len);
 | 
|  |     52 | 
 | 
|  |     53 | in
 | 
|  |     54 | 
 | 
|  |     55 | fun setup_xsymbols_output () =
 | 
|  |     56 |   Symbol.add_mode proof_generalN (pgml_output o xsymbols_output, Symbol.default_indent);
 | 
|  |     57 | 
 | 
|  |     58 | end;
 | 
|  |     59 | 
 | 
|  |     60 | 
 | 
|  |     61 | (* token translations *)
 | 
|  |     62 | 
 | 
|  |     63 | local
 | 
|  |     64 | 
 | 
|  |     65 | val end_tag = oct_char "350";
 | 
|  |     66 | val class_tag = ("class", oct_char "351");
 | 
|  |     67 | val tfree_tag = ("tfree", oct_char "352");
 | 
|  |     68 | val tvar_tag = ("tvar", oct_char "353");
 | 
|  |     69 | val free_tag = ("free", oct_char "354");
 | 
|  |     70 | val bound_tag = ("bound", oct_char "355");
 | 
|  |     71 | val var_tag = ("var", oct_char "356");
 | 
|  |     72 | val skolem_tag = ("skolem", oct_char "357");
 | 
|  |     73 | 
 | 
|  |     74 | fun xml_atom kind x = XML.element "atom" [("kind", kind)] [XML.text x];
 | 
|  |     75 | 
 | 
|  |     76 | fun tagit (kind, bg_tag) x =
 | 
|  |     77 |   (if pgml () then xml_atom kind x else bg_tag ^ x ^ end_tag,
 | 
|  |     78 |     real (Symbol.length (Symbol.explode x)));
 | 
|  |     79 | 
 | 
|  |     80 | fun free_or_skolem x =
 | 
|  |     81 |   (case try Syntax.dest_skolem x of
 | 
|  |     82 |     None => tagit free_tag x
 | 
|  |     83 |   | Some x' => tagit skolem_tag x');
 | 
|  |     84 | 
 | 
|  |     85 | fun var_or_skolem s =
 | 
|  |     86 |   (case Syntax.read_var s of
 | 
|  |     87 |     Var ((x, i), _) =>
 | 
|  |     88 |       (case try Syntax.dest_skolem x of
 | 
|  |     89 |         None => tagit var_tag s
 | 
|  |     90 |       | Some x' => tagit skolem_tag (Syntax.string_of_vname (x', i)))
 | 
|  |     91 |   | _ => tagit var_tag s);
 | 
|  |     92 | 
 | 
|  |     93 | val proof_general_trans =
 | 
|  |     94 |  Syntax.tokentrans_mode proof_generalN
 | 
|  |     95 |   [("class", tagit class_tag),
 | 
|  |     96 |    ("tfree", tagit tfree_tag),
 | 
|  |     97 |    ("tvar", tagit tvar_tag),
 | 
|  |     98 |    ("free", free_or_skolem),
 | 
|  |     99 |    ("bound", tagit bound_tag),
 | 
|  |    100 |    ("var", var_or_skolem)];
 | 
|  |    101 | 
 | 
|  |    102 | in val setup = [Theory.add_tokentrfuns proof_general_trans] end;
 | 
|  |    103 | 
 | 
|  |    104 | 
 | 
|  |    105 | 
 | 
|  |    106 | (* messages and notification *)
 | 
|  |    107 | 
 | 
|  |    108 | local
 | 
|  |    109 | 
 | 
|  |    110 | fun decorated_output bg en prfx =
 | 
|  |    111 |   writeln_default o enclose bg en o prefix_lines prfx;
 | 
|  |    112 | 
 | 
|  |    113 | fun message kind bg en prfx s =
 | 
|  |    114 |   if pgml () then writeln_default (XML.element kind [] [prefix_lines prfx s])
 | 
|  |    115 |   else decorated_output bg en prfx s;
 | 
|  |    116 | 
 | 
|  |    117 | val notify = message "notify" (oct_char "360") (oct_char "361") "";
 | 
|  |    118 | 
 | 
|  |    119 | in
 | 
|  |    120 | 
 | 
|  |    121 | fun setup_messages () =
 | 
|  |    122 |  (writeln_fn := message "output" "" "" "";
 | 
|  |    123 |   priority_fn := message "information" (oct_char "360") (oct_char "361") "";
 | 
|  |    124 |   tracing_fn := message "tracing" (oct_char "360" ^ oct_char "375") (oct_char "361") "";
 | 
|  |    125 |   warning_fn := message "warning" (oct_char "362") (oct_char "363") "### ";
 | 
|  |    126 |   error_fn := message "error" (oct_char "364") (oct_char "365") "*** ");
 | 
|  |    127 | 
 | 
|  |    128 | fun tell_clear_goals () = notify "Proof General, please clear the goals buffer.";
 | 
|  |    129 | fun tell_clear_response () = notify "Proof General, please clear the response buffer.";
 | 
|  |    130 | fun tell_file msg path = notify ("Proof General, " ^ msg ^ " " ^ quote (File.sysify_path path));
 | 
|  |    131 | 
 | 
|  |    132 | end;
 | 
|  |    133 | 
 | 
|  |    134 | 
 | 
|  |    135 | (* theory / proof state output *)
 | 
|  |    136 | 
 | 
|  |    137 | local
 | 
|  |    138 | 
 | 
|  |    139 | fun tmp_markers f =
 | 
|  |    140 |   setmp Display.current_goals_markers (oct_char "366", oct_char "367", "") f ();
 | 
|  |    141 | 
 | 
|  |    142 | fun statedisplay prts =
 | 
|  |    143 |   writeln_default (XML.element "statedisplay" [] [Pretty.string_of (Pretty.chunks prts)]);
 | 
|  |    144 | 
 | 
|  |    145 | fun print_current_goals n m st =
 | 
|  |    146 |   if pgml () then statedisplay (Display.pretty_current_goals n m st)
 | 
|  |    147 |   else tmp_markers (fn () => Display.print_current_goals_default n m st);
 | 
|  |    148 | 
 | 
|  |    149 | fun print_state b st =
 | 
|  |    150 |   if pgml () then statedisplay (Toplevel.pretty_state b st)
 | 
|  |    151 |   else tmp_markers (fn () => Toplevel.print_state_default b st);
 | 
|  |    152 | 
 | 
|  |    153 | in
 | 
|  |    154 | 
 | 
|  |    155 | fun setup_state () =
 | 
|  |    156 |  (Display.print_current_goals_fn := print_current_goals;
 | 
|  |    157 |   Toplevel.print_state_fn := print_state;
 | 
|  |    158 |   Toplevel.prompt_state_fn := (suffix (oct_char "372") o Toplevel.prompt_state_default));
 | 
|  |    159 | 
 | 
|  |    160 | end;
 | 
|  |    161 | 
 | 
|  |    162 | 
 | 
|  |    163 | (* theory loader actions *)
 | 
|  |    164 | 
 | 
|  |    165 | local
 | 
|  |    166 | 
 | 
|  |    167 | fun add_master_files name files =
 | 
|  |    168 |   let val masters = [ThyLoad.thy_path name, ThyLoad.ml_path name]
 | 
|  |    169 |   in masters @ gen_rems (op = o pairself Path.base) (files, masters) end;
 | 
|  |    170 | 
 | 
|  |    171 | fun trace_action action name =
 | 
|  |    172 |   if action = ThyInfo.Update then
 | 
|  |    173 |     seq (tell_file "this file is loaded:") (ThyInfo.loaded_files name)
 | 
|  |    174 |   else if action = ThyInfo.Outdate orelse action = ThyInfo.Remove then
 | 
|  |    175 |     seq (tell_file "you can unlock the file") (add_master_files name (ThyInfo.loaded_files name))
 | 
|  |    176 |   else ();
 | 
|  |    177 | 
 | 
|  |    178 | in
 | 
|  |    179 |   fun setup_thy_loader () = ThyInfo.add_hook trace_action;
 | 
|  |    180 |   fun sync_thy_loader () = seq (trace_action ThyInfo.Update) (ThyInfo.names ());
 | 
|  |    181 | end;
 | 
|  |    182 | 
 | 
|  |    183 | 
 | 
|  |    184 | (* prepare theory context *)
 | 
|  |    185 | 
 | 
|  |    186 | val thy_name = Path.pack o Path.drop_ext o Path.base o Path.unpack;
 | 
|  |    187 | val update_thy_only = setmp MetaSimplifier.trace_simp false ThyInfo.update_thy_only;
 | 
|  |    188 | 
 | 
|  |    189 | fun which_context () =
 | 
|  |    190 |   (case Context.get_context () of
 | 
|  |    191 |     Some thy => "  Using current (dynamic!) one: " ^
 | 
|  |    192 |       (case try PureThy.get_name thy of Some name => quote name | None => "<unnamed>")
 | 
|  |    193 |   | None => "");
 | 
|  |    194 | 
 | 
|  |    195 | fun try_update_thy_only file =
 | 
|  |    196 |   ThyLoad.cond_add_path (Path.dir (Path.unpack file)) (fn () =>
 | 
|  |    197 |     let val name = thy_name file in
 | 
|  |    198 |       if is_some (ThyLoad.check_file (ThyLoad.thy_path name)) then update_thy_only name
 | 
|  |    199 |       else warning ("Unkown theory context of ML file." ^ which_context ())
 | 
|  |    200 |     end) ();
 | 
|  |    201 | 
 | 
|  |    202 | 
 | 
|  |    203 | (* get informed about files *)
 | 
|  |    204 | 
 | 
|  |    205 | val touch_child_thys = ThyInfo.if_known_thy ThyInfo.touch_child_thys;
 | 
|  |    206 | 
 | 
|  |    207 | val inform_file_processed = touch_child_thys o thy_name;
 | 
|  |    208 | val inform_file_retracted = touch_child_thys o thy_name;
 | 
|  |    209 | 
 | 
|  |    210 | fun proper_inform_file_processed file state =
 | 
|  |    211 |   let val name = thy_name file in
 | 
|  |    212 |     touch_child_thys name;
 | 
|  |    213 |     if not (Toplevel.is_toplevel state) then
 | 
|  |    214 |       warning ("Not at toplevel -- cannot register theory " ^ quote name)
 | 
|  |    215 |     else Library.transform_error ThyInfo.pretend_use_thy_only name handle ERROR_MESSAGE msg =>
 | 
|  |    216 |       (warning msg; warning ("Failed to register theory " ^ quote name))
 | 
|  |    217 |   end;
 | 
|  |    218 | 
 | 
|  |    219 | 
 | 
|  |    220 | (* misc commands for ProofGeneral/isa *)
 | 
|  |    221 | 
 | 
|  |    222 | fun thms_containing ss =
 | 
|  |    223 |   let
 | 
|  |    224 |     val thy = the_context ();
 | 
|  |    225 |     fun read_term s = Thm.term_of (Thm.read_cterm (Theory.sign_of thy) (s, TypeInfer.logicT));
 | 
|  |    226 |   in ThmDatabase.print_thms_containing thy (map read_term ss) end;
 | 
|  |    227 | 
 | 
|  |    228 | val welcome = priority o Session.welcome;
 | 
|  |    229 | val help = welcome;
 | 
|  |    230 | val show_context = Context.the_context;
 | 
|  |    231 | 
 | 
|  |    232 | fun kill_goal () = (Goals.reset_goals (); tell_clear_goals ());
 | 
|  |    233 | 
 | 
|  |    234 | fun no_print_goals f = setmp Display.print_current_goals_fn (fn _ => fn _ => fn _ => ()) f;
 | 
|  |    235 | 
 | 
|  |    236 | fun repeat_undo 0 = ()
 | 
|  |    237 |   | repeat_undo 1 = undo ()
 | 
|  |    238 |   | repeat_undo n = (no_print_goals undo (); repeat_undo (n - 1));
 | 
|  |    239 | 
 | 
|  |    240 | 
 | 
|  |    241 | (* restart top-level loop (keeps most state information) *)
 | 
|  |    242 | 
 | 
|  |    243 | local
 | 
|  |    244 | 
 | 
|  |    245 | fun restart isar =
 | 
|  |    246 |  (if isar then tell_clear_goals () else kill_goal ();
 | 
|  |    247 |   tell_clear_response ();
 | 
|  |    248 |   welcome ());
 | 
|  |    249 | 
 | 
|  |    250 | in
 | 
|  |    251 | 
 | 
|  |    252 | fun isa_restart () = restart false;
 | 
|  |    253 | fun isar_restart () = (sync_thy_loader (); restart true; raise Toplevel.RESTART);
 | 
|  |    254 | 
 | 
|  |    255 | end;
 | 
|  |    256 | 
 | 
|  |    257 | 
 | 
| 12833 |    258 | fun full_proofs true = proofs := 2
 | 
|  |    259 |   | full_proofs false = proofs := 1;
 | 
|  |    260 | 
 | 
|  |    261 | 
 | 
| 12778 |    262 | (* outer syntax *)
 | 
|  |    263 | 
 | 
|  |    264 | local structure P = OuterParse and K = OuterSyntax.Keyword in
 | 
|  |    265 | 
 | 
|  |    266 | val old_undoP = (*same name for compatibility with PG/Isabelle99*)
 | 
|  |    267 |   OuterSyntax.improper_command "undo" "undo last command (no output)" K.control
 | 
|  |    268 |     (Scan.succeed (Toplevel.no_timing o IsarCmd.undo));
 | 
|  |    269 | 
 | 
|  |    270 | val undoP =
 | 
|  |    271 |   OuterSyntax.improper_command "ProofGeneral.undo" "(internal)" K.control
 | 
|  |    272 |     (Scan.succeed (Toplevel.no_timing o IsarCmd.undo));
 | 
|  |    273 | 
 | 
|  |    274 | val context_thy_onlyP =
 | 
|  |    275 |   OuterSyntax.improper_command "ProofGeneral.context_thy_only" "(internal)" K.control
 | 
|  |    276 |     (P.name >> (Toplevel.no_timing oo IsarThy.init_context update_thy_only));
 | 
|  |    277 | 
 | 
|  |    278 | val try_context_thy_onlyP =
 | 
|  |    279 |   OuterSyntax.improper_command "ProofGeneral.try_context_thy_only" "(internal)" K.control
 | 
|  |    280 |     (P.name >> (Toplevel.no_timing oo
 | 
|  |    281 |       (Toplevel.imperative (K ()) oo IsarThy.init_context try_update_thy_only)));
 | 
|  |    282 | 
 | 
|  |    283 | val restartP =
 | 
|  |    284 |   OuterSyntax.improper_command "ProofGeneral.restart" "(internal)" K.control
 | 
|  |    285 |     (P.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative isar_restart)));
 | 
|  |    286 | 
 | 
|  |    287 | val kill_proofP =
 | 
|  |    288 |   OuterSyntax.improper_command "ProofGeneral.kill_proof" "(internal)" K.control
 | 
|  |    289 |     (Scan.succeed (Toplevel.no_timing o IsarCmd.kill_proof_notify tell_clear_goals));
 | 
|  |    290 | 
 | 
|  |    291 | val inform_file_processedP =
 | 
|  |    292 |   OuterSyntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" K.control
 | 
|  |    293 |     (P.name >> (Toplevel.no_timing oo
 | 
|  |    294 |       (fn name => Toplevel.keep (proper_inform_file_processed name))));
 | 
|  |    295 | 
 | 
|  |    296 | val inform_file_retractedP =
 | 
|  |    297 |   OuterSyntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" K.control
 | 
|  |    298 |     (P.name >> (Toplevel.no_timing oo
 | 
|  |    299 |       (fn name => Toplevel.imperative (fn () => inform_file_retracted name))));
 | 
|  |    300 | 
 | 
|  |    301 | fun init_outer_syntax () = OuterSyntax.add_parsers
 | 
|  |    302 |  [old_undoP, undoP, restartP, kill_proofP, context_thy_onlyP, try_context_thy_onlyP,
 | 
|  |    303 |   inform_file_processedP, inform_file_retractedP];
 | 
|  |    304 | 
 | 
|  |    305 | end;
 | 
|  |    306 | 
 | 
|  |    307 | 
 | 
|  |    308 | (* init *)
 | 
|  |    309 | 
 | 
|  |    310 | val initialized = ref false;
 | 
|  |    311 | 
 | 
|  |    312 | fun init isar =
 | 
|  |    313 |  (conditional (not (! initialized)) (fn () =>
 | 
|  |    314 |    (if isar then setmp warning_fn (K ()) init_outer_syntax () else ();
 | 
|  |    315 |     setup_xsymbols_output ();
 | 
|  |    316 |     setup_messages ();
 | 
|  |    317 |     setup_state ();
 | 
|  |    318 |     setup_thy_loader ();
 | 
|  |    319 |     set initialized; ()));
 | 
|  |    320 |   sync_thy_loader ();
 | 
|  |    321 |   print_mode := proof_generalN :: (! print_mode \ proof_generalN);
 | 
|  |    322 |   set quick_and_dirty;
 | 
|  |    323 |   ThmDeps.enable ();
 | 
|  |    324 |   if isar then ml_prompts "ML> " "ML# "
 | 
|  |    325 |   else ml_prompts ("> " ^ oct_char "372") ("- " ^ oct_char "373");
 | 
|  |    326 |   if isar then (welcome (); Isar.sync_main ()) else isa_restart ());
 | 
|  |    327 | 
 | 
|  |    328 | 
 | 
|  |    329 | 
 | 
|  |    330 | (** generate keyword classification file **)
 | 
|  |    331 | 
 | 
|  |    332 | local
 | 
|  |    333 | 
 | 
|  |    334 | val regexp_meta = explode ".*+?[]^$";
 | 
|  |    335 | val regexp_quote = implode o map (fn c => if c mem regexp_meta then "\\\\" ^ c else c) o explode;
 | 
|  |    336 | 
 | 
|  |    337 | fun defconst name strs =
 | 
|  |    338 |   "\n(defconst isar-keywords-" ^ name ^
 | 
|  |    339 |   "\n  '(" ^ space_implode "\n    " (map (quote o regexp_quote) strs) ^ "))\n";
 | 
|  |    340 | 
 | 
|  |    341 | fun make_elisp_commands commands kind =
 | 
|  |    342 |   defconst kind (mapfilter (fn (c, _, k, _) => if k = kind then Some c else None) commands);
 | 
|  |    343 | 
 | 
|  |    344 | fun make_elisp_syntax (keywords, commands) =
 | 
|  |    345 |   ";;\n\
 | 
|  |    346 |   \;; Keyword classification tables for Isabelle/Isar.\n\
 | 
|  |    347 |   \;; This file was generated by " ^ Session.name () ^ " -- DO NOT EDIT!\n\
 | 
|  |    348 |   \;;\n\
 | 
|  |    349 |   \;; $" ^ "Id$\n\
 | 
|  |    350 |   \;;\n" ^
 | 
|  |    351 |   defconst "major" (map #1 commands) ^
 | 
|  |    352 |   defconst "minor" (filter Syntax.is_identifier keywords) ^
 | 
|  |    353 |   implode (map (make_elisp_commands commands) OuterSyntax.Keyword.kinds) ^
 | 
|  |    354 |   "\n(provide 'isar-keywords)\n";
 | 
|  |    355 | 
 | 
|  |    356 | in
 | 
|  |    357 | 
 | 
|  |    358 | fun write_keywords s =
 | 
|  |    359 |   (init_outer_syntax ();
 | 
|  |    360 |     File.write (Path.unpack ("isar-keywords" ^ (if s = "" then "" else "-" ^ s) ^".el"))
 | 
|  |    361 |       (make_elisp_syntax (OuterSyntax.dest_keywords (), OuterSyntax.dest_parsers ())));
 | 
|  |    362 | 
 | 
|  |    363 | end;
 | 
|  |    364 | 
 | 
|  |    365 | 
 | 
|  |    366 | end;
 | 
|  |    367 | 
 | 
|  |    368 | (*a hack for Proof General 3.2 to avoid problems with escapes in ML commands*)
 | 
|  |    369 | infix \\\\ val op \\\\ = op \\;
 |