renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
authorwenzelm
Sat, 15 May 2010 23:40:00 +0200
changeset 36953 2af1ad9aa1a3
parent 36952 338c3f8229e4
child 36954 ef698bd61057
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
NEWS
src/Pure/Isar/isar_document.ML
src/Pure/Isar/isar_syn.ML
src/Pure/Isar/outer_syntax.ML
src/Pure/Proof/extraction.ML
src/Pure/ProofGeneral/proof_general_emacs.ML
src/Pure/ProofGeneral/proof_general_pgip.ML
src/Pure/System/isar.ML
src/Pure/System/session.ML
src/Pure/Thy/thy_info.ML
src/Pure/Tools/find_consts.ML
src/Pure/Tools/find_theorems.ML
src/Pure/codegen.ML
src/Pure/pure_setup.ML
--- a/NEWS	Sat May 15 23:32:15 2010 +0200
+++ b/NEWS	Sat May 15 23:40:00 2010 +0200
@@ -505,6 +505,7 @@
 * Renamed some important ML structures, while keeping the old names as
 legacy aliases for some time:
 
+  OuterSyntax   ~>  Outer_Syntax
   OuterKeyword  ~>  Keyword
   OuterParse    ~>  Parse
   SpecParse     ~>  Parse_Spec
--- a/src/Pure/Isar/isar_document.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/Isar/isar_document.ML	Sat May 15 23:40:00 2010 +0200
@@ -276,23 +276,23 @@
 (** concrete syntax **)
 
 val _ =
-  OuterSyntax.internal_command "Isar.define_command"
+  Outer_Syntax.internal_command "Isar.define_command"
     (Parse.string -- Parse.string >> (fn (id, text) =>
       Toplevel.position (Position.id_only id) o
       Toplevel.imperative (fn () =>
-        define_command id (OuterSyntax.prepare_command (Position.id id) text))));
+        define_command id (Outer_Syntax.prepare_command (Position.id id) text))));
 
 val _ =
-  OuterSyntax.internal_command "Isar.begin_document"
+  Outer_Syntax.internal_command "Isar.begin_document"
     (Parse.string -- Parse.string >> (fn (id, path) =>
       Toplevel.imperative (fn () => begin_document id (Path.explode path))));
 
 val _ =
-  OuterSyntax.internal_command "Isar.end_document"
+  Outer_Syntax.internal_command "Isar.end_document"
     (Parse.string >> (fn id => Toplevel.imperative (fn () => end_document id)));
 
 val _ =
