just one ProofGeneral module;
authorwenzelm
Wed, 15 May 2013 17:39:41 +0200
changeset 52006 9402221f77dd
parent 51996 26aecb553c74
child 52007 0b1183012a3c
just one ProofGeneral module;
src/HOL/Tools/Nitpick/nitpick_isar.ML
src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML
src/HOL/Tools/try0.ML
src/Pure/ProofGeneral/proof_general.ML
src/Pure/ProofGeneral/proof_general_emacs.ML
src/Pure/ProofGeneral/proof_general_pgip.ML
src/Pure/ROOT
src/Pure/ROOT.ML
src/Tools/quickcheck.ML
src/Tools/solve_direct.ML
src/Tools/try.ML
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML	Wed May 15 12:13:38 2013 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML	Wed May 15 17:39:41 2013 +0200
@@ -36,7 +36,7 @@
 val auto_try_max_scopes = 6
 
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
+  ProofGeneral.add_preference Preferences.category_tracing
       (Preferences.bool_pref auto "auto-nitpick" "Run Nitpick automatically.")
 
 type raw_param = string * string list
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Wed May 15 12:13:38 2013 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Wed May 15 17:39:41 2013 +0200
@@ -45,7 +45,7 @@
 val auto = Unsynchronized.ref false
 
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
+  ProofGeneral.add_preference Preferences.category_tracing
       (Preferences.bool_pref auto "auto-sledgehammer"
            "Run Sledgehammer automatically.")
 
@@ -66,13 +66,13 @@
 val timeout = Unsynchronized.ref 30
 
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_proof
+  ProofGeneral.add_preference Preferences.category_proof
       (Preferences.string_pref provers
           "Sledgehammer: Provers"
           "Default automatic provers (separated by whitespace)")
 
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_proof
+  ProofGeneral.add_preference Preferences.category_proof
       (Preferences.int_pref timeout
           "Sledgehammer: Time Limit"
           "ATPs will be interrupted after this time (in seconds)")
