--- a/src/Pure/ProofGeneral/proof_general.ML Wed May 15 20:28:43 2013 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,468 +0,0 @@
-(* Title: Pure/ProofGeneral/proof_general.ML
- Author: David Aspinall
- Author: Makarius
-
-Isabelle/Isar configuration for Proof General / Emacs.
-See also http://proofgeneral.inf.ed.ac.uk
-*)
-
-signature PROOF_GENERAL =
-sig
- type category = string
- val category_display: category
- val category_advanced_display: category
- val category_tracing: category
- val category_proof: category
- type pgiptype = string
- val pgipbool: pgiptype
- val pgipint: pgiptype
- val pgipfloat: pgiptype
- val pgipstring: pgiptype
- val preference_raw: category ->
- (unit -> string) -> (string -> unit) -> string -> string -> string -> unit
- val preference_bool: category -> bool Unsynchronized.ref -> string -> string -> unit
- val preference_int: category -> int Unsynchronized.ref -> string -> string -> unit
- val preference_real: category -> real Unsynchronized.ref -> string -> string -> unit
- val preference_string: category -> string Unsynchronized.ref -> string -> string -> unit
- val preference_option: category -> string -> string -> string -> unit
- structure ThyLoad: sig val add_path: string -> unit end
- val init: unit -> unit
- val thm_depsN: string
-end;
-
-structure ProofGeneral: PROOF_GENERAL =
-struct
-
-(** preferences **)
-
-(* type preference *)
-
-type category = string;
-val category_display = "Display";
-val category_advanced_display = "Advanced Display";
-val category_tracing = "Tracing";
-val category_proof = "Proof";
-
-type pgiptype = string;
-val pgipbool = "pgipbool";
-val pgipint = "pgipint";
-val pgipfloat = "pgipint"; (*NB: PG 3.7.x and 4.0 lack pgipfloat, but accept floats as pgipint*)
-val pgipstring = "pgipstring";
-
-type preference =
- {name: string,
- descr: string,
- default: string,
- pgiptype: pgiptype,
- get: unit -> string,
- set: string -> unit};
-
-
-(* global preferences *)
-
-local
- val preferences =
- Synchronized.var "ProofGeneral.preferences"
- ([(category_display, []),
- (category_advanced_display, []),
- (category_tracing, []),
- (category_proof, [])]: (category * preference list) list);
-in
-
-fun add_preference category (pref: preference) =
- Synchronized.change preferences (map (fn (cat, prefs) =>
- if cat <> category then (cat, prefs)
- else
- if exists (fn {name, ...} => name = #name pref) prefs
- then (warning ("Preference already exists: " ^ quote (#name pref)); (cat, prefs))
- else (cat, prefs @ [pref])));
-
-fun get_preferences () = Synchronized.value preferences;
-
-end;
-
-fun set_preference name value =
- let val prefs = map (`(#name)) (maps snd (get_preferences ())) in
- (case AList.lookup (op =) prefs name of
- SOME {set, ...} => set value
- | NONE => error ("Unknown prover preference: " ^ quote name))
- end;
-
-
-(* raw preferences *)
-
-fun preference_raw category raw_get raw_set typ name descr =
- let
- fun get () = CRITICAL raw_get;
- fun set x = CRITICAL (fn () => raw_set x);
- val pref = {name = name, descr = descr, pgiptype = typ, get = get, set = set, default = get ()};
- in add_preference category pref end;
-
-fun preference_ref category read write typ r =
- preference_raw category (fn () => read (! r)) (fn x => r := write x) typ;
-
-fun preference_bool category = preference_ref category Markup.print_bool Markup.parse_bool pgipbool;
-fun preference_int category = preference_ref category Markup.print_int Markup.parse_int pgipint;
-fun preference_real category = preference_ref category Markup.print_real Markup.parse_real pgipfloat;
-fun preference_string category = preference_ref category I I pgipstring;
-
-
-(* system options *)
-
-fun preference_option category option_name pgip_name descr =
- let
- val typ = Options.default_typ option_name;
- val pgiptype =
- if typ = Options.boolT then pgipbool
- else if typ = Options.intT then pgipint
- else if typ = Options.realT then pgipfloat
- else pgipstring;
- in
- add_preference category
- {name = pgip_name,
- descr = descr,
- default = Options.get_default option_name,
- pgiptype = pgiptype,
- get = fn () => Options.get_default option_name,
- set = Options.put_default option_name}
- end;
-
-
-(* minimal PGIP support for <askprefs>, <haspref>, <setpref> *)
-
-local
-
-fun get_attr attrs name =
- (case Properties.get attrs name of
- SOME value => value
- | NONE => raise Fail ("Missing attribute: " ^ quote name));
-
-fun attr x y = [(x, y)] : XML.attributes;
-
-fun opt_attr _ NONE = []
- | opt_attr name (SOME value) = attr name value;
-
-val pgip_id = "dummy";
-val pgip_serial = Synchronized.counter ();
-
-fun output_pgip refid refseq content =
- XML.Elem (("pgip",
- attr "tag" "Isabelle/Isar" @
- attr "id" pgip_id @
- opt_attr "destid" refid @
- attr "class" "pg" @
- opt_attr "refid" refid @
- attr "refseq" refseq @
- attr "seq" (string_of_int (pgip_serial ()))), content)
- |> XML.string_of
- |> Output.urgent_message;
-
-
-fun invalid_pgip () = raise Fail "Invalid PGIP packet";
-
-fun haspref ({name, descr, default, pgiptype, ...}: preference) =
- XML.Elem (("haspref", [("name", name), ("descr", descr), ("default", default)]),
- [XML.Elem ((pgiptype, []), [])]);
-
-fun process_element refid refseq (XML.Elem (("askprefs", _), _)) =
- get_preferences () |> List.app (fn (category, prefs) =>
- output_pgip refid refseq
- [XML.Elem (("hasprefs", [("prefcategory", category)]), map haspref prefs)])
- | process_element _ _ (XML.Elem (("setpref", attrs), data)) =
- let
- val name =
- (case Properties.get attrs "name" of
- SOME name => name
- | NONE => invalid_pgip ());
- val value = XML.content_of data;
- in set_preference name value end
- | process_element _ _ _ = invalid_pgip ();
-
-in
-
-fun process_pgip str =
- (case XML.parse str of
- XML.Elem (("pgip", attrs), pgips) =>
- let
- val class = get_attr attrs "class";
- val dest = Properties.get attrs "destid";
- val refid = Properties.get attrs "id";
- val refseq = get_attr attrs "seq";
- val processit =
- (case dest of
- NONE => class = "pa"
- | SOME id => id = pgip_id);
- in if processit then List.app (process_element refid refseq) pgips else () end
- | _ => invalid_pgip ())
- handle Fail msg => raise Fail (msg ^ "\n" ^ str);
-
-end;
-
-
-(** messages **)
-
-(* render markup *)
-
-fun special ch = chr 1 ^ ch;
-
-local
-
-fun render_trees ts = fold render_tree ts
-and render_tree t =
- (case XML.unwrap_elem t of
- SOME (_, ts) => render_trees ts
- | NONE =>
- (case t of
- XML.Text s => Buffer.add s
- | XML.Elem ((name, props), ts) =>
- let
- val (bg, en) =
- if null ts then Markup.no_output
- else if name = Markup.stateN then (special "O" ^ "\n", "\n" ^ special "P")
- else if name = Markup.sendbackN then (special "W", special "X")
- else if name = Markup.intensifyN then (special "0", special "1")
- else if name = Markup.tfreeN then (special "C", special "A")
- else if name = Markup.tvarN then (special "D", special "A")
- else if name = Markup.freeN then (special "E", special "A")
- else if name = Markup.boundN then (special "F", special "A")
- else if name = Markup.varN then (special "G", special "A")
- else if name = Markup.skolemN then (special "H", special "A")
- else
- (case Markup.get_entity_kind (name, props) of
- SOME kind =>
- if kind = Markup.classN then (special "B", special "A")
- else Markup.no_output
- | NONE => Markup.no_output);
- in Buffer.add bg #> render_trees ts #> Buffer.add en end));
-
-in
-
-fun render text =
- Buffer.content (render_trees (YXML.parse_body text) Buffer.empty);
-
-end;
-
-
-(* hooks *)
-
-fun message bg en prfx text =
- (case render text of
- "" => ()
- | s => Output.physical_writeln (enclose bg en (prefix_lines prfx s)));
-
-fun setup_messages () =
- (Output.Private_Hooks.writeln_fn := message "" "" "";
- Output.Private_Hooks.status_fn := (fn _ => ());
- Output.Private_Hooks.report_fn := (fn _ => ());
- Output.Private_Hooks.urgent_message_fn := message (special "I") (special "J") "";
- Output.Private_Hooks.tracing_fn := message (special "I" ^ special "V") (special "J") "";
- Output.Private_Hooks.warning_fn := message (special "K") (special "L") "### ";
- Output.Private_Hooks.error_fn := (fn (_, s) => message (special "M") (special "N") "*** " s);
- Output.Private_Hooks.prompt_fn := (fn s => Output.physical_stdout (render s ^ special "S")));
-
-
-(* notification *)
-
-val emacs_notify = message (special "I") (special "J") "";
-
-fun tell_clear_goals () =
- emacs_notify "Proof General, please clear the goals buffer.";
-
-fun tell_clear_response () =
- emacs_notify "Proof General, please clear the response buffer.";
-
-fun tell_file_loaded path =
- emacs_notify ("Proof General, this file is loaded: " ^ quote (File.platform_path path));
-
-fun tell_file_retracted path =
- emacs_notify ("Proof General, you can unlock the file " ^ quote (File.platform_path path));
-
-
-
-(** theory loader **)
-
-(* fake old ThyLoad -- with new semantics *)
-
-structure ThyLoad =
-struct
- val add_path = Thy_Load.set_master_path o Path.explode;
-end;
-
-
-(* actions *)
-
-local
-
-fun trace_action action name =
- if action = Thy_Info.Update then
- List.app tell_file_loaded (Thy_Info.loaded_files name)
- else if action = Thy_Info.Remove then
- List.app tell_file_retracted (Thy_Info.loaded_files name)
- else ();
-
-in
- fun setup_thy_loader () = Thy_Info.add_hook trace_action;
- fun sync_thy_loader () = List.app (trace_action Thy_Info.Update) (Thy_Info.get_names ());
-end;
-
-
-(* get informed about files *)
-
-(*liberal low-level version*)
-val thy_name = perhaps (try (unsuffix ".thy")) o List.last o space_explode "/";
-
-val inform_file_retracted = Thy_Info.kill_thy o thy_name;
-
-fun inform_file_processed file =
- let
- val name = thy_name file;
- val _ = name = "" andalso error ("Bad file name: " ^ quote file);
- val _ =
- Thy_Info.register_thy (Toplevel.end_theory Position.none (Isar.state ()))
- handle ERROR msg =>
- (warning (cat_lines ["Failed to register theory: " ^ quote name, msg]);
- tell_file_retracted (Thy_Load.thy_path (Path.basic name)))
- val _ = Isar.init ();
- in () end;
-
-
-(* restart top-level loop (keeps most state information) *)
-
-val welcome = Output.urgent_message o Session.welcome;
-
-fun restart () =
- (sync_thy_loader ();
- tell_clear_goals ();
- tell_clear_response ();
- Isar.init ();
- welcome ());
-
-
-(** theorem dependencies **)
-
-(* thm_deps *)
-
-local
-
-fun add_proof_body (PBody {thms, ...}) =
- thms |> fold (fn (_, (name, _, _)) => name <> "" ? Symtab.update (name, ()));
-
-fun add_thm th =
- (case Thm.proof_body_of th of
- PBody {proof = PThm (_, ((name, _, _), body)), ...} =>
- if Thm.has_name_hint th andalso Thm.get_name_hint th = name
- then add_proof_body (Future.join body)
- else I
- | body => add_proof_body body);
-
-in
-
-fun thm_deps ths =
- let
- (* FIXME proper derivation names!? *)
- val names = map Thm.get_name_hint (filter Thm.has_name_hint ths);
- val deps = Symtab.keys (fold add_thm ths Symtab.empty);
- in (names, deps) end;
-
-end;
-
-
-(* report via hook *)
-
-val thm_depsN = "thm_deps";
-
-local
-
-val spaces_quote = space_implode " " o map quote;
-
-fun thm_deps_message (thms, deps) =
- emacs_notify ("Proof General, theorem dependencies of " ^ thms ^ " are " ^ deps);
-
-in
-
-fun setup_present_hook () = Toplevel.add_hook (fn _ => fn state => fn state' =>
- if print_mode_active thm_depsN andalso
- can Toplevel.theory_of state andalso Toplevel.is_theory state'
- then
- let
- val prev_facts = Global_Theory.facts_of (Toplevel.theory_of state);
- val facts = Global_Theory.facts_of (Toplevel.theory_of state');
- val (names, deps) = thm_deps (maps #2 (Facts.dest_static [prev_facts] facts));
- in
- if null names orelse null deps then ()
- else thm_deps_message (spaces_quote names, spaces_quote deps)
- end
- else ());
-
-end;
-
-
-
-(** Isar command syntax **)
-
-val _ =
- Outer_Syntax.improper_command
- (("ProofGeneral.process_pgip", Keyword.control), Position.none) "(internal)"
- (Parse.text >> (fn str => Toplevel.imperative (fn () => process_pgip str)));
-
-val _ =
- Outer_Syntax.improper_command
- (("ProofGeneral.pr", Keyword.diag), Position.none) "print state (internal)"
- (Scan.succeed (Toplevel.keep (fn state =>
- if Toplevel.is_toplevel state orelse Toplevel.is_theory state then tell_clear_goals ()
- else (Toplevel.quiet := false; Toplevel.print_state true state))));
-
-val _ = (*undo without output -- historical*)
- Outer_Syntax.improper_command
- (("ProofGeneral.undo", Keyword.control), Position.none) "(internal)"
- (Scan.succeed (Toplevel.imperative (fn () => Isar.undo 1)));
-
-val _ =
- Outer_Syntax.improper_command
- (("ProofGeneral.restart", Keyword.control), Position.none) "(internal)"
- (Parse.opt_unit >> (K (Toplevel.imperative restart)));
-
-val _ =
- Outer_Syntax.improper_command
- (("ProofGeneral.kill_proof", Keyword.control), Position.none) "(internal)"
- (Scan.succeed (Toplevel.imperative (fn () => (Isar.kill_proof (); tell_clear_goals ()))));
-
-val _ =
- Outer_Syntax.improper_command
- (("ProofGeneral.inform_file_processed", Keyword.control), Position.none) "(internal)"
- (Parse.name >> (fn file => Toplevel.imperative (fn () => inform_file_processed file)));
-
-val _ =
- Outer_Syntax.improper_command
- (("ProofGeneral.inform_file_retracted", Keyword.control), Position.none) "(internal)"
- (Parse.name >> (fn file => Toplevel.imperative (fn () => inform_file_retracted file)));
-
-
-
-(** init **)
-
-val proof_general_emacsN = "ProofGeneralEmacs";
-
-val initialized = Unsynchronized.ref false;
-
-fun init () =
- (if ! initialized then ()
- else
- (Output.add_mode Symbol.xsymbolsN Symbol.output Output.default_escape;
- Output.add_mode proof_general_emacsN
- Output.default_output Output.default_escape;
- Markup.add_mode proof_general_emacsN YXML.output_markup;
- setup_messages ();
- setup_thy_loader ();
- setup_present_hook ();
- initialized := true);
- sync_thy_loader ();
- Unsynchronized.change print_mode (update (op =) proof_general_emacsN);
- Secure.PG_setup ();
- Isar.toplevel_loop TextIO.stdIn
- {init = true, welcome = true, sync = true, secure = Secure.is_secure ()});
-
-end;
-
-
-
--- a/src/Pure/ProofGeneral/proof_general_pure.ML Wed May 15 20:28:43 2013 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,158 +0,0 @@
-(* Title: Pure/ProofGeneral/proof_general_pure.ML
- Author: David Aspinall
- Author: Makarius
-
-Proof General preferences for Isabelle/Pure.
-*)
-
-(* display *)
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_display
- Printer.show_types_default
- "show-types"
- "Include types in display of Isabelle terms";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_display
- Printer.show_sorts_default
- "show-sorts"
- "Include sorts in display of Isabelle terms";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_display
- Goal_Display.show_consts_default
- "show-consts"
- "Show types of consts in Isabelle goal display";
-
-val _ =
- ProofGeneral.preference_option ProofGeneral.category_display
- @{option names_long}
- "long-names"
- "Show fully qualified names in Isabelle terms";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_display
- Printer.show_brackets_default
- "show-brackets"
- "Show full bracketing in Isabelle terms";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_display
- Goal_Display.show_main_goal_default
- "show-main-goal"
- "Show main goal in proof state display";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_display
- Syntax_Trans.eta_contract_default
- "eta-contract"
- "Print terms eta-contracted";
-
-
-(* advanced display *)
-
-val _ =
- ProofGeneral.preference_option ProofGeneral.category_advanced_display
- @{option goals_limit}
- "goals-limit"
- "Setting for maximum number of subgoals to be printed";
-
-val _ =
- ProofGeneral.preference_raw ProofGeneral.category_advanced_display
- (Markup.print_int o get_print_depth)
- (print_depth o Markup.parse_int)
- ProofGeneral.pgipint
- "print-depth"
- "Setting for the ML print depth";
-
-val _ =
- ProofGeneral.preference_option ProofGeneral.category_advanced_display
- @{option show_question_marks}
- "show-question-marks"
- "Show leading question mark of variable name";
-
-
-(* tracing *)
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_tracing
- Raw_Simplifier.simp_trace_default
- "trace-simplifier"
- "Trace simplification rules";
-
-val _ =
- ProofGeneral.preference_int ProofGeneral.category_tracing
- Raw_Simplifier.simp_trace_depth_limit_default
- "trace-simplifier-depth"
- "Trace simplifier depth limit";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_tracing
- Pattern.trace_unify_fail
- "trace-unification"
- "Output error diagnostics during unification";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_tracing
- Toplevel.timing
- "global-timing"
- "Whether to enable timing in Isabelle";
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_tracing
- Toplevel.debug
- "debugging"
- "Whether to enable debugging";
-
-val _ =
- ProofGeneral.preference_raw ProofGeneral.category_tracing
- (fn () => Markup.print_bool (print_mode_active ProofGeneral.thm_depsN))
- (fn s =>
- if Markup.parse_bool s
- then Unsynchronized.change print_mode (insert (op =) ProofGeneral.thm_depsN)
- else Unsynchronized.change print_mode (remove (op =) ProofGeneral.thm_depsN))
- ProofGeneral.pgipbool
- "theorem-dependencies"
- "Track theorem dependencies within Proof General";
-
-
-(* proof *)
-
-val _ =
- Unsynchronized.setmp quick_and_dirty true (fn () =>
- ProofGeneral.preference_bool ProofGeneral.category_proof
- quick_and_dirty
- "quick-and-dirty"
- "Take a few short cuts") ();
-
-val _ =
- ProofGeneral.preference_bool ProofGeneral.category_proof
- Goal.skip_proofs
- "skip-proofs"
- "Skip over proofs";
-
-val _ =
- Unsynchronized.setmp Proofterm.proofs 1 (fn () =>
- ProofGeneral.preference_raw ProofGeneral.category_proof
- (Markup.print_bool o Proofterm.proofs_enabled)
- (fn s => Proofterm.proofs := (if Markup.parse_bool s then 2 else 1))
- ProofGeneral.pgipbool
- "full-proofs"
- "Record full proof objects internally") ();
-
-val _ =
- Unsynchronized.setmp Multithreading.max_threads 0 (fn () =>
- ProofGeneral.preference_int ProofGeneral.category_proof
- Multithreading.max_threads
- "max-threads"
- "Maximum number of threads") ();
-
-val _ =
- ProofGeneral.preference_raw ProofGeneral.category_proof
- (fn () => Markup.print_bool (! Goal.parallel_proofs >= 1))
- (fn s => Goal.parallel_proofs := (if Markup.parse_bool s then 1 else 0))
- ProofGeneral.pgipbool
- "parallel-proofs"
- "Check proofs in parallel";
-
--- a/src/Pure/Pure.thy Wed May 15 20:28:43 2013 +0200
+++ b/src/Pure/Pure.thy Wed May 15 20:34:42 2013 +0200
@@ -98,7 +98,7 @@
ML_file "Isar/isar_syn.ML"
ML_file "Tools/find_theorems.ML"
ML_file "Tools/find_consts.ML"
-ML_file "ProofGeneral/proof_general_pure.ML"
+ML_file "Tools/proof_general_pure.ML"
section {* Further content for the Pure theory *}
--- a/src/Pure/ROOT Wed May 15 20:28:43 2013 +0200
+++ b/src/Pure/ROOT Wed May 15 20:34:42 2013 +0200
@@ -160,7 +160,6 @@
"Proof/proof_rewrite_rules.ML"
"Proof/proof_syntax.ML"
"Proof/reconstruct.ML"
- "ProofGeneral/proof_general.ML"
"ROOT.ML"
"Syntax/ast.ML"
"Syntax/lexicon.ML"
@@ -196,6 +195,7 @@
"Tools/build.ML"
"Tools/named_thms.ML"
"Tools/legacy_xml_syntax.ML"
+ "Tools/proof_general.ML"
"assumption.ML"
"axclass.ML"
"config.ML"
--- a/src/Pure/ROOT.ML Wed May 15 20:28:43 2013 +0200
+++ b/src/Pure/ROOT.ML Wed May 15 20:34:42 2013 +0200
@@ -293,13 +293,10 @@
use "Tools/build.ML";
use "Tools/named_thms.ML";
+use "Tools/proof_general.ML";
use "Tools/legacy_xml_syntax.ML";
-(* configuration for Proof General *)
-use "ProofGeneral/proof_general.ML";
-
-
(* ML toplevel pretty printing *)
toplevel_pp ["Pretty", "T"] "(fn _: Pretty.T => Pretty.str \"<pretty>\")";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Tools/proof_general.ML Wed May 15 20:34:42 2013 +0200
@@ -0,0 +1,468 @@
+(* Title: Pure/Tools/proof_general.ML
+ Author: David Aspinall
+ Author: Makarius
+
+Isabelle/Isar configuration for Proof General / Emacs.
+See also http://proofgeneral.inf.ed.ac.uk
+*)
+
+signature PROOF_GENERAL =
+sig
+ type category = string
+ val category_display: category
+ val category_advanced_display: category
+ val category_tracing: category
+ val category_proof: category
+ type pgiptype = string
+ val pgipbool: pgiptype
+ val pgipint: pgiptype
+ val pgipfloat: pgiptype
+ val pgipstring: pgiptype
+ val preference_raw: category ->
+ (unit -> string) -> (string -> unit) -> string -> string -> string -> unit
+ val preference_bool: category -> bool Unsynchronized.ref -> string -> string -> unit
+ val preference_int: category -> int Unsynchronized.ref -> string -> string -> unit
+ val preference_real: category -> real Unsynchronized.ref -> string -> string -> unit
+ val preference_string: category -> string Unsynchronized.ref -> string -> string -> unit
+ val preference_option: category -> string -> string -> string -> unit
+ structure ThyLoad: sig val add_path: string -> unit end
+ val init: unit -> unit
+ val thm_depsN: string
+end;
+
+structure ProofGeneral: PROOF_GENERAL =
+struct
+
+(** preferences **)
+
+(* type preference *)
+
+type category = string;
+val category_display = "Display";
+val category_advanced_display = "Advanced Display";
+val category_tracing = "Tracing";
+val category_proof = "Proof";
+
+type pgiptype = string;
+val pgipbool = "pgipbool";
+val pgipint = "pgipint";
+val pgipfloat = "pgipint"; (*NB: PG 3.7.x and 4.0 lack pgipfloat, but accept floats as pgipint*)
+val pgipstring = "pgipstring";
+
+type preference =
+ {name: string,
+ descr: string,
+ default: string,
+ pgiptype: pgiptype,
+ get: unit -> string,
+ set: string -> unit};
+
+
+(* global preferences *)
+
+local
+ val preferences =
+ Synchronized.var "ProofGeneral.preferences"
+ ([(category_display, []),
+ (category_advanced_display, []),
+ (category_tracing, []),
+ (category_proof, [])]: (category * preference list) list);
+in
+
+fun add_preference category (pref: preference) =
+ Synchronized.change preferences (map (fn (cat, prefs) =>
+ if cat <> category then (cat, prefs)
+ else
+ if exists (fn {name, ...} => name = #name pref) prefs
+ then (warning ("Preference already exists: " ^ quote (#name pref)); (cat, prefs))
+ else (cat, prefs @ [pref])));
+
+fun get_preferences () = Synchronized.value preferences;
+
+end;
+
+fun set_preference name value =
+ let val prefs = map (`(#name)) (maps snd (get_preferences ())) in
+ (case AList.lookup (op =) prefs name of
+ SOME {set, ...} => set value
+ | NONE => error ("Unknown prover preference: " ^ quote name))
+ end;
+
+
+(* raw preferences *)
+
+fun preference_raw category raw_get raw_set typ name descr =
+ let
+ fun get () = CRITICAL raw_get;
+ fun set x = CRITICAL (fn () => raw_set x);
+ val pref = {name = name, descr = descr, pgiptype = typ, get = get, set = set, default = get ()};
+ in add_preference category pref end;
+
+fun preference_ref category read write typ r =
+ preference_raw category (fn () => read (! r)) (fn x => r := write x) typ;
+
+fun preference_bool category = preference_ref category Markup.print_bool Markup.parse_bool pgipbool;
+fun preference_int category = preference_ref category Markup.print_int Markup.parse_int pgipint;
+fun preference_real category = preference_ref category Markup.print_real Markup.parse_real pgipfloat;
+fun preference_string category = preference_ref category I I pgipstring;
+
+
+(* system options *)
+
+fun preference_option category option_name pgip_name descr =
+ let
+ val typ = Options.default_typ option_name;
+ val pgiptype =
+ if typ = Options.boolT then pgipbool
+ else if typ = Options.intT then pgipint
+ else if typ = Options.realT then pgipfloat
+ else pgipstring;
+ in
+ add_preference category
+ {name = pgip_name,
+ descr = descr,
+ default = Options.get_default option_name,
+ pgiptype = pgiptype,
+ get = fn () => Options.get_default option_name,
+ set = Options.put_default option_name}
+ end;
+
+
+(* minimal PGIP support for <askprefs>, <haspref>, <setpref> *)
+
+local
+
+fun get_attr attrs name =
+ (case Properties.get attrs name of
+ SOME value => value
+ | NONE => raise Fail ("Missing attribute: " ^ quote name));
+
+fun attr x y = [(x, y)] : XML.attributes;
+
+fun opt_attr _ NONE = []
+ | opt_attr name (SOME value) = attr name value;
+
+val pgip_id = "dummy";
+val pgip_serial = Synchronized.counter ();
+
+fun output_pgip refid refseq content =
+ XML.Elem (("pgip",
+ attr "tag" "Isabelle/Isar" @
+ attr "id" pgip_id @
+ opt_attr "destid" refid @
+ attr "class" "pg" @
+ opt_attr "refid" refid @
+ attr "refseq" refseq @
+ attr "seq" (string_of_int (pgip_serial ()))), content)
+ |> XML.string_of
+ |> Output.urgent_message;
+
+
+fun invalid_pgip () = raise Fail "Invalid PGIP packet";
+
+fun haspref ({name, descr, default, pgiptype, ...}: preference) =
+ XML.Elem (("haspref", [("name", name), ("descr", descr), ("default", default)]),
+ [XML.Elem ((pgiptype, []), [])]);
+
+fun process_element refid refseq (XML.Elem (("askprefs", _), _)) =
+ get_preferences () |> List.app (fn (category, prefs) =>
+ output_pgip refid refseq
+ [XML.Elem (("hasprefs", [("prefcategory", category)]), map haspref prefs)])
+ | process_element _ _ (XML.Elem (("setpref", attrs), data)) =
+ let
+ val name =
+ (case Properties.get attrs "name" of
+ SOME name => name
+ | NONE => invalid_pgip ());
+ val value = XML.content_of data;
+ in set_preference name value end
+ | process_element _ _ _ = invalid_pgip ();
+
+in
+
+fun process_pgip str =
+ (case XML.parse str of
+ XML.Elem (("pgip", attrs), pgips) =>
+ let
+ val class = get_attr attrs "class";
+ val dest = Properties.get attrs "destid";
+ val refid = Properties.get attrs "id";
+ val refseq = get_attr attrs "seq";
+ val processit =
+ (case dest of
+ NONE => class = "pa"
+ | SOME id => id = pgip_id);
+ in if processit then List.app (process_element refid refseq) pgips else () end
+ | _ => invalid_pgip ())
+ handle Fail msg => raise Fail (msg ^ "\n" ^ str);
+
+end;
+
+
+(** messages **)
+
+(* render markup *)
+
+fun special ch = chr 1 ^ ch;
+
+local
+
+fun render_trees ts = fold render_tree ts
+and render_tree t =
+ (case XML.unwrap_elem t of
+ SOME (_, ts) => render_trees ts
+ | NONE =>
+ (case t of
+ XML.Text s => Buffer.add s
+ | XML.Elem ((name, props), ts) =>
+ let
+ val (bg, en) =
+ if null ts then Markup.no_output
+ else if name = Markup.stateN then (special "O" ^ "\n", "\n" ^ special "P")
+ else if name = Markup.sendbackN then (special "W", special "X")
+ else if name = Markup.intensifyN then (special "0", special "1")
+ else if name = Markup.tfreeN then (special "C", special "A")
+ else if name = Markup.tvarN then (special "D", special "A")
+ else if name = Markup.freeN then (special "E", special "A")
+ else if name = Markup.boundN then (special "F", special "A")
+ else if name = Markup.varN then (special "G", special "A")
+ else if name = Markup.skolemN then (special "H", special "A")
+ else
+ (case Markup.get_entity_kind (name, props) of
+ SOME kind =>
+ if kind = Markup.classN then (special "B", special "A")
+ else Markup.no_output
+ | NONE => Markup.no_output);
+ in Buffer.add bg #> render_trees ts #> Buffer.add en end));
+
+in
+
+fun render text =
+ Buffer.content (render_trees (YXML.parse_body text) Buffer.empty);
+
+end;
+
+
+(* hooks *)
+
+fun message bg en prfx text =
+ (case render text of
+ "" => ()
+ | s => Output.physical_writeln (enclose bg en (prefix_lines prfx s)));
+
+fun setup_messages () =
+ (Output.Private_Hooks.writeln_fn := message "" "" "";
+ Output.Private_Hooks.status_fn := (fn _ => ());
+ Output.Private_Hooks.report_fn := (fn _ => ());
+ Output.Private_Hooks.urgent_message_fn := message (special "I") (special "J") "";
+ Output.Private_Hooks.tracing_fn := message (special "I" ^ special "V") (special "J") "";
+ Output.Private_Hooks.warning_fn := message (special "K") (special "L") "### ";
+ Output.Private_Hooks.error_fn := (fn (_, s) => message (special "M") (special "N") "*** " s);
+ Output.Private_Hooks.prompt_fn := (fn s => Output.physical_stdout (render s ^ special "S")));
+
+
+(* notification *)
+
+val emacs_notify = message (special "I") (special "J") "";
+
+fun tell_clear_goals () =
+ emacs_notify "Proof General, please clear the goals buffer.";
+
+fun tell_clear_response () =
+ emacs_notify "Proof General, please clear the response buffer.";
+
+fun tell_file_loaded path =
+ emacs_notify ("Proof General, this file is loaded: " ^ quote (File.platform_path path));
+
+fun tell_file_retracted path =
+ emacs_notify ("Proof General, you can unlock the file " ^ quote (File.platform_path path));
+
+
+
+(** theory loader **)
+
+(* fake old ThyLoad -- with new semantics *)
+
+structure ThyLoad =
+struct
+ val add_path = Thy_Load.set_master_path o Path.explode;
+end;
+
+
+(* actions *)
+
+local
+
+fun trace_action action name =
+ if action = Thy_Info.Update then
+ List.app tell_file_loaded (Thy_Info.loaded_files name)
+ else if action = Thy_Info.Remove then
+ List.app tell_file_retracted (Thy_Info.loaded_files name)
+ else ();
+
+in
+ fun setup_thy_loader () = Thy_Info.add_hook trace_action;
+ fun sync_thy_loader () = List.app (trace_action Thy_Info.Update) (Thy_Info.get_names ());
+end;
+
+
+(* get informed about files *)
+
+(*liberal low-level version*)
+val thy_name = perhaps (try (unsuffix ".thy")) o List.last o space_explode "/";
+
+val inform_file_retracted = Thy_Info.kill_thy o thy_name;
+
+fun inform_file_processed file =
+ let
+ val name = thy_name file;
+ val _ = name = "" andalso error ("Bad file name: " ^ quote file);
+ val _ =
+ Thy_Info.register_thy (Toplevel.end_theory Position.none (Isar.state ()))
+ handle ERROR msg =>
+ (warning (cat_lines ["Failed to register theory: " ^ quote name, msg]);
+ tell_file_retracted (Thy_Load.thy_path (Path.basic name)))
+ val _ = Isar.init ();
+ in () end;
+
+
+(* restart top-level loop (keeps most state information) *)
+
+val welcome = Output.urgent_message o Session.welcome;
+
+fun restart () =
+ (sync_thy_loader ();
+ tell_clear_goals ();
+ tell_clear_response ();
+ Isar.init ();
+ welcome ());
+
+
+(** theorem dependencies **)
+
+(* thm_deps *)
+
+local
+
+fun add_proof_body (PBody {thms, ...}) =
+ thms |> fold (fn (_, (name, _, _)) => name <> "" ? Symtab.update (name, ()));
+
+fun add_thm th =
+ (case Thm.proof_body_of th of
+ PBody {proof = PThm (_, ((name, _, _), body)), ...} =>
+ if Thm.has_name_hint th andalso Thm.get_name_hint th = name
+ then add_proof_body (Future.join body)
+ else I
+ | body => add_proof_body body);
+
+in
+
+fun thm_deps ths =
+ let
+ (* FIXME proper derivation names!? *)
+ val names = map Thm.get_name_hint (filter Thm.has_name_hint ths);
+ val deps = Symtab.keys (fold add_thm ths Symtab.empty);
+ in (names, deps) end;
+
+end;
+
+
+(* report via hook *)
+
+val thm_depsN = "thm_deps";
+
+local
+
+val spaces_quote = space_implode " " o map quote;
+
+fun thm_deps_message (thms, deps) =
+ emacs_notify ("Proof General, theorem dependencies of " ^ thms ^ " are " ^ deps);
+
+in
+
+fun setup_present_hook () = Toplevel.add_hook (fn _ => fn state => fn state' =>
+ if print_mode_active thm_depsN andalso
+ can Toplevel.theory_of state andalso Toplevel.is_theory state'
+ then
+ let
+ val prev_facts = Global_Theory.facts_of (Toplevel.theory_of state);
+ val facts = Global_Theory.facts_of (Toplevel.theory_of state');
+ val (names, deps) = thm_deps (maps #2 (Facts.dest_static [prev_facts] facts));
+ in
+ if null names orelse null deps then ()
+ else thm_deps_message (spaces_quote names, spaces_quote deps)
+ end
+ else ());
+
+end;
+
+
+
+(** Isar command syntax **)
+
+val _ =
+ Outer_Syntax.improper_command
+ (("ProofGeneral.process_pgip", Keyword.control), Position.none) "(internal)"
+ (Parse.text >> (fn str => Toplevel.imperative (fn () => process_pgip str)));
+
+val _ =
+ Outer_Syntax.improper_command
+ (("ProofGeneral.pr", Keyword.diag), Position.none) "print state (internal)"
+ (Scan.succeed (Toplevel.keep (fn state =>
+ if Toplevel.is_toplevel state orelse Toplevel.is_theory state then tell_clear_goals ()
+ else (Toplevel.quiet := false; Toplevel.print_state true state))));
+
+val _ = (*undo without output -- historical*)
+ Outer_Syntax.improper_command
+ (("ProofGeneral.undo", Keyword.control), Position.none) "(internal)"
+ (Scan.succeed (Toplevel.imperative (fn () => Isar.undo 1)));
+
+val _ =
+ Outer_Syntax.improper_command
+ (("ProofGeneral.restart", Keyword.control), Position.none) "(internal)"
+ (Parse.opt_unit >> (K (Toplevel.imperative restart)));
+
+val _ =
+ Outer_Syntax.improper_command
+ (("ProofGeneral.kill_proof", Keyword.control), Position.none) "(internal)"
+ (Scan.succeed (Toplevel.imperative (fn () => (Isar.kill_proof (); tell_clear_goals ()))));
+
+val _ =
+ Outer_Syntax.improper_command
+ (("ProofGeneral.inform_file_processed", Keyword.control), Position.none) "(internal)"
+ (Parse.name >> (fn file => Toplevel.imperative (fn () => inform_file_processed file)));
+
+val _ =
+ Outer_Syntax.improper_command
+ (("ProofGeneral.inform_file_retracted", Keyword.control), Position.none) "(internal)"
+ (Parse.name >> (fn file => Toplevel.imperative (fn () => inform_file_retracted file)));
+
+
+
+(** init **)
+
+val proof_general_emacsN = "ProofGeneralEmacs";
+
+val initialized = Unsynchronized.ref false;
+
+fun init () =
+ (if ! initialized then ()
+ else
+ (Output.add_mode Symbol.xsymbolsN Symbol.output Output.default_escape;
+ Output.add_mode proof_general_emacsN
+ Output.default_output Output.default_escape;
+ Markup.add_mode proof_general_emacsN YXML.output_markup;
+ setup_messages ();
+ setup_thy_loader ();
+ setup_present_hook ();
+ initialized := true);
+ sync_thy_loader ();
+ Unsynchronized.change print_mode (update (op =) proof_general_emacsN);
+ Secure.PG_setup ();
+ Isar.toplevel_loop TextIO.stdIn
+ {init = true, welcome = true, sync = true, secure = Secure.is_secure ()});
+
+end;
+
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Tools/proof_general_pure.ML Wed May 15 20:34:42 2013 +0200
@@ -0,0 +1,158 @@
+(* Title: Pure/Tools/proof_general_pure.ML
+ Author: David Aspinall
+ Author: Makarius
+
+Proof General preferences for Isabelle/Pure.
+*)
+
+(* display *)
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_display
+ Printer.show_types_default
+ "show-types"
+ "Include types in display of Isabelle terms";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_display
+ Printer.show_sorts_default
+ "show-sorts"
+ "Include sorts in display of Isabelle terms";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_display
+ Goal_Display.show_consts_default
+ "show-consts"
+ "Show types of consts in Isabelle goal display";
+
+val _ =
+ ProofGeneral.preference_option ProofGeneral.category_display
+ @{option names_long}
+ "long-names"
+ "Show fully qualified names in Isabelle terms";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_display
+ Printer.show_brackets_default
+ "show-brackets"
+ "Show full bracketing in Isabelle terms";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_display
+ Goal_Display.show_main_goal_default
+ "show-main-goal"
+ "Show main goal in proof state display";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_display
+ Syntax_Trans.eta_contract_default
+ "eta-contract"
+ "Print terms eta-contracted";
+
+
+(* advanced display *)
+
+val _ =
+ ProofGeneral.preference_option ProofGeneral.category_advanced_display
+ @{option goals_limit}
+ "goals-limit"
+ "Setting for maximum number of subgoals to be printed";
+
+val _ =
+ ProofGeneral.preference_raw ProofGeneral.category_advanced_display
+ (Markup.print_int o get_print_depth)
+ (print_depth o Markup.parse_int)
+ ProofGeneral.pgipint
+ "print-depth"
+ "Setting for the ML print depth";
+
+val _ =
+ ProofGeneral.preference_option ProofGeneral.category_advanced_display
+ @{option show_question_marks}
+ "show-question-marks"
+ "Show leading question mark of variable name";
+
+
+(* tracing *)
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_tracing
+ Raw_Simplifier.simp_trace_default
+ "trace-simplifier"
+ "Trace simplification rules";
+
+val _ =
+ ProofGeneral.preference_int ProofGeneral.category_tracing
+ Raw_Simplifier.simp_trace_depth_limit_default
+ "trace-simplifier-depth"
+ "Trace simplifier depth limit";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_tracing
+ Pattern.trace_unify_fail
+ "trace-unification"
+ "Output error diagnostics during unification";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_tracing
+ Toplevel.timing
+ "global-timing"
+ "Whether to enable timing in Isabelle";
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_tracing
+ Toplevel.debug
+ "debugging"
+ "Whether to enable debugging";
+
+val _ =
+ ProofGeneral.preference_raw ProofGeneral.category_tracing
+ (fn () => Markup.print_bool (print_mode_active ProofGeneral.thm_depsN))
+ (fn s =>
+ if Markup.parse_bool s
+ then Unsynchronized.change print_mode (insert (op =) ProofGeneral.thm_depsN)
+ else Unsynchronized.change print_mode (remove (op =) ProofGeneral.thm_depsN))
+ ProofGeneral.pgipbool
+ "theorem-dependencies"
+ "Track theorem dependencies within Proof General";
+
+
+(* proof *)
+
+val _ =
+ Unsynchronized.setmp quick_and_dirty true (fn () =>
+ ProofGeneral.preference_bool ProofGeneral.category_proof
+ quick_and_dirty
+ "quick-and-dirty"
+ "Take a few short cuts") ();
+
+val _ =
+ ProofGeneral.preference_bool ProofGeneral.category_proof
+ Goal.skip_proofs
+ "skip-proofs"
+ "Skip over proofs";
+
+val _ =
+ Unsynchronized.setmp Proofterm.proofs 1 (fn () =>
+ ProofGeneral.preference_raw ProofGeneral.category_proof
+ (Markup.print_bool o Proofterm.proofs_enabled)
+ (fn s => Proofterm.proofs := (if Markup.parse_bool s then 2 else 1))
+ ProofGeneral.pgipbool
+ "full-proofs"
+ "Record full proof objects internally") ();
+
+val _ =
+ Unsynchronized.setmp Multithreading.max_threads 0 (fn () =>
+ ProofGeneral.preference_int ProofGeneral.category_proof
+ Multithreading.max_threads
+ "max-threads"
+ "Maximum number of threads") ();
+
+val _ =
+ ProofGeneral.preference_raw ProofGeneral.category_proof
+ (fn () => Markup.print_bool (! Goal.parallel_proofs >= 1))
+ (fn s => Goal.parallel_proofs := (if Markup.parse_bool s then 1 else 0))
+ ProofGeneral.pgipbool
+ "parallel-proofs"
+ "Check proofs in parallel";
+