-  OuterSyntax.internal_command "Isar.edit_document"
+  Outer_Syntax.internal_command "Isar.edit_document"
     (Parse.string -- Parse.string --
         Parse_Value.list (Parse.string -- (Parse.string >> SOME) || Parse.string >> rpair NONE)
       >> (fn ((id, new_id), edits) =>
--- a/src/Pure/Isar/isar_syn.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/Isar/isar_syn.ML	Sat May 15 23:40:00 2010 +0200
@@ -27,52 +27,52 @@
 (** init and exit **)
 
 val _ =
-  OuterSyntax.command "theory" "begin theory" (Keyword.tag_theory Keyword.thy_begin)
+  Outer_Syntax.command "theory" "begin theory" (Keyword.tag_theory Keyword.thy_begin)
     (Thy_Header.args >> (Toplevel.print oo IsarCmd.init_theory));
 
 val _ =
-  OuterSyntax.command "end" "end (local) theory" (Keyword.tag_theory Keyword.thy_end)
+  Outer_Syntax.command "end" "end (local) theory" (Keyword.tag_theory Keyword.thy_end)
     (Scan.succeed (Toplevel.exit o Toplevel.end_local_theory));
 
 
 
 (** markup commands **)
 
-val _ = OuterSyntax.markup_command ThyOutput.Markup "header" "theory header" Keyword.diag
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "header" "theory header" Keyword.diag
   (Parse.doc_source >> IsarCmd.header_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Markup "chapter" "chapter heading"
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "chapter" "chapter heading"
   Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Markup "section" "section heading"
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "section" "section heading"
   Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Markup "subsection" "subsection heading"
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "subsection" "subsection heading"
   Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
 
 val _ =
-  OuterSyntax.markup_command ThyOutput.Markup "subsubsection" "subsubsection heading"
+  Outer_Syntax.markup_command ThyOutput.Markup "subsubsection" "subsubsection heading"
   Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.MarkupEnv "text" "formal comment (theory)"
+val _ = Outer_Syntax.markup_command ThyOutput.MarkupEnv "text" "formal comment (theory)"
   Keyword.thy_decl (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Verbatim "text_raw" "raw document preparation text"
+val _ = Outer_Syntax.markup_command ThyOutput.Verbatim "text_raw" "raw document preparation text"
   Keyword.thy_decl (Parse.opt_target -- Parse.doc_source >> IsarCmd.local_theory_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Markup "sect" "formal comment (proof)"
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "sect" "formal comment (proof)"
   (Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> IsarCmd.proof_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Markup "subsect" "formal comment (proof)"
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "subsect" "formal comment (proof)"
   (Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> IsarCmd.proof_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Markup "subsubsect" "formal comment (proof)"
+val _ = Outer_Syntax.markup_command ThyOutput.Markup "subsubsect" "formal comment (proof)"
   (Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> IsarCmd.proof_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.MarkupEnv "txt" "formal comment (proof)"
+val _ = Outer_Syntax.markup_command ThyOutput.MarkupEnv "txt" "formal comment (proof)"
   (Keyword.tag_proof Keyword.prf_decl) (Parse.doc_source >> IsarCmd.proof_markup);
 
-val _ = OuterSyntax.markup_command ThyOutput.Verbatim "txt_raw"
+val _ = Outer_Syntax.markup_command ThyOutput.Verbatim "txt_raw"
   "raw document preparation text (proof)" (Keyword.tag_proof Keyword.prf_decl)
   (Parse.doc_source >> IsarCmd.proof_markup);
 
@@ -83,19 +83,19 @@
 (* classes and sorts *)
 
 val _ =
-  OuterSyntax.command "classes" "declare type classes" Keyword.thy_decl
+  Outer_Syntax.command "classes" "declare type classes" Keyword.thy_decl
     (Scan.repeat1 (Parse.binding -- Scan.optional ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") |--
         Parse.!!! (Parse.list1 Parse.xname)) [])
       >> (Toplevel.theory o fold AxClass.axiomatize_class_cmd));
 
 val _ =
-  OuterSyntax.command "classrel" "state inclusion of type classes (axiomatic!)" Keyword.thy_decl
+  Outer_Syntax.command "classrel" "state inclusion of type classes (axiomatic!)" Keyword.thy_decl
     (Parse.and_list1 (Parse.xname -- ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<")
         |-- Parse.!!! Parse.xname))
     >> (Toplevel.theory o AxClass.axiomatize_classrel_cmd));
 
 val _ =
-  OuterSyntax.local_theory "default_sort" "declare default sort for explicit type variables"
+  Outer_Syntax.local_theory "default_sort" "declare default sort for explicit type variables"
     Keyword.thy_decl
     (Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy));
 
@@ -103,40 +103,40 @@
 (* types *)
 
 val _ =
-  OuterSyntax.local_theory "typedecl" "type declaration" Keyword.thy_decl
+  Outer_Syntax.local_theory "typedecl" "type declaration" Keyword.thy_decl
     (Parse.type_args -- Parse.binding -- Parse.opt_mixfix
       >> (fn ((args, a), mx) => Typedecl.typedecl (a, map (rpair dummyS) args, mx) #> snd));
 
 val _ =
-  OuterSyntax.local_theory "types" "declare type abbreviations" Keyword.thy_decl
+  Outer_Syntax.local_theory "types" "declare type abbreviations" Keyword.thy_decl
     (Scan.repeat1
       (Parse.type_args -- Parse.binding --
         (Parse.$$$ "=" |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix')))
       >> (fold (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs)));
 
 val _ =
-  OuterSyntax.command "nonterminals" "declare types treated as grammar nonterminal symbols"
+  Outer_Syntax.command "nonterminals" "declare types treated as grammar nonterminal symbols"
     Keyword.thy_decl (Scan.repeat1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals));
 
 val _ =
-  OuterSyntax.command "arities" "state type arities (axiomatic!)" Keyword.thy_decl
+  Outer_Syntax.command "arities" "state type arities (axiomatic!)" Keyword.thy_decl
     (Scan.repeat1 Parse.arity >> (Toplevel.theory o fold AxClass.axiomatize_arity_cmd));
 
 
 (* consts and syntax *)
 
 val _ =
-  OuterSyntax.command "judgment" "declare object-logic judgment" Keyword.thy_decl
+  Outer_Syntax.command "judgment" "declare object-logic judgment" Keyword.thy_decl
     (Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd));
 
 val _ =
-  OuterSyntax.command "consts" "declare constants" Keyword.thy_decl
+  Outer_Syntax.command "consts" "declare constants" Keyword.thy_decl
     (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts));
 
 val opt_overloaded = Parse.opt_keyword "overloaded";
 
 val _ =
-  OuterSyntax.command "finalconsts" "declare constants as final" Keyword.thy_decl
+  Outer_Syntax.command "finalconsts" "declare constants as final" Keyword.thy_decl
     (opt_overloaded -- Scan.repeat1 Parse.term >> (uncurry (Toplevel.theory oo Theory.add_finals)));
 
 val mode_spec =
@@ -147,11 +147,11 @@
   Scan.optional (Parse.$$$ "(" |-- Parse.!!! (mode_spec --| Parse.$$$ ")")) Syntax.mode_default;
 
 val _ =
-  OuterSyntax.command "syntax" "declare syntactic constants" Keyword.thy_decl
+  Outer_Syntax.command "syntax" "declare syntactic constants" Keyword.thy_decl
     (opt_mode -- Scan.repeat1 Parse.const >> (Toplevel.theory o uncurry Sign.add_modesyntax));
 
 val _ =
-  OuterSyntax.command "no_syntax" "delete syntax declarations" Keyword.thy_decl
+  Outer_Syntax.command "no_syntax" "delete syntax declarations" Keyword.thy_decl
     (opt_mode -- Scan.repeat1 Parse.const >> (Toplevel.theory o uncurry Sign.del_modesyntax));
 
 
@@ -171,18 +171,18 @@
     >> (fn (left, (arr, right)) => arr (left, right));
 
 val _ =
-  OuterSyntax.command "translations" "declare syntax translation rules" Keyword.thy_decl
+  Outer_Syntax.command "translations" "declare syntax translation rules" Keyword.thy_decl
     (Scan.repeat1 trans_line >> (Toplevel.theory o Sign.add_trrules));
 
 val _ =
-  OuterSyntax.command "no_translations" "remove syntax translation rules" Keyword.thy_decl
+  Outer_Syntax.command "no_translations" "remove syntax translation rules" Keyword.thy_decl
     (Scan.repeat1 trans_line >> (Toplevel.theory o Sign.del_trrules));
 
 
 (* axioms and definitions *)
 
 val _ =
-  OuterSyntax.command "axioms" "state arbitrary propositions (axiomatic!)" Keyword.thy_decl
+  Outer_Syntax.command "axioms" "state arbitrary propositions (axiomatic!)" Keyword.thy_decl
     (Scan.repeat1 Parse_Spec.spec >>
       (Toplevel.theory o
         (IsarCmd.add_axioms o
@@ -194,7 +194,7 @@
       Parse.$$$ "overloaded" >> K (false, true)) --| Parse.$$$ ")")) (false, false);
 
 val _ =
-  OuterSyntax.command "defs" "define constants" Keyword.thy_decl
+  Outer_Syntax.command "defs" "define constants" Keyword.thy_decl
     (opt_unchecked_overloaded --
       Scan.repeat1 (Parse_Spec.thm_name ":" -- Parse.prop >> (fn ((x, y), z) => ((x, z), y)))
       >> (Toplevel.theory o IsarCmd.add_defs));
@@ -215,41 +215,41 @@
     |-- Parse.!!! (Parse.simple_fixes --| Parse.$$$ ")")) [];
 
 val _ =
-  OuterSyntax.command "constdefs" "old-style constant definition" Keyword.thy_decl
+  Outer_Syntax.command "constdefs" "old-style constant definition" Keyword.thy_decl
     (structs -- Scan.repeat1 old_constdef >> (Toplevel.theory o Constdefs.add_constdefs));
 
 
 (* constant definitions and abbreviations *)
 
 val _ =
-  OuterSyntax.local_theory "definition" "constant definition" Keyword.thy_decl
+  Outer_Syntax.local_theory "definition" "constant definition" Keyword.thy_decl
     (Parse_Spec.constdef >> (fn args => #2 o Specification.definition_cmd args));
 
 val _ =
-  OuterSyntax.local_theory "abbreviation" "constant abbreviation" Keyword.thy_decl
+  Outer_Syntax.local_theory "abbreviation" "constant abbreviation" Keyword.thy_decl
     (opt_mode -- (Scan.option Parse_Spec.constdecl -- Parse.prop)
       >> (fn (mode, args) => Specification.abbreviation_cmd mode args));
 
 val _ =
-  OuterSyntax.local_theory "type_notation" "add concrete syntax for type constructors"
+  Outer_Syntax.local_theory "type_notation" "add concrete syntax for type constructors"
     Keyword.thy_decl
     (opt_mode -- Parse.and_list1 (Parse.xname -- Parse.mixfix)
       >> (fn (mode, args) => Specification.type_notation_cmd true mode args));
 
 val _ =
-  OuterSyntax.local_theory "no_type_notation" "delete concrete syntax for type constructors"
+  Outer_Syntax.local_theory "no_type_notation" "delete concrete syntax for type constructors"
     Keyword.thy_decl
     (opt_mode -- Parse.and_list1 (Parse.xname -- Parse.mixfix)
       >> (fn (mode, args) => Specification.type_notation_cmd false mode args));
 
 val _ =
-  OuterSyntax.local_theory "notation" "add concrete syntax for constants / fixed variables"
+  Outer_Syntax.local_theory "notation" "add concrete syntax for constants / fixed variables"
     Keyword.thy_decl
     (opt_mode -- Parse.and_list1 (Parse.xname -- Parse_Spec.locale_mixfix)
       >> (fn (mode, args) => Specification.notation_cmd true mode args));
 
 val _ =
-  OuterSyntax.local_theory "no_notation" "delete concrete syntax for constants / fixed variables"
+  Outer_Syntax.local_theory "no_notation" "delete concrete syntax for constants / fixed variables"
     Keyword.thy_decl
     (opt_mode -- Parse.and_list1 (Parse.xname -- Parse_Spec.locale_mixfix)
       >> (fn (mode, args) => Specification.notation_cmd false mode args));
@@ -258,7 +258,7 @@
 (* constant specifications *)
 
 val _ =
-  OuterSyntax.command "axiomatization" "axiomatic constant specification" Keyword.thy_decl
+  Outer_Syntax.command "axiomatization" "axiomatic constant specification" Keyword.thy_decl
     (Scan.optional Parse.fixes [] --
       Scan.optional (Parse.where_ |-- Parse.!!! (Parse.and_list1 Parse_Spec.specs)) []
       >> (fn (x, y) => Toplevel.theory (#2 o Specification.axiomatization_cmd x y)));
@@ -270,13 +270,13 @@
   Parse_Spec.name_facts >> (fn args => #2 o Specification.theorems_cmd kind args);
 
 val _ =
-  OuterSyntax.local_theory "theorems" "define theorems" Keyword.thy_decl (theorems Thm.theoremK);
+  Outer_Syntax.local_theory "theorems" "define theorems" Keyword.thy_decl (theorems Thm.theoremK);
 
 val _ =
-  OuterSyntax.local_theory "lemmas" "define lemmas" Keyword.thy_decl (theorems Thm.lemmaK);
+  Outer_Syntax.local_theory "lemmas" "define lemmas" Keyword.thy_decl (theorems Thm.lemmaK);
 
 val _ =
-  OuterSyntax.local_theory "declare" "declare theorems" Keyword.thy_decl
+  Outer_Syntax.local_theory "declare" "declare theorems" Keyword.thy_decl
     (Parse.and_list1 Parse_Spec.xthms1
       >> (fn args => #2 o Specification.theorems_cmd "" [(Attrib.empty_binding, flat args)]));
 
@@ -284,15 +284,15 @@
 (* name space entry path *)
 
 val _ =
-  OuterSyntax.command "global" "disable prefixing of theory name" Keyword.thy_decl
+  Outer_Syntax.command "global" "disable prefixing of theory name" Keyword.thy_decl
     (Scan.succeed (Toplevel.theory Sign.root_path));
 
 val _ =
-  OuterSyntax.command "local" "enable prefixing of theory name" Keyword.thy_decl
+  Outer_Syntax.command "local" "enable prefixing of theory name" Keyword.thy_decl
     (Scan.succeed (Toplevel.theory Sign.local_path));
 
 fun hide_names name hide what =
-  OuterSyntax.command name ("hide " ^ what ^ " from name space") Keyword.thy_decl
+  Outer_Syntax.command name ("hide " ^ what ^ " from name space") Keyword.thy_decl
     ((Parse.opt_keyword "open" >> not) -- Scan.repeat1 Parse.xname >>
       (Toplevel.theory o uncurry hide));
 
@@ -312,55 +312,55 @@
   (Context.proof_map (ML_Env.inherit (Context.Proof (Proof.context_of prf)))) prf;
 
 val _ =
-  OuterSyntax.command "use" "ML text from file" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.command "use" "ML text from file" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.path >>
       (fn path => Toplevel.generic_theory (ThyInfo.exec_file false path #> propagate_env)));
 
 val _ =
-  OuterSyntax.command "ML" "ML text within theory or local theory"
+  Outer_Syntax.command "ML" "ML text within theory or local theory"
     (Keyword.tag_ml Keyword.thy_decl)
     (Parse.ML_source >> (fn (txt, pos) =>
       Toplevel.generic_theory
         (ML_Context.exec (fn () => ML_Context.eval true pos txt) #> propagate_env)));
 
 val _ =
-  OuterSyntax.command "ML_prf" "ML text within proof" (Keyword.tag_proof Keyword.prf_decl)
+  Outer_Syntax.command "ML_prf" "ML text within proof" (Keyword.tag_proof Keyword.prf_decl)
     (Parse.ML_source >> (fn (txt, pos) =>
       Toplevel.proof (Proof.map_context (Context.proof_map
         (ML_Context.exec (fn () => ML_Context.eval true pos txt))) #> propagate_env_prf)));
 
 val _ =
-  OuterSyntax.command "ML_val" "diagnostic ML text" (Keyword.tag_ml Keyword.diag)
+  Outer_Syntax.command "ML_val" "diagnostic ML text" (Keyword.tag_ml Keyword.diag)
     (Parse.ML_source >> IsarCmd.ml_diag true);
 
 val _ =
-  OuterSyntax.command "ML_command" "diagnostic ML text (silent)" (Keyword.tag_ml Keyword.diag)
+  Outer_Syntax.command "ML_command" "diagnostic ML text (silent)" (Keyword.tag_ml Keyword.diag)
     (Parse.ML_source >> (Toplevel.no_timing oo IsarCmd.ml_diag false));
 
 val _ =
-  OuterSyntax.command "setup" "ML theory setup" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.command "setup" "ML theory setup" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.ML_source >> (Toplevel.theory o IsarCmd.global_setup));
 
 val _ =
-  OuterSyntax.local_theory "local_setup" "ML local theory setup" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.local_theory "local_setup" "ML local theory setup" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.ML_source >> IsarCmd.local_setup);
 
 val _ =
-  OuterSyntax.command "attribute_setup" "define attribute in ML" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.command "attribute_setup" "define attribute in ML" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.position Parse.name -- Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Parse.text)
       >> (fn (name, (txt, cmt)) => Toplevel.theory (Attrib.attribute_setup name txt cmt)));
 
 val _ =
-  OuterSyntax.command "method_setup" "define proof method in ML" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.command "method_setup" "define proof method in ML" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.position Parse.name -- Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Parse.text)
       >> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt)));
 
 val _ =
-  OuterSyntax.local_theory "declaration" "generic ML declaration" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.local_theory "declaration" "generic ML declaration" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.opt_keyword "pervasive" -- Parse.ML_source >> uncurry IsarCmd.declaration);
 
 val _ =
-  OuterSyntax.local_theory "simproc_setup" "define simproc in ML" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.local_theory "simproc_setup" "define simproc in ML" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.name --
       (Parse.$$$ "(" |-- Parse.enum1 "|" Parse.term --| Parse.$$$ ")" --| Parse.$$$ "=") --
       Parse.ML_source -- Scan.optional (Parse.$$$ "identifier" |-- Scan.repeat1 Parse.xname) []
@@ -372,27 +372,27 @@
 val trfun = Parse.opt_keyword "advanced" -- Parse.ML_source;
 
 val _ =
-  OuterSyntax.command "parse_ast_translation" "install parse ast translation functions"
+  Outer_Syntax.command "parse_ast_translation" "install parse ast translation functions"
     (Keyword.tag_ml Keyword.thy_decl)
     (trfun >> (Toplevel.theory o IsarCmd.parse_ast_translation));
 
 val _ =
-  OuterSyntax.command "parse_translation" "install parse translation functions"
+  Outer_Syntax.command "parse_translation" "install parse translation functions"
     (Keyword.tag_ml Keyword.thy_decl)
     (trfun >> (Toplevel.theory o IsarCmd.parse_translation));
 
 val _ =
-  OuterSyntax.command "print_translation" "install print translation functions"
+  Outer_Syntax.command "print_translation" "install print translation functions"
     (Keyword.tag_ml Keyword.thy_decl)
     (trfun >> (Toplevel.theory o IsarCmd.print_translation));
 
 val _ =
-  OuterSyntax.command "typed_print_translation" "install typed print translation functions"
+  Outer_Syntax.command "typed_print_translation" "install typed print translation functions"
     (Keyword.tag_ml Keyword.thy_decl)
     (trfun >> (Toplevel.theory o IsarCmd.typed_print_translation));
 
 val _ =
-  OuterSyntax.command "print_ast_translation" "install print ast translation functions"
+  Outer_Syntax.command "print_ast_translation" "install print ast translation functions"
     (Keyword.tag_ml Keyword.thy_decl)
     (trfun >> (Toplevel.theory o IsarCmd.print_ast_translation));
 
@@ -400,7 +400,7 @@
 (* oracles *)
 
 val _ =
-  OuterSyntax.command "oracle" "declare oracle" (Keyword.tag_ml Keyword.thy_decl)
+  Outer_Syntax.command "oracle" "declare oracle" (Keyword.tag_ml Keyword.thy_decl)
     (Parse.position Parse.name -- (Parse.$$$ "=" |-- Parse.ML_source) >>
       (fn (x, y) => Toplevel.theory (IsarCmd.oracle x y)));
 
@@ -408,7 +408,7 @@
 (* local theories *)
 
 val _ =
-  OuterSyntax.command "context" "enter local theory context" Keyword.thy_decl
+  Outer_Syntax.command "context" "enter local theory context" Keyword.thy_decl
     (Parse.name --| Parse.begin >> (fn name =>
       Toplevel.print o Toplevel.begin_local_theory true (Theory_Target.context_cmd name)));
 
@@ -421,7 +421,7 @@
   Scan.repeat1 Parse_Spec.context_element >> pair ([], []);
 
 val _ =
-  OuterSyntax.command "locale" "define named proof context" Keyword.thy_decl
+  Outer_Syntax.command "locale" "define named proof context" Keyword.thy_decl
     (Parse.binding --
       Scan.optional (Parse.$$$ "=" |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin
       >> (fn ((name, (expr, elems)), begin) =>
@@ -429,7 +429,7 @@
             (Expression.add_locale_cmd name Binding.empty expr elems #> snd)));
 
 val _ =
-  OuterSyntax.command "sublocale"
+  Outer_Syntax.command "sublocale"
     "prove sublocale relation between a locale and a locale expression" Keyword.thy_goal
     (Parse.xname --| (Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") --
       Parse.!!! (Parse_Spec.locale_expression false)
@@ -437,7 +437,7 @@
           Toplevel.print o Toplevel.theory_to_proof (Expression.sublocale_cmd loc expr)));
 
 val _ =
-  OuterSyntax.command "interpretation"
+  Outer_Syntax.command "interpretation"
     "prove interpretation of locale expression in theory" Keyword.thy_goal
     (Parse.!!! (Parse_Spec.locale_expression true) --
       Scan.optional (Parse.where_ |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []
@@ -445,7 +445,7 @@
           Toplevel.theory_to_proof (Expression.interpretation_cmd expr equations)));
 
 val _ =
-  OuterSyntax.command "interpret"
+  Outer_Syntax.command "interpret"
     "prove interpretation of locale expression in proof context"
     (Keyword.tag_proof Keyword.prf_goal)
     (Parse.!!! (Parse_Spec.locale_expression true)
@@ -460,24 +460,24 @@
   Scan.repeat1 Parse_Spec.context_element >> pair [];
 
 val _ =
-  OuterSyntax.command "class" "define type class" Keyword.thy_decl
+  Outer_Syntax.command "class" "define type class" Keyword.thy_decl
    (Parse.binding -- Scan.optional (Parse.$$$ "=" |-- class_val) ([], []) -- Parse.opt_begin
     >> (fn ((name, (supclasses, elems)), begin) =>
         (begin ? Toplevel.print) o Toplevel.begin_local_theory begin
           (Class.class_cmd name supclasses elems #> snd)));
 
 val _ =
-  OuterSyntax.local_theory_to_proof "subclass" "prove a subclass relation" Keyword.thy_goal
+  Outer_Syntax.local_theory_to_proof "subclass" "prove a subclass relation" Keyword.thy_goal
     (Parse.xname >> Class.subclass_cmd);
 
 val _ =
-  OuterSyntax.command "instantiation" "instantiate and prove type arity" Keyword.thy_decl
+  Outer_Syntax.command "instantiation" "instantiate and prove type arity" Keyword.thy_decl
    (Parse.multi_arity --| Parse.begin
      >> (fn arities => Toplevel.print o
          Toplevel.begin_local_theory true (Theory_Target.instantiation_cmd arities)));
 
 val _ =
-  OuterSyntax.command "instance" "prove type arity or subclass relation" Keyword.thy_goal
+  Outer_Syntax.command "instance" "prove type arity or subclass relation" Keyword.thy_goal
   ((Parse.xname -- ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") |-- Parse.!!! Parse.xname)
         >> Class.classrel_cmd ||
     Parse.multi_arity >> Class.instance_arity_cmd)
@@ -489,7 +489,7 @@
 (* arbitrary overloading *)
 
 val _ =
-  OuterSyntax.command "overloading" "overloaded definitions" Keyword.thy_decl
+  Outer_Syntax.command "overloading" "overloaded definitions" Keyword.thy_decl
    (Scan.repeat1 (Parse.name --| (Parse.$$$ "\\<equiv>" || Parse.$$$ "==") -- Parse.term --
       Scan.optional (Parse.$$$ "(" |-- (Parse.$$$ "unchecked" >> K false) --| Parse.$$$ ")") true
       >> Parse.triple1) --| Parse.begin
@@ -500,7 +500,7 @@
 (* code generation *)
 
 val _ =
-  OuterSyntax.command "code_datatype" "define set of code datatype constructors" Keyword.thy_decl
+  Outer_Syntax.command "code_datatype" "define set of code datatype constructors" Keyword.thy_decl
     (Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd));
 
 
@@ -510,7 +510,7 @@
 (* statements *)
 
 fun gen_theorem schematic kind =
-  OuterSyntax.local_theory_to_proof'
+  Outer_Syntax.local_theory_to_proof'
     (if schematic then "schematic_" ^ kind else kind)
     ("state " ^ (if schematic then "schematic " ^ kind else kind))
     (if schematic then Keyword.thy_schematic_goal else Keyword.thy_goal)
@@ -528,29 +528,29 @@
 val _ = gen_theorem true Thm.corollaryK;
 
 val _ =
-  OuterSyntax.local_theory_to_proof "example_proof"
+  Outer_Syntax.local_theory_to_proof "example_proof"
     "example proof body, without any result" Keyword.thy_schematic_goal
     (Scan.succeed
       (Specification.schematic_theorem_cmd "" NONE (K I)
         Attrib.empty_binding [] (Element.Shows []) false #> Proof.enter_forward));
 
 val _ =
-  OuterSyntax.command "have" "state local goal"
+  Outer_Syntax.command "have" "state local goal"
     (Keyword.tag_proof Keyword.prf_goal)
     (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.have));
 
 val _ =
-  OuterSyntax.command "hence" "abbreviates \"then have\""
+  Outer_Syntax.command "hence" "abbreviates \"then have\""
     (Keyword.tag_proof Keyword.prf_goal)
     (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.hence));
 
 val _ =
-  OuterSyntax.command "show" "state local goal, solving current obligation"
+  Outer_Syntax.command "show" "state local goal, solving current obligation"
     (Keyword.tag_proof Keyword.prf_asm_goal)
     (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.show));
 
 val _ =
-  OuterSyntax.command "thus" "abbreviates \"then show\""
+  Outer_Syntax.command "thus" "abbreviates \"then show\""
     (Keyword.tag_proof Keyword.prf_asm_goal)
     (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.thus));
 
@@ -560,32 +560,32 @@
 val facts = Parse.and_list1 Parse_Spec.xthms1;
 
 val _ =
-  OuterSyntax.command "then" "forward chaining"
+  Outer_Syntax.command "then" "forward chaining"
     (Keyword.tag_proof Keyword.prf_chain)
     (Scan.succeed (Toplevel.print o Toplevel.proof Proof.chain));
 
 val _ =
-  OuterSyntax.command "from" "forward chaining from given facts"
+  Outer_Syntax.command "from" "forward chaining from given facts"
     (Keyword.tag_proof Keyword.prf_chain)
     (facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss_cmd)));
 
 val _ =
-  OuterSyntax.command "with" "forward chaining from given and current facts"
+  Outer_Syntax.command "with" "forward chaining from given and current facts"
     (Keyword.tag_proof Keyword.prf_chain)
     (facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss_cmd)));
 
 val _ =
-  OuterSyntax.command "note" "define facts"
+  Outer_Syntax.command "note" "define facts"
     (Keyword.tag_proof Keyword.prf_decl)
     (Parse_Spec.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss_cmd)));
 
 val _ =
-  OuterSyntax.command "using" "augment goal facts"
+  Outer_Syntax.command "using" "augment goal facts"
     (Keyword.tag_proof Keyword.prf_decl)
     (facts >> (Toplevel.print oo (Toplevel.proof o Proof.using_cmd)));
 
 val _ =
-  OuterSyntax.command "unfolding" "unfold definitions in goal and facts"
+  Outer_Syntax.command "unfolding" "unfold definitions in goal and facts"
     (Keyword.tag_proof Keyword.prf_decl)
     (facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding_cmd)));
 
@@ -593,22 +593,22 @@
 (* proof context *)
 
 val _ =
-  OuterSyntax.command "fix" "fix local variables (Skolem constants)"
+  Outer_Syntax.command "fix" "fix local variables (Skolem constants)"
     (Keyword.tag_proof Keyword.prf_asm)
     (Parse.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix_cmd)));
 
 val _ =
-  OuterSyntax.command "assume" "assume propositions"
+  Outer_Syntax.command "assume" "assume propositions"
     (Keyword.tag_proof Keyword.prf_asm)
     (Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume_cmd)));
 
 val _ =
-  OuterSyntax.command "presume" "assume propositions, to be established later"
+  Outer_Syntax.command "presume" "assume propositions, to be established later"
     (Keyword.tag_proof Keyword.prf_asm)
     (Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume_cmd)));
 
 val _ =
-  OuterSyntax.command "def" "local definition"
+  Outer_Syntax.command "def" "local definition"
     (Keyword.tag_proof Keyword.prf_asm)
     (Parse.and_list1
       (Parse_Spec.opt_thm_name ":" --
@@ -617,24 +617,24 @@
     >> (Toplevel.print oo (Toplevel.proof o Proof.def_cmd)));
 
 val _ =
-  OuterSyntax.command "obtain" "generalized existence"
+  Outer_Syntax.command "obtain" "generalized existence"
     (Keyword.tag_proof Keyword.prf_asm_goal)
     (Parse.parname -- Scan.optional (Parse.fixes --| Parse.where_) [] -- Parse_Spec.statement
       >> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain_cmd x y z)));
 
 val _ =
-  OuterSyntax.command "guess" "wild guessing (unstructured)"
+  Outer_Syntax.command "guess" "wild guessing (unstructured)"
     (Keyword.tag_proof Keyword.prf_asm_goal)
     (Scan.optional Parse.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess_cmd)));
 
 val _ =
-  OuterSyntax.command "let" "bind text variables"
+  Outer_Syntax.command "let" "bind text variables"
     (Keyword.tag_proof Keyword.prf_decl)
     (Parse.and_list1 (Parse.and_list1 Parse.term -- (Parse.$$$ "=" |-- Parse.term))
       >> (Toplevel.print oo (Toplevel.proof o Proof.let_bind_cmd)));
 
 val _ =
-  OuterSyntax.command "write" "add concrete syntax for constants / fixed variables"
+  Outer_Syntax.command "write" "add concrete syntax for constants / fixed variables"
     (Keyword.tag_proof Keyword.prf_decl)
     (opt_mode -- Parse.and_list1 (Parse.xname -- Parse_Spec.locale_mixfix)
     >> (fn (mode, args) => Toplevel.print o Toplevel.proof (Proof.write_cmd mode args)));
@@ -645,7 +645,7 @@
     Parse.xname >> rpair []) -- Parse_Spec.opt_attribs >> Parse.triple1;
 
 val _ =
-  OuterSyntax.command "case" "invoke local context"
+  Outer_Syntax.command "case" "invoke local context"
     (Keyword.tag_proof Keyword.prf_asm)
     (case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case_cmd)));
 