--- a/src/HOL/Tools/try0.ML	Wed May 15 12:13:38 2013 +0200
+++ b/src/HOL/Tools/try0.ML	Wed May 15 17:39:41 2013 +0200
@@ -27,7 +27,7 @@
 val auto = Unsynchronized.ref false
 
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
+  ProofGeneral.add_preference Preferences.category_tracing
       (Preferences.bool_pref auto "auto-try0" "Try standard proof methods.")
 
 val default_timeout = seconds 5.0
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ProofGeneral/proof_general.ML	Wed May 15 17:39:41 2013 +0200
@@ -0,0 +1,355 @@
+(*  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
+  val add_preference: string -> Preferences.preference -> unit
+  val init: unit -> unit
+  structure ThyLoad: sig val add_path: string -> unit end
+end;
+
+structure ProofGeneral: PROOF_GENERAL =
+struct
+
+(* preferences *)
+
+val preferences = Unsynchronized.ref Preferences.pure_preferences;
+
+fun add_preference cat pref =
+  CRITICAL (fn () => Unsynchronized.change preferences (Preferences.add cat pref));
+
+fun set_preference pref value =
+  (case AList.lookup (op =) (map (fn p => (#name p, p)) (maps snd (! preferences))) pref of
+    SOME {set, ...} => set value
+  | NONE => error ("Unknown prover preference: " ^ quote pref));
+
+
+(* process_pgip: 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, ...}: Preferences.preference) =
+  XML.Elem (("haspref", [("name", name), ("descr", descr), ("default", default)]),
+    [XML.Elem ((pgiptype, []), [])]);
+
+fun process_element refid refseq (XML.Elem (("askprefs", _), _)) =
+      ! 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;
+
+
+(* 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;
+
+
+(* messages *)
+
+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 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 *)
+
+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 theorem dependencies *)
+
+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 Preferences.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;
+
+
+(* additional outer syntax for Isar *)
+
+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 ()});
+
+
+(* fake old ThyLoad -- with new semantics *)
+
+structure ThyLoad =
+struct
+  val add_path = Thy_Load.set_master_path o Path.explode;
+end;
+
+end;
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Wed May 15 12:13:38 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,264 +0,0 @@
-(*  Title:      Pure/ProofGeneral/proof_general_emacs.ML
-    Author:     David Aspinall and Markus Wenzel
-
-Isabelle/Isar configuration for Emacs Proof General.
-See also http://proofgeneral.inf.ed.ac.uk
-*)
-
-signature PROOF_GENERAL =
-sig
-  val init: unit -> unit
-  structure ThyLoad: sig val add_path: string -> unit end
-end;
-
-structure ProofGeneral: PROOF_GENERAL =
-struct
-
-(* 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;
-
-
-(* messages *)
-
-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 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 *)
-
-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 theorem dependencies *)
-
-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 Preferences.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;
-
-
-(* additional outer syntax for Isar *)
-
-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 ()});
-
-
-(* fake old ThyLoad -- with new semantics *)
-
-structure ThyLoad =
-struct
-  val add_path = Thy_Load.set_master_path o Path.explode;
-end;
-
-end;
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Wed May 15 12:13:38 2013 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,105 +0,0 @@
-(*  Title:      Pure/ProofGeneral/proof_general_pgip.ML
-    Author:     David Aspinall and Markus Wenzel
-
-Rudiments of PGIP for ProofGeneral preferences.
-*)
-
-signature PROOF_GENERAL_PGIP =
-sig
-  val add_preference: string -> Preferences.preference -> unit
-end
-
-structure ProofGeneralPgip : PROOF_GENERAL_PGIP =
-struct
-
-(* output PGIP packets *)
-
-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;
-
-
-(* preferences *)
-
-val preferences = Unsynchronized.ref Preferences.pure_preferences;
-
-fun add_preference cat pref =
-  CRITICAL (fn () => Unsynchronized.change preferences (Preferences.add cat pref));
-
-fun set_preference pref value =
-  (case AList.lookup (op =) (map (fn p => (#name p, p)) (maps snd (! preferences))) pref of
-    SOME {set, ...} => set value
-  | NONE => error ("Unknown prover preference: " ^ quote pref));
-
-
-(* process_pgip *)
-
-local
-
-fun invalid_pgip () = raise Fail "Invalid PGIP packet";
-
-fun haspref ({name, descr, default, pgiptype, ...}: Preferences.preference) =
-  XML.Elem (("haspref", [("name", name), ("descr", descr), ("default", default)]),
-    [XML.Elem ((pgiptype, []), [])]);
-
-fun process_element refid refseq (XML.Elem (("askprefs", _), _)) =
-      ! 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 ();
-
-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);
-
-in
-
-val _ =
-  Outer_Syntax.improper_command
-    (("ProofGeneral.process_pgip", Keyword.control), Position.none) "(internal)"
-    (Parse.text >> (fn str => Toplevel.imperative (fn () => process_pgip str)));
-
-end;
-
-end;
--- a/src/Pure/ROOT	Wed May 15 12:13:38 2013 +0200
+++ b/src/Pure/ROOT	Wed May 15 17:39:41 2013 +0200
@@ -161,8 +161,7 @@
     "Proof/proof_syntax.ML"
     "Proof/reconstruct.ML"
     "ProofGeneral/preferences.ML"
-    "ProofGeneral/proof_general_emacs.ML"
-    "ProofGeneral/proof_general_pgip.ML"
+    "ProofGeneral/proof_general.ML"
     "ROOT.ML"
     "Syntax/ast.ML"
     "Syntax/lexicon.ML"
--- a/src/Pure/ROOT.ML	Wed May 15 12:13:38 2013 +0200
+++ b/src/Pure/ROOT.ML	Wed May 15 17:39:41 2013 +0200
@@ -303,8 +303,7 @@
   |> Unsynchronized.setmp Multithreading.max_threads 0)
   "ProofGeneral/preferences.ML";
 
-use "ProofGeneral/proof_general_pgip.ML";
-use "ProofGeneral/proof_general_emacs.ML";
+use "ProofGeneral/proof_general.ML";
 
 
 (* ML toplevel pretty printing *)
--- a/src/Tools/quickcheck.ML	Wed May 15 12:13:38 2013 +0200
+++ b/src/Tools/quickcheck.ML	Wed May 15 17:39:41 2013 +0200
@@ -97,7 +97,7 @@
 val auto = Unsynchronized.ref false;
 
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
+  ProofGeneral.add_preference Preferences.category_tracing
   (Unsynchronized.setmp auto true (fn () =>
     Preferences.bool_pref auto
       "auto-quickcheck"
--- a/src/Tools/solve_direct.ML	Wed May 15 12:13:38 2013 +0200
+++ b/src/Tools/solve_direct.ML	Wed May 15 17:39:41 2013 +0200
@@ -37,7 +37,7 @@
 val max_solutions = Unsynchronized.ref 5;
 
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
+  ProofGeneral.add_preference Preferences.category_tracing
   (Unsynchronized.setmp auto true (fn () =>
     Preferences.bool_pref auto
       "auto-solve-direct"
--- a/src/Tools/try.ML	Wed May 15 12:13:38 2013 +0200
+++ b/src/Tools/try.ML	Wed May 15 17:39:41 2013 +0200
@@ -36,7 +36,7 @@
 
 val auto_try_time_limitN = "auto-try-time-limit"
 val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
+  ProofGeneral.add_preference Preferences.category_tracing
     (Preferences.real_pref auto_time_limit
       auto_try_time_limitN "Time limit for automatically tried tools (in seconds).")