@@ -653,17 +653,17 @@
 (* proof structure *)
 
 val _ =
-  OuterSyntax.command "{" "begin explicit proof block"
+  Outer_Syntax.command "{" "begin explicit proof block"
     (Keyword.tag_proof Keyword.prf_open)
     (Scan.succeed (Toplevel.print o Toplevel.proof Proof.begin_block));
 
 val _ =
-  OuterSyntax.command "}" "end explicit proof block"
+  Outer_Syntax.command "}" "end explicit proof block"
     (Keyword.tag_proof Keyword.prf_close)
     (Scan.succeed (Toplevel.print o Toplevel.proof Proof.end_block));
 
 val _ =
-  OuterSyntax.command "next" "enter next proof block"
+  Outer_Syntax.command "next" "enter next proof block"
     (Keyword.tag_proof Keyword.prf_block)
     (Scan.succeed (Toplevel.print o Toplevel.proof Proof.next_block));
 
@@ -671,37 +671,37 @@
 (* end proof *)
 
 val _ =
-  OuterSyntax.command "qed" "conclude (sub-)proof"
+  Outer_Syntax.command "qed" "conclude (sub-)proof"
     (Keyword.tag_proof Keyword.qed_block)
     (Scan.option Method.parse >> IsarCmd.qed);
 
 val _ =
-  OuterSyntax.command "by" "terminal backward proof"
+  Outer_Syntax.command "by" "terminal backward proof"
     (Keyword.tag_proof Keyword.qed)
     (Method.parse -- Scan.option Method.parse >> IsarCmd.terminal_proof);
 
 val _ =
-  OuterSyntax.command ".." "default proof"
+  Outer_Syntax.command ".." "default proof"
     (Keyword.tag_proof Keyword.qed)
     (Scan.succeed IsarCmd.default_proof);
 
 val _ =
-  OuterSyntax.command "." "immediate proof"
+  Outer_Syntax.command "." "immediate proof"
     (Keyword.tag_proof Keyword.qed)
     (Scan.succeed IsarCmd.immediate_proof);
 
 val _ =
-  OuterSyntax.command "done" "done proof"
+  Outer_Syntax.command "done" "done proof"
     (Keyword.tag_proof Keyword.qed)
     (Scan.succeed IsarCmd.done_proof);
 
 val _ =
-  OuterSyntax.command "sorry" "skip proof (quick-and-dirty mode only!)"
+  Outer_Syntax.command "sorry" "skip proof (quick-and-dirty mode only!)"
     (Keyword.tag_proof Keyword.qed)
     (Scan.succeed IsarCmd.skip_proof);
 
 val _ =
-  OuterSyntax.command "oops" "forget proof"
+  Outer_Syntax.command "oops" "forget proof"
     (Keyword.tag_proof Keyword.qed_global)
     (Scan.succeed Toplevel.forget_proof);
 
@@ -709,27 +709,27 @@
 (* proof steps *)
 
 val _ =
-  OuterSyntax.command "defer" "shuffle internal proof state"
+  Outer_Syntax.command "defer" "shuffle internal proof state"
     (Keyword.tag_proof Keyword.prf_script)
     (Scan.option Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.defer)));
 
 val _ =
-  OuterSyntax.command "prefer" "shuffle internal proof state"
+  Outer_Syntax.command "prefer" "shuffle internal proof state"
     (Keyword.tag_proof Keyword.prf_script)
     (Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.prefer)));
 
 val _ =
-  OuterSyntax.command "apply" "initial refinement step (unstructured)"
+  Outer_Syntax.command "apply" "initial refinement step (unstructured)"
     (Keyword.tag_proof Keyword.prf_script)
     (Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply)));
 
 val _ =
-  OuterSyntax.command "apply_end" "terminal refinement (unstructured)"
+  Outer_Syntax.command "apply_end" "terminal refinement (unstructured)"
     (Keyword.tag_proof Keyword.prf_script)
     (Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply_end)));
 
 val _ =
-  OuterSyntax.command "proof" "backward proof"
+  Outer_Syntax.command "proof" "backward proof"
     (Keyword.tag_proof Keyword.prf_block)
     (Scan.option Method.parse >> (fn m => Toplevel.print o
       Toplevel.actual_proof (Proof_Node.applys (Proof.proof m)) o
@@ -742,22 +742,22 @@
   Scan.option (Parse.$$$ "(" |-- Parse.!!! ((Parse_Spec.xthms1 --| Parse.$$$ ")")));
 
 val _ =
-  OuterSyntax.command "also" "combine calculation and current facts"
+  Outer_Syntax.command "also" "combine calculation and current facts"
     (Keyword.tag_proof Keyword.prf_decl)
     (calc_args >> (Toplevel.proofs' o Calculation.also_cmd));
 
 val _ =
-  OuterSyntax.command "finally" "combine calculation and current facts, exhibit result"
+  Outer_Syntax.command "finally" "combine calculation and current facts, exhibit result"
     (Keyword.tag_proof Keyword.prf_chain)
     (calc_args >> (Toplevel.proofs' o Calculation.finally_cmd));
 
 val _ =
-  OuterSyntax.command "moreover" "augment calculation by current facts"
+  Outer_Syntax.command "moreover" "augment calculation by current facts"
     (Keyword.tag_proof Keyword.prf_decl)
     (Scan.succeed (Toplevel.proof' Calculation.moreover));
 
 val _ =
-  OuterSyntax.command "ultimately" "augment calculation by current facts, exhibit result"
+  Outer_Syntax.command "ultimately" "augment calculation by current facts, exhibit result"
     (Keyword.tag_proof Keyword.prf_chain)
     (Scan.succeed (Toplevel.proof' Calculation.ultimately));
 
@@ -765,7 +765,7 @@
 (* proof navigation *)
 
 val _ =
-  OuterSyntax.command "back" "backtracking of proof command"
+  Outer_Syntax.command "back" "backtracking of proof command"
     (Keyword.tag_proof Keyword.prf_script)
     (Scan.succeed (Toplevel.print o Toplevel.actual_proof Proof_Node.back o Toplevel.skip_proof I));
 
@@ -778,9 +778,9 @@
       (Position.of_properties (Position.default_properties pos props), str));
 
 val _ =
-  OuterSyntax.improper_command "Isabelle.command" "nested Isabelle command" Keyword.control
+  Outer_Syntax.improper_command "Isabelle.command" "nested Isabelle command" Keyword.control
     (props_text :|-- (fn (pos, str) =>
-      (case OuterSyntax.parse pos str of
+      (case Outer_Syntax.parse pos str of
         [tr] => Scan.succeed (K tr)
       | _ => Scan.fail_with (K "exactly one command expected"))
       handle ERROR msg => Scan.fail_with (K msg)));
@@ -795,145 +795,145 @@
 val opt_bang = Scan.optional (Parse.$$$ "!" >> K true) false;
 
 val _ =
-  OuterSyntax.improper_command "pretty_setmargin" "change default margin for pretty printing"
+  Outer_Syntax.improper_command "pretty_setmargin" "change default margin for pretty printing"
     Keyword.diag (Parse.nat >> (Toplevel.no_timing oo IsarCmd.pretty_setmargin));
 
 val _ =
-  OuterSyntax.improper_command "help" "print outer syntax commands" Keyword.diag
-    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative OuterSyntax.print_outer_syntax));
+  Outer_Syntax.improper_command "help" "print outer syntax commands" Keyword.diag
+    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax));
 
 val _ =
-  OuterSyntax.improper_command "print_commands" "print outer syntax commands" Keyword.diag
-    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative OuterSyntax.print_outer_syntax));
+  Outer_Syntax.improper_command "print_commands" "print outer syntax commands" Keyword.diag
+    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax));
 
 val _ =
-  OuterSyntax.improper_command "print_configs" "print configuration options" Keyword.diag
+  Outer_Syntax.improper_command "print_configs" "print configuration options" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_configs));
 
 val _ =
-  OuterSyntax.improper_command "print_context" "print theory context name" Keyword.diag
+  Outer_Syntax.improper_command "print_context" "print theory context name" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_context));
 
 val _ =
-  OuterSyntax.improper_command "print_theory" "print logical theory contents (verbose!)"
+  Outer_Syntax.improper_command "print_theory" "print logical theory contents (verbose!)"
     Keyword.diag (opt_bang >> (Toplevel.no_timing oo IsarCmd.print_theory));
 
 val _ =
-  OuterSyntax.improper_command "print_syntax" "print inner syntax of context (verbose!)"
+  Outer_Syntax.improper_command "print_syntax" "print inner syntax of context (verbose!)"
     Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.print_syntax));
 
 val _ =
-  OuterSyntax.improper_command "print_abbrevs" "print constant abbreviation of context"
+  Outer_Syntax.improper_command "print_abbrevs" "print constant abbreviation of context"
     Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.print_abbrevs));
 
 val _ =
-  OuterSyntax.improper_command "print_theorems"
+  Outer_Syntax.improper_command "print_theorems"
       "print theorems of local theory or proof context" Keyword.diag
     (opt_bang >> (Toplevel.no_timing oo IsarCmd.print_theorems));
 
 val _ =
-  OuterSyntax.improper_command "print_locales" "print locales of this theory" Keyword.diag
+  Outer_Syntax.improper_command "print_locales" "print locales of this theory" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_locales));
 
 val _ =
-  OuterSyntax.improper_command "print_classes" "print classes of this theory" Keyword.diag
+  Outer_Syntax.improper_command "print_classes" "print classes of this theory" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory
       o Toplevel.keep (Class.print_classes o Toplevel.theory_of)));
 
 val _ =
-  OuterSyntax.improper_command "print_locale" "print locale of this theory" Keyword.diag
+  Outer_Syntax.improper_command "print_locale" "print locale of this theory" Keyword.diag
     (opt_bang -- Parse.xname >> (Toplevel.no_timing oo IsarCmd.print_locale));
 
 val _ =
-  OuterSyntax.improper_command "print_interps"
+  Outer_Syntax.improper_command "print_interps"
     "print interpretations of locale for this theory" Keyword.diag
     (Parse.xname >> (Toplevel.no_timing oo IsarCmd.print_registrations));
 
 val _ =
-  OuterSyntax.improper_command "print_attributes" "print attributes of this theory" Keyword.diag
+  Outer_Syntax.improper_command "print_attributes" "print attributes of this theory" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_attributes));
 
 val _ =
-  OuterSyntax.improper_command "print_simpset" "print context of Simplifier" Keyword.diag
+  Outer_Syntax.improper_command "print_simpset" "print context of Simplifier" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_simpset));
 
 val _ =
-  OuterSyntax.improper_command "print_rules" "print intro/elim rules" Keyword.diag
+  Outer_Syntax.improper_command "print_rules" "print intro/elim rules" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_rules));
 
 val _ =
-  OuterSyntax.improper_command "print_trans_rules" "print transitivity rules" Keyword.diag
+  Outer_Syntax.improper_command "print_trans_rules" "print transitivity rules" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_trans_rules));
 
 val _ =
-  OuterSyntax.improper_command "print_methods" "print methods of this theory" Keyword.diag
+  Outer_Syntax.improper_command "print_methods" "print methods of this theory" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_methods));
 
 val _ =
-  OuterSyntax.improper_command "print_antiquotations" "print antiquotations (global)" Keyword.diag
+  Outer_Syntax.improper_command "print_antiquotations" "print antiquotations (global)" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_antiquotations));
 
 val _ =
-  OuterSyntax.improper_command "thy_deps" "visualize theory dependencies"
+  Outer_Syntax.improper_command "thy_deps" "visualize theory dependencies"
     Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.thy_deps));
 
 val _ =
-  OuterSyntax.improper_command "class_deps" "visualize class dependencies"
+  Outer_Syntax.improper_command "class_deps" "visualize class dependencies"
     Keyword.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.class_deps));
 
 val _ =
-  OuterSyntax.improper_command "thm_deps" "visualize theorem dependencies"
+  Outer_Syntax.improper_command "thm_deps" "visualize theorem dependencies"
     Keyword.diag (Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.thm_deps));
 
 val _ =
-  OuterSyntax.improper_command "print_binds" "print term bindings of proof context" Keyword.diag
+  Outer_Syntax.improper_command "print_binds" "print term bindings of proof context" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_binds));
 
 val _ =
-  OuterSyntax.improper_command "print_facts" "print facts of proof context" Keyword.diag
+  Outer_Syntax.improper_command "print_facts" "print facts of proof context" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_facts));
 
 val _ =
-  OuterSyntax.improper_command "print_cases" "print cases of proof context" Keyword.diag
+  Outer_Syntax.improper_command "print_cases" "print cases of proof context" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.print_cases));
 
 val _ =
-  OuterSyntax.improper_command "print_statement" "print theorems as long statements" Keyword.diag
+  Outer_Syntax.improper_command "print_statement" "print theorems as long statements" Keyword.diag
     (opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_stmts));
 
 val _ =
-  OuterSyntax.improper_command "thm" "print theorems" Keyword.diag
+  Outer_Syntax.improper_command "thm" "print theorems" Keyword.diag
     (opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_thms));
 
 val _ =
-  OuterSyntax.improper_command "prf" "print proof terms of theorems" Keyword.diag
+  Outer_Syntax.improper_command "prf" "print proof terms of theorems" Keyword.diag
     (opt_modes -- Scan.option Parse_Spec.xthms1
       >> (Toplevel.no_timing oo IsarCmd.print_prfs false));
 
 val _ =
-  OuterSyntax.improper_command "full_prf" "print full proof terms of theorems" Keyword.diag
+  Outer_Syntax.improper_command "full_prf" "print full proof terms of theorems" Keyword.diag
     (opt_modes -- Scan.option Parse_Spec.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_prfs true));
 
 val _ =
-  OuterSyntax.improper_command "prop" "read and print proposition" Keyword.diag
+  Outer_Syntax.improper_command "prop" "read and print proposition" Keyword.diag
     (opt_modes -- Parse.term >> (Toplevel.no_timing oo IsarCmd.print_prop));
 
 val _ =
-  OuterSyntax.improper_command "term" "read and print term" Keyword.diag
+  Outer_Syntax.improper_command "term" "read and print term" Keyword.diag
     (opt_modes -- Parse.term >> (Toplevel.no_timing oo IsarCmd.print_term));
 
 val _ =
-  OuterSyntax.improper_command "typ" "read and print type" Keyword.diag
+  Outer_Syntax.improper_command "typ" "read and print type" Keyword.diag
     (opt_modes -- Parse.typ >> (Toplevel.no_timing oo IsarCmd.print_type));
 
 val _ =
-  OuterSyntax.improper_command "print_codesetup" "print code generator setup" Keyword.diag
+  Outer_Syntax.improper_command "print_codesetup" "print code generator setup" Keyword.diag
     (Scan.succeed
       (Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep
         (Code.print_codesetup o Toplevel.theory_of)));
 
 val _ =
-  OuterSyntax.improper_command "unused_thms" "find unused theorems" Keyword.diag
+  Outer_Syntax.improper_command "unused_thms" "find unused theorems" Keyword.diag
     (Scan.option ((Scan.repeat1 (Scan.unless Parse.minus Parse.name) --| Parse.minus) --
        Scan.option (Scan.repeat1 (Scan.unless Parse.minus Parse.name))) >>
          (Toplevel.no_timing oo IsarCmd.unused_thms));
@@ -943,66 +943,66 @@
 (** system commands (for interactive mode only) **)
 
 val _ =
-  OuterSyntax.improper_command "cd" "change current working directory" Keyword.diag
+  Outer_Syntax.improper_command "cd" "change current working directory" Keyword.diag
     (Parse.path >> (Toplevel.no_timing oo IsarCmd.cd));
 
 val _ =
-  OuterSyntax.improper_command "pwd" "print current working directory" Keyword.diag
+  Outer_Syntax.improper_command "pwd" "print current working directory" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.pwd));
 
 val _ =
-  OuterSyntax.improper_command "use_thy" "use theory file" Keyword.diag
+  Outer_Syntax.improper_command "use_thy" "use theory file" Keyword.diag
     (Parse.name >> (fn name =>
       Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.use_thy name)));
 
 val _ =
-  OuterSyntax.improper_command "touch_thy" "outdate theory, including descendants" Keyword.diag
+  Outer_Syntax.improper_command "touch_thy" "outdate theory, including descendants" Keyword.diag
     (Parse.name >> (fn name =>
       Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.touch_thy name)));
 
 val _ =
-  OuterSyntax.improper_command "remove_thy" "remove theory from loader database" Keyword.diag
+  Outer_Syntax.improper_command "remove_thy" "remove theory from loader database" Keyword.diag
     (Parse.name >> (fn name =>
       Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.remove_thy name)));
 
 val _ =
-  OuterSyntax.improper_command "kill_thy" "kill theory -- try to remove from loader database"
+  Outer_Syntax.improper_command "kill_thy" "kill theory -- try to remove from loader database"
     Keyword.diag (Parse.name >> (fn name =>
       Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.kill_thy name)));
 
 val _ =
-  OuterSyntax.improper_command "display_drafts" "display raw source files with symbols"
+  Outer_Syntax.improper_command "display_drafts" "display raw source files with symbols"
     Keyword.diag (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo IsarCmd.display_drafts));
 
 val _ =
-  OuterSyntax.improper_command "print_drafts" "print raw source files with symbols"
+  Outer_Syntax.improper_command "print_drafts" "print raw source files with symbols"
     Keyword.diag (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo IsarCmd.print_drafts));
 
 val opt_limits =
   Scan.option Parse.nat -- Scan.option (Parse.$$$ "," |-- Parse.!!! Parse.nat);
 
 val _ =
-  OuterSyntax.improper_command "pr" "print current proof state (if present)" Keyword.diag
+  Outer_Syntax.improper_command "pr" "print current proof state (if present)" Keyword.diag
     (opt_modes -- opt_limits >> (Toplevel.no_timing oo IsarCmd.pr));
 
 val _ =
-  OuterSyntax.improper_command "disable_pr" "disable printing of toplevel state" Keyword.diag
+  Outer_Syntax.improper_command "disable_pr" "disable printing of toplevel state" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.disable_pr));
 
 val _ =
-  OuterSyntax.improper_command "enable_pr" "enable printing of toplevel state" Keyword.diag
+  Outer_Syntax.improper_command "enable_pr" "enable printing of toplevel state" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o IsarCmd.enable_pr));
 
 val _ =
-  OuterSyntax.improper_command "commit" "commit current session to ML database" Keyword.diag
+  Outer_Syntax.improper_command "commit" "commit current session to ML database" Keyword.diag
     (Parse.opt_unit >> K (Toplevel.no_timing o Toplevel.imperative Secure.commit));
 
 val _ =
-  OuterSyntax.improper_command "quit" "quit Isabelle" Keyword.control
+  Outer_Syntax.improper_command "quit" "quit Isabelle" Keyword.control
     (Parse.opt_unit >> (Toplevel.no_timing oo K IsarCmd.quit));
 
 val _ =
-  OuterSyntax.improper_command "exit" "exit Isar loop" Keyword.control
+  Outer_Syntax.improper_command "exit" "exit Isar loop" Keyword.control
     (Scan.succeed (Toplevel.no_timing o IsarCmd.exit));
 
 end;
--- a/src/Pure/Isar/outer_syntax.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/Isar/outer_syntax.ML	Sat May 15 23:40:00 2010 +0200
@@ -34,7 +34,7 @@
   val load_thy: string -> Position.T -> string list -> bool -> unit -> unit
 end;
 
-structure OuterSyntax: OUTER_SYNTAX =
+structure Outer_Syntax: OUTER_SYNTAX =
 struct
 
 structure T = OuterLex;
@@ -296,3 +296,7 @@
   in after_load end;
 
 end;
+
+(*legacy alias*)
+structure OuterSyntax = Outer_Syntax;
+
--- a/src/Pure/Proof/extraction.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/Proof/extraction.ML	Sat May 15 23:40:00 2010 +0200
@@ -753,7 +753,7 @@
 val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") [];
 
 val _ =
-  OuterSyntax.command "realizers"
+  Outer_Syntax.command "realizers"
   "specify realizers for primitive axioms / theorems, together with correctness proof"
   Keyword.thy_decl
     (Scan.repeat1 (Parse.xname -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >>
@@ -761,17 +761,17 @@
        (map (fn (((a, vs), s1), s2) => (PureThy.get_thm thy a, (vs, s1, s2))) xs) thy)));
 
 val _ =
-  OuterSyntax.command "realizability"
+  Outer_Syntax.command "realizability"
   "add equations characterizing realizability" Keyword.thy_decl
   (Scan.repeat1 Parse.string >> (Toplevel.theory o add_realizes_eqns));
 
 val _ =
-  OuterSyntax.command "extract_type"
+  Outer_Syntax.command "extract_type"
   "add equations characterizing type of extracted program" Keyword.thy_decl
   (Scan.repeat1 Parse.string >> (Toplevel.theory o add_typeof_eqns));
 
 val _ =
-  OuterSyntax.command "extract" "extract terms from proofs" Keyword.thy_decl
+  Outer_Syntax.command "extract" "extract terms from proofs" Keyword.thy_decl
     (Scan.repeat1 (Parse.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy =>
       extract (map (apfst (PureThy.get_thm thy)) xs) thy)));
 
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML	Sat May 15 23:40:00 2010 +0200
@@ -189,36 +189,36 @@
 (* additional outer syntax for Isar *)
 
 fun prP () =
-  OuterSyntax.improper_command "ProofGeneral.pr" "print state (internal)" Keyword.diag
+  Outer_Syntax.improper_command "ProofGeneral.pr" "print state (internal)" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o 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))));
 
 fun undoP () = (*undo without output -- historical*)
-  OuterSyntax.improper_command "ProofGeneral.undo" "(internal)" Keyword.control
+  Outer_Syntax.improper_command "ProofGeneral.undo" "(internal)" Keyword.control
     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Isar.undo 1)));
 
 fun restartP () =
-  OuterSyntax.improper_command "ProofGeneral.restart" "(internal)" Keyword.control
+  Outer_Syntax.improper_command "ProofGeneral.restart" "(internal)" Keyword.control
     (Parse.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative restart)));
 
 fun kill_proofP () =
-  OuterSyntax.improper_command "ProofGeneral.kill_proof" "(internal)" Keyword.control
+  Outer_Syntax.improper_command "ProofGeneral.kill_proof" "(internal)" Keyword.control
     (Scan.succeed (Toplevel.no_timing o
       Toplevel.imperative (fn () => (Isar.kill_proof (); tell_clear_goals ()))));
 
 fun inform_file_processedP () =
-  OuterSyntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" Keyword.control
+  Outer_Syntax.improper_command "ProofGeneral.inform_file_processed" "(internal)" Keyword.control
     (Parse.name >> (fn file =>
       Toplevel.no_timing o Toplevel.imperative (fn () => inform_file_processed file)));
 
 fun inform_file_retractedP () =
-  OuterSyntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" Keyword.control
+  Outer_Syntax.improper_command "ProofGeneral.inform_file_retracted" "(internal)" Keyword.control
     (Parse.name >> (Toplevel.no_timing oo
       (fn file => Toplevel.imperative (fn () => inform_file_retracted file))));
 
 fun process_pgipP () =
-  OuterSyntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
+  Outer_Syntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
     (Parse.text >> (Toplevel.no_timing oo
       (fn txt => Toplevel.imperative (fn () => ProofGeneralPgip.process_pgip txt))));
 
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Sat May 15 23:40:00 2010 +0200
@@ -382,7 +382,7 @@
 
 (* Sending commands to Isar *)
 
-fun isarcmd s = Isar.>>> (OuterSyntax.parse Position.none s);
+fun isarcmd s = Isar.>>> (Outer_Syntax.parse Position.none s);
 
 (* TODO:
     - apply a command given a transition function;
@@ -1013,7 +1013,7 @@
 (* Extra command for embedding prover-control inside document (obscure/debug usage). *)
 
 fun init_outer_syntax () =
-  OuterSyntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
+  Outer_Syntax.improper_command "ProofGeneral.process_pgip" "(internal)" Keyword.control
     (Parse.text >> (Toplevel.no_timing oo
       (fn txt => Toplevel.imperative (fn () => process_pgip_plain txt))));
 
--- a/src/Pure/System/isar.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/System/isar.ML	Sat May 15 23:40:00 2010 +0200
@@ -141,7 +141,7 @@
   Secure.open_unsynchronized ();
   if do_init then init () else ();
   if welcome then writeln (Session.welcome ()) else ();
-  uninterruptible (fn _ => fn () => raw_loop secure (OuterSyntax.isar sync)) ());
+  uninterruptible (fn _ => fn () => raw_loop secure (Outer_Syntax.isar sync)) ());
 
 end;
 
@@ -164,35 +164,35 @@
 (* global history *)
 
 val _ =
-  OuterSyntax.improper_command "init_toplevel" "init toplevel point-of-interest" Keyword.control
+  Outer_Syntax.improper_command "init_toplevel" "init toplevel point-of-interest" Keyword.control
     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative init));
 
 val _ =
-  OuterSyntax.improper_command "linear_undo" "undo commands" Keyword.control
+  Outer_Syntax.improper_command "linear_undo" "undo commands" Keyword.control
     (Scan.optional Parse.nat 1 >>
       (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => linear_undo n)));
 
 val _ =
-  OuterSyntax.improper_command "undo" "undo commands (skipping closed proofs)" Keyword.control
+  Outer_Syntax.improper_command "undo" "undo commands (skipping closed proofs)" Keyword.control
     (Scan.optional Parse.nat 1 >>
       (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => undo n)));
 
 val _ =
-  OuterSyntax.improper_command "undos_proof" "undo commands (skipping closed proofs)"
+  Outer_Syntax.improper_command "undos_proof" "undo commands (skipping closed proofs)"
     Keyword.control
     (Scan.optional Parse.nat 1 >> (fn n => Toplevel.no_timing o
       Toplevel.keep (fn state =>
         if Toplevel.is_proof state then (undo n; print ()) else raise Toplevel.UNDEF)));
 
 val _ =
-  OuterSyntax.improper_command "cannot_undo" "partial undo -- Proof General legacy"
+  Outer_Syntax.improper_command "cannot_undo" "partial undo -- Proof General legacy"
     Keyword.control
     (Parse.name >>
       (fn "end" => Toplevel.no_timing o Toplevel.imperative (fn () => undo 1)
         | txt => Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt))));
 
 val _ =
-  OuterSyntax.improper_command "kill" "kill partial proof or theory development" Keyword.control
+  Outer_Syntax.improper_command "kill" "kill partial proof or theory development" Keyword.control
     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
 
 end;
--- a/src/Pure/System/session.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/System/session.ML	Sat May 15 23:40:00 2010 +0200
@@ -48,7 +48,7 @@
     (if Distribution.changelog <> "" then "\nSee also " ^ Distribution.changelog else "");
 
 val _ =
-  OuterSyntax.improper_command "welcome" "print welcome message" Keyword.diag
+  Outer_Syntax.improper_command "welcome" "print welcome message" Keyword.diag
     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (writeln o welcome)));
 
 
--- a/src/Pure/Thy/thy_info.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/Thy/thy_info.ML	Sat May 15 23:40:00 2010 +0200
@@ -346,7 +346,7 @@
     val _ = CRITICAL (fn () =>
       change_deps name (Option.map (fn {master, text, parents, files, ...} =>
         make_deps upd_time master text parents files)));
-    val after_load = OuterSyntax.load_thy name pos text (time orelse ! Output.timing);
+    val after_load = Outer_Syntax.load_thy name pos text (time orelse ! Output.timing);
     val _ =
       CRITICAL (fn () =>
        (change_deps name
--- a/src/Pure/Tools/find_consts.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/Tools/find_consts.ML	Sat May 15 23:40:00 2010 +0200
@@ -154,7 +154,7 @@
   Parse.xname >> Loose;
 
 val _ =
-  OuterSyntax.improper_command "find_consts" "search constants by type pattern" Keyword.diag
+  Outer_Syntax.improper_command "find_consts" "search constants by type pattern" Keyword.diag
     (Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion)
       >> (Toplevel.no_timing oo find_consts_cmd));
 
--- a/src/Pure/Tools/find_theorems.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/Tools/find_theorems.ML	Sat May 15 23:40:00 2010 +0200
@@ -483,7 +483,7 @@
 in
 
 val _ =
-  OuterSyntax.improper_command "find_theorems" "print theorems meeting specified criteria"
+  Outer_Syntax.improper_command "find_theorems" "print theorems meeting specified criteria"
     Keyword.diag
     (options -- Scan.repeat (((Scan.option Parse.minus >> is_none) -- criterion))
       >> (Toplevel.no_timing oo find_theorems_cmd));
--- a/src/Pure/codegen.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/codegen.ML	Sat May 15 23:40:00 2010 +0200
@@ -969,7 +969,7 @@
     (Parse.verbatim >> strip_whitespace));
 
 val _ =
-  OuterSyntax.command "types_code"
+  Outer_Syntax.command "types_code"
   "associate types with target language types" Keyword.thy_decl
     (Scan.repeat1 (Parse.xname --| Parse.$$$ "(" -- Parse.string --| Parse.$$$ ")" -- parse_attach) >>
      (fn xs => Toplevel.theory (fn thy => fold (assoc_type o
@@ -977,7 +977,7 @@
          (Syntax.read_typ_global thy) mfx, aux)))) xs thy)));
 
 val _ =
-  OuterSyntax.command "consts_code"
+  Outer_Syntax.command "consts_code"
   "associate constants with target language code" Keyword.thy_decl
     (Scan.repeat1
        (Parse.term --|
@@ -1018,12 +1018,12 @@
   #> add_preprocessor unfold_preprocessor;
 
 val _ =
-  OuterSyntax.command "code_library"
+  Outer_Syntax.command "code_library"
     "generates code for terms (one structure for each theory)" Keyword.thy_decl
     (parse_code true);
 
 val _ =
-  OuterSyntax.command "code_module"
+  Outer_Syntax.command "code_module"
     "generates code for terms (single structure, incremental)" Keyword.thy_decl
     (parse_code false);
 
--- a/src/Pure/pure_setup.ML	Sat May 15 23:32:15 2010 +0200
+++ b/src/Pure/pure_setup.ML	Sat May 15 23:40:00 2010 +0200
@@ -9,7 +9,7 @@
 val theory = ThyInfo.get_theory;
 
 Context.>> (Context.map_theory
- (OuterSyntax.process_file (Path.explode "Pure.thy") #>
+ (Outer_Syntax.process_file (Path.explode "Pure.thy") #>
   Theory.end_theory));
 structure Pure = struct val thy = ML_Context.the_global_context () end;
 Context.set_thread_data NONE;