# HG changeset patch # User wenzelm # Date 1343917072 -7200 # Node ID 875a4657523ebec8013de5c5165d759043d83efa # Parent ee9cba42d83d58cdfa48e0efa5f902a90c749f09# Parent 47330b712f8f97e6e2f24ae6ab72d83a8bf4ca3f merged diff -r ee9cba42d83d -r 875a4657523e doc-src/System/Thy/Sessions.thy --- a/doc-src/System/Thy/Sessions.thy Thu Aug 02 10:10:29 2012 +0200 +++ b/doc-src/System/Thy/Sessions.thy Thu Aug 02 16:17:52 2012 +0200 @@ -195,11 +195,7 @@ component directories (\secref{sec:components}), augmented by more directories given via options @{verbatim "-d"}~@{text "DIR"} on the command line. Each such directory may contain a session - \texttt{ROOT} file and an additional catalog file @{verbatim - "etc/sessions"} with further sub-directories (list of lines). Note - that a single \texttt{ROOT} file usually defines many sessions; - catalogs are only required for extra scalability and modularity - of large libraries. + \texttt{ROOT} file with several session specifications. \medskip The subset of sessions to be managed is determined via individual @{text "SESSIONS"} given as command-line arguments, or diff -r ee9cba42d83d -r 875a4657523e doc-src/System/Thy/document/Sessions.tex --- a/doc-src/System/Thy/document/Sessions.tex Thu Aug 02 10:10:29 2012 +0200 +++ b/doc-src/System/Thy/document/Sessions.tex Thu Aug 02 16:17:52 2012 +0200 @@ -311,10 +311,7 @@ component directories (\secref{sec:components}), augmented by more directories given via options \verb|-d|~\isa{{\isaliteral{22}{\isachardoublequote}}DIR{\isaliteral{22}{\isachardoublequote}}} on the command line. Each such directory may contain a session - \texttt{ROOT} file and an additional catalog file \verb|etc/sessions| with further sub-directories (list of lines). Note - that a single \texttt{ROOT} file usually defines many sessions; - catalogs are only required for extra scalability and modularity - of large libraries. + \texttt{ROOT} file with several session specifications. \medskip The subset of sessions to be managed is determined via individual \isa{{\isaliteral{22}{\isachardoublequote}}SESSIONS{\isaliteral{22}{\isachardoublequote}}} given as command-line arguments, or diff -r ee9cba42d83d -r 875a4657523e src/HOL/ex/Interpretation_with_Defs.thy --- a/src/HOL/ex/Interpretation_with_Defs.thy Thu Aug 02 10:10:29 2012 +0200 +++ b/src/HOL/ex/Interpretation_with_Defs.thy Thu Aug 02 16:17:52 2012 +0200 @@ -6,7 +6,6 @@ theory Interpretation_with_Defs imports Pure -keywords "interpretation" :: thy_goal uses "~~/src/Tools/interpretation_with_defs.ML" begin diff -r ee9cba42d83d -r 875a4657523e src/Pure/General/graph.scala --- a/src/Pure/General/graph.scala Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/General/graph.scala Thu Aug 02 16:17:52 2012 +0200 @@ -124,15 +124,12 @@ def new_node(x: Key, info: A): Graph[Key, A] = { - if (rep.isDefinedAt(x)) throw new Graph.Duplicate(x) + if (defined(x)) throw new Graph.Duplicate(x) else new Graph[Key, A](rep + (x -> (info, (empty_keys, empty_keys)))) } def default_node(x: Key, info: A): Graph[Key, A] = - { - if (rep.isDefinedAt(x)) this - else new_node(x, info) - } + if (defined(x)) this else new_node(x, info) private def del_adjacent(fst: Boolean, x: Key)(map: SortedMap[Key, Entry], y: Key) : SortedMap[Key, Entry] = diff -r ee9cba42d83d -r 875a4657523e src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/Isar/isar_syn.ML Thu Aug 02 16:17:52 2012 +0200 @@ -7,98 +7,66 @@ structure Isar_Syn: sig end = struct -(** keywords **) - -val _ = Context.>> (Context.map_theory - (fold (fn name => (Keyword.define (name, NONE); Thy_Header.declare_keyword (name, NONE))) - ["!!", "!", "%", "(", ")", "+", ",", "--", ":", "::", ";", "<", "<=", - "=", "==", "=>", "?", "[", "\\", "\\", - "\\", "\\", "\\", "]", - "advanced", "and", "assumes", "attach", "begin", "binder", - "constrains", "defines", "fixes", "for", "identifier", "if", - "imports", "in", "includes", "infix", "infixl", "infixr", "is", "keywords", - "notes", "obtains", "open", "output", "overloaded", "pervasive", - "shows", "structure", "unchecked", "uses", "where", "|"])); - - - -(** init and exit **) - -val _ = - Outer_Syntax.command ("theory", Keyword.tag_theory Keyword.thy_begin) "begin theory context" - (Thy_Header.args >> (fn header => - Toplevel.print o - Toplevel.init_theory - (fn () => Thy_Info.toplevel_begin_theory (Thy_Load.get_master_path ()) header))); - -val _ = - Outer_Syntax.command ("end", Keyword.tag_theory Keyword.thy_end) "end context" - (Scan.succeed - (Toplevel.exit o Toplevel.end_local_theory o Toplevel.close_target o - Toplevel.end_proof (K Proof.end_notepad))); - - - (** markup commands **) val _ = Outer_Syntax.markup_command Thy_Output.Markup - ("header", Keyword.diag) "theory header" + @{command_spec "header"} "theory header" (Parse.doc_source >> Isar_Cmd.header_markup); val _ = Outer_Syntax.markup_command Thy_Output.Markup - ("chapter", Keyword.thy_heading1) "chapter heading" + @{command_spec "chapter"} "chapter heading" + (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); + +val _ = + Outer_Syntax.markup_command Thy_Output.Markup + @{command_spec "section"} "section heading" + (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); + +val _ = + Outer_Syntax.markup_command Thy_Output.Markup + @{command_spec "subsection"} "subsection heading" (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); val _ = Outer_Syntax.markup_command Thy_Output.Markup - ("section", Keyword.thy_heading2) "section heading" - (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); - -val _ = - Outer_Syntax.markup_command Thy_Output.Markup - ("subsection", Keyword.thy_heading3) "subsection heading" - (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); - -val _ = - Outer_Syntax.markup_command Thy_Output.Markup - ("subsubsection", Keyword.thy_heading4) "subsubsection heading" + @{command_spec "subsubsection"} "subsubsection heading" (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); val _ = Outer_Syntax.markup_command Thy_Output.MarkupEnv - ("text", Keyword.thy_decl) "formal comment (theory)" + @{command_spec "text"} "formal comment (theory)" (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); val _ = Outer_Syntax.markup_command Thy_Output.Verbatim - ("text_raw", Keyword.thy_decl) "raw document preparation text" + @{command_spec "text_raw"} "raw document preparation text" (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup); val _ = Outer_Syntax.markup_command Thy_Output.Markup - ("sect", Keyword.tag_proof Keyword.prf_heading2) "formal comment (proof)" + @{command_spec "sect"} "formal comment (proof)" (Parse.doc_source >> Isar_Cmd.proof_markup); val _ = Outer_Syntax.markup_command Thy_Output.Markup - ("subsect", Keyword.tag_proof Keyword.prf_heading3) "formal comment (proof)" + @{command_spec "subsect"} "formal comment (proof)" (Parse.doc_source >> Isar_Cmd.proof_markup); val _ = Outer_Syntax.markup_command Thy_Output.Markup - ("subsubsect", Keyword.tag_proof Keyword.prf_heading4) "formal comment (proof)" + @{command_spec "subsubsect"} "formal comment (proof)" (Parse.doc_source >> Isar_Cmd.proof_markup); val _ = Outer_Syntax.markup_command Thy_Output.MarkupEnv - ("txt", Keyword.tag_proof Keyword.prf_decl) "formal comment (proof)" + @{command_spec "txt"} "formal comment (proof)" (Parse.doc_source >> Isar_Cmd.proof_markup); val _ = Outer_Syntax.markup_command Thy_Output.Verbatim - ("txt_raw", Keyword.tag_proof Keyword.prf_decl) "raw document preparation text (proof)" + @{command_spec "txt_raw"} "raw document preparation text (proof)" (Parse.doc_source >> Isar_Cmd.proof_markup); @@ -108,19 +76,19 @@ (* classes and sorts *) val _ = - Outer_Syntax.command ("classes", Keyword.thy_decl) "declare type classes" - (Scan.repeat1 (Parse.binding -- Scan.optional ((Parse.$$$ "\\" || Parse.$$$ "<") |-- + Outer_Syntax.command @{command_spec "classes"} "declare type classes" + (Scan.repeat1 (Parse.binding -- Scan.optional ((@{keyword "\"} || @{keyword "<"}) |-- Parse.!!! (Parse.list1 Parse.class)) []) >> (Toplevel.theory o fold AxClass.axiomatize_class_cmd)); val _ = - Outer_Syntax.command ("classrel", Keyword.thy_decl) "state inclusion of type classes (axiomatic!)" - (Parse.and_list1 (Parse.class -- ((Parse.$$$ "\\" || Parse.$$$ "<") + Outer_Syntax.command @{command_spec "classrel"} "state inclusion of type classes (axiomatic!)" + (Parse.and_list1 (Parse.class -- ((@{keyword "\"} || @{keyword "<"}) |-- Parse.!!! Parse.class)) >> (Toplevel.theory o AxClass.axiomatize_classrel_cmd)); val _ = - Outer_Syntax.local_theory ("default_sort", Keyword.thy_decl) + Outer_Syntax.local_theory @{command_spec "default_sort"} "declare default sort for explicit type variables" (Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy)); @@ -128,51 +96,49 @@ (* types *) val _ = - Outer_Syntax.local_theory ("typedecl", Keyword.thy_decl) "type declaration" + Outer_Syntax.local_theory @{command_spec "typedecl"} "type declaration" (Parse.type_args -- Parse.binding -- Parse.opt_mixfix >> (fn ((args, a), mx) => Typedecl.typedecl (a, map (rpair dummyS) args, mx) #> snd)); val _ = - Outer_Syntax.local_theory ("type_synonym", Keyword.thy_decl) "declare type abbreviation" + Outer_Syntax.local_theory @{command_spec "type_synonym"} "declare type abbreviation" (Parse.type_args -- Parse.binding -- - (Parse.$$$ "=" |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix')) + (@{keyword "="} |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix')) >> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs)); val _ = - Outer_Syntax.command ("nonterminal", Keyword.thy_decl) + Outer_Syntax.command @{command_spec "nonterminal"} "declare syntactic type constructors (grammar nonterminal symbols)" (Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global)); val _ = - Outer_Syntax.command ("arities", Keyword.thy_decl) "state type arities (axiomatic!)" + Outer_Syntax.command @{command_spec "arities"} "state type arities (axiomatic!)" (Scan.repeat1 Parse.arity >> (Toplevel.theory o fold AxClass.axiomatize_arity_cmd)); (* consts and syntax *) val _ = - Outer_Syntax.command ("judgment", Keyword.thy_decl) "declare object-logic judgment" + Outer_Syntax.command @{command_spec "judgment"} "declare object-logic judgment" (Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd)); val _ = - Outer_Syntax.command ("consts", Keyword.thy_decl) "declare constants" + Outer_Syntax.command @{command_spec "consts"} "declare constants" (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts)); -val opt_overloaded = Parse.opt_keyword "overloaded"; - val mode_spec = - (Parse.$$$ "output" >> K ("", false)) || - Parse.name -- Scan.optional (Parse.$$$ "output" >> K false) true; + (@{keyword "output"} >> K ("", false)) || + Parse.name -- Scan.optional (@{keyword "output"} >> K false) true; val opt_mode = - Scan.optional (Parse.$$$ "(" |-- Parse.!!! (mode_spec --| Parse.$$$ ")")) Syntax.mode_default; + Scan.optional (@{keyword "("} |-- Parse.!!! (mode_spec --| @{keyword ")"})) Syntax.mode_default; val _ = - Outer_Syntax.command ("syntax", Keyword.thy_decl) "declare syntactic constants" + Outer_Syntax.command @{command_spec "syntax"} "declare syntactic constants" (opt_mode -- Scan.repeat1 Parse.const_decl >> (Toplevel.theory o uncurry Sign.add_modesyntax)); val _ = - Outer_Syntax.command ("no_syntax", Keyword.thy_decl) "delete syntax declarations" + Outer_Syntax.command @{command_spec "no_syntax"} "delete syntax declarations" (opt_mode -- Scan.repeat1 Parse.const_decl >> (Toplevel.theory o uncurry Sign.del_modesyntax)); @@ -180,43 +146,44 @@ val trans_pat = Scan.optional - (Parse.$$$ "(" |-- Parse.!!! (Parse.xname --| Parse.$$$ ")")) "logic" + (@{keyword "("} |-- Parse.!!! (Parse.xname --| @{keyword ")"})) "logic" -- Parse.inner_syntax Parse.string; fun trans_arrow toks = - ((Parse.$$$ "\\" || Parse.$$$ "=>") >> K Syntax.Parse_Rule || - (Parse.$$$ "\\" || Parse.$$$ "<=") >> K Syntax.Print_Rule || - (Parse.$$$ "\\" || Parse.$$$ "==") >> K Syntax.Parse_Print_Rule) toks; + ((@{keyword "\"} || @{keyword "=>"}) >> K Syntax.Parse_Rule || + (@{keyword "\"} || @{keyword "<="}) >> K Syntax.Print_Rule || + (@{keyword "\"} || @{keyword "=="}) >> K Syntax.Parse_Print_Rule) toks; val trans_line = trans_pat -- Parse.!!! (trans_arrow -- trans_pat) >> (fn (left, (arr, right)) => arr (left, right)); val _ = - Outer_Syntax.command ("translations", Keyword.thy_decl) "declare syntax translation rules" + Outer_Syntax.command @{command_spec "translations"} "declare syntax translation rules" (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.translations)); val _ = - Outer_Syntax.command ("no_translations", Keyword.thy_decl) "remove syntax translation rules" + Outer_Syntax.command @{command_spec "no_translations"} "remove syntax translation rules" (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.no_translations)); (* axioms and definitions *) val _ = - Outer_Syntax.command ("axioms", Keyword.thy_decl) "state arbitrary propositions (axiomatic!)" + Outer_Syntax.command @{command_spec "axioms"} "state arbitrary propositions (axiomatic!)" (Scan.repeat1 Parse_Spec.spec >> (fn spec => Toplevel.theory (fn thy => (legacy_feature "Old 'axioms' command -- use 'axiomatization' instead"; Isar_Cmd.add_axioms spec thy)))); val opt_unchecked_overloaded = - Scan.optional (Parse.$$$ "(" |-- Parse.!!! - (((Parse.$$$ "unchecked" >> K true) -- Scan.optional (Parse.$$$ "overloaded" >> K true) false || - Parse.$$$ "overloaded" >> K (false, true)) --| Parse.$$$ ")")) (false, false); + Scan.optional (@{keyword "("} |-- Parse.!!! + (((@{keyword "unchecked"} >> K true) -- + Scan.optional (@{keyword "overloaded"} >> K true) false || + @{keyword "overloaded"} >> K (false, true)) --| @{keyword ")"})) (false, false); val _ = - Outer_Syntax.command ("defs", Keyword.thy_decl) "define constants" + Outer_Syntax.command @{command_spec "defs"} "define constants" (opt_unchecked_overloaded -- Scan.repeat1 (Parse_Spec.thm_name ":" -- Parse.prop >> (fn ((x, y), z) => ((x, z), y))) >> (Toplevel.theory o Isar_Cmd.add_defs)); @@ -225,34 +192,34 @@ (* constant definitions and abbreviations *) val _ = - Outer_Syntax.local_theory' ("definition", Keyword.thy_decl) "constant definition" + Outer_Syntax.local_theory' @{command_spec "definition"} "constant definition" (Parse_Spec.constdef >> (fn args => #2 oo Specification.definition_cmd args)); val _ = - Outer_Syntax.local_theory' ("abbreviation", Keyword.thy_decl) "constant abbreviation" + Outer_Syntax.local_theory' @{command_spec "abbreviation"} "constant abbreviation" (opt_mode -- (Scan.option Parse_Spec.constdecl -- Parse.prop) >> (fn (mode, args) => Specification.abbreviation_cmd mode args)); val _ = - Outer_Syntax.local_theory ("type_notation", Keyword.thy_decl) + Outer_Syntax.local_theory @{command_spec "type_notation"} "add concrete syntax for type constructors" (opt_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) >> (fn (mode, args) => Specification.type_notation_cmd true mode args)); val _ = - Outer_Syntax.local_theory ("no_type_notation", Keyword.thy_decl) + Outer_Syntax.local_theory @{command_spec "no_type_notation"} "delete concrete syntax for type constructors" (opt_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) >> (fn (mode, args) => Specification.type_notation_cmd false mode args)); val _ = - Outer_Syntax.local_theory ("notation", Keyword.thy_decl) + Outer_Syntax.local_theory @{command_spec "notation"} "add concrete syntax for constants / fixed variables" (opt_mode -- Parse.and_list1 (Parse.const -- Parse_Spec.locale_mixfix) >> (fn (mode, args) => Specification.notation_cmd true mode args)); val _ = - Outer_Syntax.local_theory ("no_notation", Keyword.thy_decl) + Outer_Syntax.local_theory @{command_spec "no_notation"} "delete concrete syntax for constants / fixed variables" (opt_mode -- Parse.and_list1 (Parse.const -- Parse_Spec.locale_mixfix) >> (fn (mode, args) => Specification.notation_cmd false mode args)); @@ -261,7 +228,7 @@ (* constant specifications *) val _ = - Outer_Syntax.command ("axiomatization", Keyword.thy_decl) "axiomatic constant specification" + Outer_Syntax.command @{command_spec "axiomatization"} "axiomatic constant specification" (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))); @@ -274,14 +241,14 @@ >> (fn (facts, fixes) => #2 oo Specification.theorems_cmd kind facts fixes); val _ = - Outer_Syntax.local_theory' ("theorems", Keyword.thy_decl) "define theorems" + Outer_Syntax.local_theory' @{command_spec "theorems"} "define theorems" (theorems Thm.theoremK); val _ = - Outer_Syntax.local_theory' ("lemmas", Keyword.thy_decl) "define lemmas" (theorems Thm.lemmaK); + Outer_Syntax.local_theory' @{command_spec "lemmas"} "define lemmas" (theorems Thm.lemmaK); val _ = - Outer_Syntax.local_theory' ("declare", Keyword.thy_decl) "declare theorems" + Outer_Syntax.local_theory' @{command_spec "declare"} "declare theorems" (Parse.and_list1 Parse_Spec.xthms1 -- Parse.for_fixes >> (fn (facts, fixes) => #2 oo Specification.theorems_cmd "" [(Attrib.empty_binding, flat facts)] fixes)); @@ -289,83 +256,79 @@ (* name space entry path *) -fun hide_names name hide what = - Outer_Syntax.command (name, Keyword.thy_decl) ("hide " ^ what ^ " from name space") +fun hide_names spec hide what = + Outer_Syntax.command spec ("hide " ^ what ^ " from name space") ((Parse.opt_keyword "open" >> not) -- Scan.repeat1 Parse.xname >> (Toplevel.theory o uncurry hide)); -val _ = hide_names "hide_class" Isar_Cmd.hide_class "classes"; -val _ = hide_names "hide_type" Isar_Cmd.hide_type "types"; -val _ = hide_names "hide_const" Isar_Cmd.hide_const "constants"; -val _ = hide_names "hide_fact" Isar_Cmd.hide_fact "facts"; +val _ = hide_names @{command_spec "hide_class"} Isar_Cmd.hide_class "classes"; +val _ = hide_names @{command_spec "hide_type"} Isar_Cmd.hide_type "types"; +val _ = hide_names @{command_spec "hide_const"} Isar_Cmd.hide_const "constants"; +val _ = hide_names @{command_spec "hide_fact"} Isar_Cmd.hide_fact "facts"; (* use ML text *) val _ = - Outer_Syntax.command ("use", Keyword.tag_ml Keyword.thy_decl) "ML text from file" + Outer_Syntax.command @{command_spec "use"} "ML text from file" (Parse.path >> (fn path => Toplevel.generic_theory (Thy_Load.exec_ml path))); val _ = - Outer_Syntax.command ("ML", Keyword.tag_ml Keyword.thy_decl) - "ML text within theory or local theory" + Outer_Syntax.command @{command_spec "ML"} "ML text within theory or local theory" (Parse.ML_source >> (fn (txt, pos) => Toplevel.generic_theory (ML_Context.exec (fn () => ML_Context.eval_text true pos txt) #> Local_Theory.propagate_ml_env))); val _ = - Outer_Syntax.command ("ML_prf", Keyword.tag_proof Keyword.prf_decl) "ML text within proof" + Outer_Syntax.command @{command_spec "ML_prf"} "ML text within proof" (Parse.ML_source >> (fn (txt, pos) => Toplevel.proof (Proof.map_context (Context.proof_map (ML_Context.exec (fn () => ML_Context.eval_text true pos txt))) #> Proof.propagate_ml_env))); val _ = - Outer_Syntax.command ("ML_val", Keyword.tag_ml Keyword.diag) "diagnostic ML text" + Outer_Syntax.command @{command_spec "ML_val"} "diagnostic ML text" (Parse.ML_source >> Isar_Cmd.ml_diag true); val _ = - Outer_Syntax.command ("ML_command", Keyword.tag_ml Keyword.diag) "diagnostic ML text (silent)" + Outer_Syntax.command @{command_spec "ML_command"} "diagnostic ML text (silent)" (Parse.ML_source >> (Toplevel.no_timing oo Isar_Cmd.ml_diag false)); val _ = - Outer_Syntax.command ("setup", Keyword.tag_ml Keyword.thy_decl) "ML theory setup" + Outer_Syntax.command @{command_spec "setup"} "ML theory setup" (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.global_setup)); val _ = - Outer_Syntax.local_theory ("local_setup", Keyword.tag_ml Keyword.thy_decl) "ML local theory setup" + Outer_Syntax.local_theory @{command_spec "local_setup"} "ML local theory setup" (Parse.ML_source >> Isar_Cmd.local_setup); val _ = - Outer_Syntax.command ("attribute_setup", Keyword.tag_ml Keyword.thy_decl) "define attribute in ML" + Outer_Syntax.command @{command_spec "attribute_setup"} "define attribute in ML" (Parse.position Parse.name -- - Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Scan.optional Parse.text "") + Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "") >> (fn (name, (txt, cmt)) => Toplevel.theory (Attrib.attribute_setup name txt cmt))); val _ = - Outer_Syntax.command ("method_setup", Keyword.tag_ml Keyword.thy_decl) "define proof method in ML" + Outer_Syntax.command @{command_spec "method_setup"} "define proof method in ML" (Parse.position Parse.name -- - Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Scan.optional Parse.text "") + Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "") >> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt))); val _ = - Outer_Syntax.local_theory ("declaration", Keyword.tag_ml Keyword.thy_decl) - "generic ML declaration" + Outer_Syntax.local_theory @{command_spec "declaration"} "generic ML declaration" (Parse.opt_keyword "pervasive" -- Parse.ML_source >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = false, pervasive = pervasive} txt)); val _ = - Outer_Syntax.local_theory ("syntax_declaration", Keyword.tag_ml Keyword.thy_decl) - "generic ML declaration" + Outer_Syntax.local_theory @{command_spec "syntax_declaration"} "generic ML syntax declaration" (Parse.opt_keyword "pervasive" -- Parse.ML_source >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = true, pervasive = pervasive} txt)); val _ = - Outer_Syntax.local_theory ("simproc_setup", Keyword.tag_ml Keyword.thy_decl) - "define simproc in ML" + Outer_Syntax.local_theory @{command_spec "simproc_setup"} "define simproc in ML" (Parse.position Parse.name -- - (Parse.$$$ "(" |-- Parse.enum1 "|" Parse.term --| Parse.$$$ ")" --| Parse.$$$ "=") -- - Parse.ML_source -- Scan.optional (Parse.$$$ "identifier" |-- Scan.repeat1 Parse.xname) [] + (@{keyword "("} |-- Parse.enum1 "|" Parse.term --| @{keyword ")"} --| @{keyword "="}) -- + Parse.ML_source -- Scan.optional (@{keyword "identifier"} |-- Scan.repeat1 Parse.xname) [] >> (fn (((a, b), c), d) => Isar_Cmd.simproc_setup a b c d)); @@ -374,27 +337,27 @@ val trfun = Parse.opt_keyword "advanced" -- Parse.ML_source; val _ = - Outer_Syntax.command ("parse_ast_translation", Keyword.tag_ml Keyword.thy_decl) + Outer_Syntax.command @{command_spec "parse_ast_translation"} "install parse ast translation functions" (trfun >> (Toplevel.theory o Isar_Cmd.parse_ast_translation)); val _ = - Outer_Syntax.command ("parse_translation", Keyword.tag_ml Keyword.thy_decl) + Outer_Syntax.command @{command_spec "parse_translation"} "install parse translation functions" (trfun >> (Toplevel.theory o Isar_Cmd.parse_translation)); val _ = - Outer_Syntax.command ("print_translation", Keyword.tag_ml Keyword.thy_decl) + Outer_Syntax.command @{command_spec "print_translation"} "install print translation functions" (trfun >> (Toplevel.theory o Isar_Cmd.print_translation)); val _ = - Outer_Syntax.command ("typed_print_translation", Keyword.tag_ml Keyword.thy_decl) + Outer_Syntax.command @{command_spec "typed_print_translation"} "install typed print translation functions" (trfun >> (Toplevel.theory o Isar_Cmd.typed_print_translation)); val _ = - Outer_Syntax.command ("print_ast_translation", Keyword.tag_ml Keyword.thy_decl) + Outer_Syntax.command @{command_spec "print_ast_translation"} "install print ast translation functions" (trfun >> (Toplevel.theory o Isar_Cmd.print_ast_translation)); @@ -402,32 +365,33 @@ (* oracles *) val _ = - Outer_Syntax.command ("oracle", Keyword.tag_ml Keyword.thy_decl) "declare oracle" - (Parse.position Parse.name -- (Parse.$$$ "=" |-- Parse.ML_source) >> + Outer_Syntax.command @{command_spec "oracle"} "declare oracle" + (Parse.position Parse.name -- (@{keyword "="} |-- Parse.ML_source) >> (fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y))); (* bundled declarations *) val _ = - Outer_Syntax.local_theory ("bundle", Keyword.thy_decl) "define bundle of declarations" - ((Parse.binding --| Parse.$$$ "=") -- Parse_Spec.xthms1 -- Parse.for_fixes + Outer_Syntax.local_theory @{command_spec "bundle"} "define bundle of declarations" + ((Parse.binding --| @{keyword "="}) -- Parse_Spec.xthms1 -- Parse.for_fixes >> (uncurry Bundle.bundle_cmd)); val _ = - Outer_Syntax.command ("include", Keyword.prf_decl) + Outer_Syntax.command @{command_spec "include"} "include declarations from bundle in proof body" (Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.print oo (Toplevel.proof o Bundle.include_cmd))); val _ = - Outer_Syntax.command ("including", Keyword.prf_decl) + Outer_Syntax.command @{command_spec "including"} "include declarations from bundle in goal refinement" (Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.print oo (Toplevel.proof o Bundle.including_cmd))); val _ = - Outer_Syntax.improper_command ("print_bundles", Keyword.diag) "print bundles of declarations" + Outer_Syntax.improper_command @{command_spec "print_bundles"} + "print bundles of declarations" (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o Toplevel.keep (Bundle.print_bundles o Toplevel.context_of))); @@ -435,7 +399,7 @@ (* local theories *) val _ = - Outer_Syntax.command ("context", Keyword.thy_decl) "begin local theory context" + Outer_Syntax.command @{command_spec "context"} "begin local theory context" ((Parse.position Parse.xname >> (fn name => Toplevel.print o Toplevel.begin_local_theory true (Named_Target.context_cmd name)) || Scan.optional Parse_Spec.includes [] -- Scan.repeat Parse_Spec.context_element @@ -447,13 +411,13 @@ val locale_val = Parse_Spec.locale_expression false -- - Scan.optional (Parse.$$$ "+" |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] || + Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] || Scan.repeat1 Parse_Spec.context_element >> pair ([], []); val _ = - Outer_Syntax.command ("locale", Keyword.thy_decl) "define named proof context" + Outer_Syntax.command @{command_spec "locale"} "define named proof context" (Parse.binding -- - Scan.optional (Parse.$$$ "=" |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin + Scan.optional (@{keyword "="} |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin >> (fn ((name, (expr, elems)), begin) => (begin ? Toplevel.print) o Toplevel.begin_local_theory begin (Expression.add_locale_cmd I name Binding.empty expr elems #> snd))); @@ -464,22 +428,22 @@ (Parse.where_ |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []; val _ = - Outer_Syntax.command ("sublocale", Keyword.thy_schematic_goal) + Outer_Syntax.command @{command_spec "sublocale"} "prove sublocale relation between a locale and a locale expression" - (Parse.position Parse.xname --| (Parse.$$$ "\\" || Parse.$$$ "<") -- + (Parse.position Parse.xname --| (@{keyword "\"} || @{keyword "<"}) -- parse_interpretation_arguments false >> (fn (loc, (expr, equations)) => Toplevel.print o Toplevel.theory_to_proof (Expression.sublocale_cmd I loc expr equations))); val _ = - Outer_Syntax.command ("interpretation", Keyword.thy_schematic_goal) + Outer_Syntax.command @{command_spec "interpretation"} "prove interpretation of locale expression in theory" (parse_interpretation_arguments true >> (fn (expr, equations) => Toplevel.print o Toplevel.theory_to_proof (Expression.interpretation_cmd expr equations))); val _ = - Outer_Syntax.command ("interpret", Keyword.tag_proof Keyword.prf_goal) + Outer_Syntax.command @{command_spec "interpret"} "prove interpretation of locale expression in proof context" (parse_interpretation_arguments true >> (fn (expr, equations) => Toplevel.print o @@ -490,30 +454,30 @@ val class_val = Parse_Spec.class_expression -- - Scan.optional (Parse.$$$ "+" |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] || + Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] || Scan.repeat1 Parse_Spec.context_element >> pair []; val _ = - Outer_Syntax.command ("class", Keyword.thy_decl) "define type class" - (Parse.binding -- Scan.optional (Parse.$$$ "=" |-- class_val) ([], []) -- Parse.opt_begin + Outer_Syntax.command @{command_spec "class"} "define type class" + (Parse.binding -- Scan.optional (@{keyword "="} |-- class_val) ([], []) -- Parse.opt_begin >> (fn ((name, (supclasses, elems)), begin) => (begin ? Toplevel.print) o Toplevel.begin_local_theory begin (Class_Declaration.class_cmd I name supclasses elems #> snd))); val _ = - Outer_Syntax.local_theory_to_proof ("subclass", Keyword.thy_goal) "prove a subclass relation" + Outer_Syntax.local_theory_to_proof @{command_spec "subclass"} "prove a subclass relation" (Parse.class >> Class_Declaration.subclass_cmd I); val _ = - Outer_Syntax.command ("instantiation", Keyword.thy_decl) "instantiate and prove type arity" + Outer_Syntax.command @{command_spec "instantiation"} "instantiate and prove type arity" (Parse.multi_arity --| Parse.begin >> (fn arities => Toplevel.print o Toplevel.begin_local_theory true (Class.instantiation_cmd arities))); val _ = - Outer_Syntax.command ("instance", Keyword.thy_goal) "prove type arity or subclass relation" + Outer_Syntax.command @{command_spec "instance"} "prove type arity or subclass relation" ((Parse.class -- - ((Parse.$$$ "\\" || Parse.$$$ "<") |-- Parse.!!! Parse.class) >> Class.classrel_cmd || + ((@{keyword "\"} || @{keyword "<"}) |-- Parse.!!! Parse.class) >> Class.classrel_cmd || Parse.multi_arity >> Class.instance_arity_cmd) >> (Toplevel.print oo Toplevel.theory_to_proof) || Scan.succeed @@ -523,9 +487,9 @@ (* arbitrary overloading *) val _ = - Outer_Syntax.command ("overloading", Keyword.thy_decl) "overloaded definitions" - (Scan.repeat1 (Parse.name --| (Parse.$$$ "\\" || Parse.$$$ "==") -- Parse.term -- - Scan.optional (Parse.$$$ "(" |-- (Parse.$$$ "unchecked" >> K false) --| Parse.$$$ ")") true + Outer_Syntax.command @{command_spec "overloading"} "overloaded definitions" + (Scan.repeat1 (Parse.name --| (@{keyword "\"} || @{keyword "=="}) -- Parse.term -- + Scan.optional (@{keyword "("} |-- (@{keyword "unchecked"} >> K false) --| @{keyword ")"}) true >> Parse.triple1) --| Parse.begin >> (fn operations => Toplevel.print o Toplevel.begin_local_theory true (Overloading.overloading_cmd operations))); @@ -534,7 +498,7 @@ (* code generation *) val _ = - Outer_Syntax.command ("code_datatype", Keyword.thy_decl) + Outer_Syntax.command @{command_spec "code_datatype"} "define set of code datatype constructors" (Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd)); @@ -544,10 +508,8 @@ (* statements *) -fun gen_theorem schematic kind = - Outer_Syntax.local_theory_to_proof' - (if schematic then ("schematic_" ^ kind, Keyword.thy_schematic_goal) - else (kind, Keyword.thy_goal)) +fun gen_theorem spec schematic kind = + Outer_Syntax.local_theory_to_proof' spec ("state " ^ (if schematic then "schematic " ^ kind else kind)) (Scan.optional (Parse_Spec.opt_thm_name ":" --| Scan.ahead (Parse_Spec.includes >> K "" || @@ -557,32 +519,32 @@ ((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd) kind NONE (K I) a includes elems concl))); -val _ = gen_theorem false Thm.theoremK; -val _ = gen_theorem false Thm.lemmaK; -val _ = gen_theorem false Thm.corollaryK; -val _ = gen_theorem true Thm.theoremK; -val _ = gen_theorem true Thm.lemmaK; -val _ = gen_theorem true Thm.corollaryK; +val _ = gen_theorem @{command_spec "theorem"} false Thm.theoremK; +val _ = gen_theorem @{command_spec "lemma"} false Thm.lemmaK; +val _ = gen_theorem @{command_spec "corollary"} false Thm.corollaryK; +val _ = gen_theorem @{command_spec "schematic_theorem"} true Thm.theoremK; +val _ = gen_theorem @{command_spec "schematic_lemma"} true Thm.lemmaK; +val _ = gen_theorem @{command_spec "schematic_corollary"} true Thm.corollaryK; val _ = - Outer_Syntax.local_theory_to_proof ("notepad", Keyword.thy_decl) "begin proof context" + Outer_Syntax.local_theory_to_proof @{command_spec "notepad"} "begin proof context" (Parse.begin >> K Proof.begin_notepad); val _ = - Outer_Syntax.command ("have", Keyword.tag_proof Keyword.prf_goal) "state local goal" + Outer_Syntax.command @{command_spec "have"} "state local goal" (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.have)); val _ = - Outer_Syntax.command ("hence", Keyword.tag_proof Keyword.prf_goal) "abbreviates \"then have\"" + Outer_Syntax.command @{command_spec "hence"} "abbreviates \"then have\"" (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.hence)); val _ = - Outer_Syntax.command ("show", Keyword.tag_proof Keyword.prf_asm_goal) + Outer_Syntax.command @{command_spec "show"} "state local goal, solving current obligation" (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.show)); val _ = - Outer_Syntax.command ("thus", Keyword.tag_proof Keyword.prf_asm_goal) "abbreviates \"then show\"" + Outer_Syntax.command @{command_spec "thus"} "abbreviates \"then show\"" (Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.thus)); @@ -591,160 +553,147 @@ val facts = Parse.and_list1 Parse_Spec.xthms1; val _ = - Outer_Syntax.command ("then", Keyword.tag_proof Keyword.prf_chain) "forward chaining" + Outer_Syntax.command @{command_spec "then"} "forward chaining" (Scan.succeed (Toplevel.print o Toplevel.proof Proof.chain)); val _ = - Outer_Syntax.command ("from", Keyword.tag_proof Keyword.prf_chain) - "forward chaining from given facts" + Outer_Syntax.command @{command_spec "from"} "forward chaining from given facts" (facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss_cmd))); val _ = - Outer_Syntax.command ("with", Keyword.tag_proof Keyword.prf_chain) - "forward chaining from given and current facts" + Outer_Syntax.command @{command_spec "with"} "forward chaining from given and current facts" (facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss_cmd))); val _ = - Outer_Syntax.command ("note", Keyword.tag_proof Keyword.prf_decl) "define facts" + Outer_Syntax.command @{command_spec "note"} "define facts" (Parse_Spec.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss_cmd))); val _ = - Outer_Syntax.command ("using", Keyword.tag_proof Keyword.prf_decl) "augment goal facts" + Outer_Syntax.command @{command_spec "using"} "augment goal facts" (facts >> (Toplevel.print oo (Toplevel.proof o Proof.using_cmd))); val _ = - Outer_Syntax.command ("unfolding", Keyword.tag_proof Keyword.prf_decl) - "unfold definitions in goal and facts" + Outer_Syntax.command @{command_spec "unfolding"} "unfold definitions in goal and facts" (facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding_cmd))); (* proof context *) val _ = - Outer_Syntax.command ("fix", Keyword.tag_proof Keyword.prf_asm) - "fix local variables (Skolem constants)" + Outer_Syntax.command @{command_spec "fix"} "fix local variables (Skolem constants)" (Parse.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix_cmd))); val _ = - Outer_Syntax.command ("assume", Keyword.tag_proof Keyword.prf_asm) "assume propositions" + Outer_Syntax.command @{command_spec "assume"} "assume propositions" (Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume_cmd))); val _ = - Outer_Syntax.command ("presume", Keyword.tag_proof Keyword.prf_asm) - "assume propositions, to be established later" + Outer_Syntax.command @{command_spec "presume"} "assume propositions, to be established later" (Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume_cmd))); val _ = - Outer_Syntax.command ("def", Keyword.tag_proof Keyword.prf_asm) "local definition" + Outer_Syntax.command @{command_spec "def"} "local definition" (Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- ((Parse.binding -- Parse.opt_mixfix) -- - ((Parse.$$$ "\\" || Parse.$$$ "==") |-- Parse.!!! Parse.termp))) + ((@{keyword "\"} || @{keyword "=="}) |-- Parse.!!! Parse.termp))) >> (Toplevel.print oo (Toplevel.proof o Proof.def_cmd))); val _ = - Outer_Syntax.command ("obtain", Keyword.tag_proof Keyword.prf_asm_goal) - "generalized elimination" + Outer_Syntax.command @{command_spec "obtain"} "generalized elimination" (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 _ = - Outer_Syntax.command ("guess", Keyword.tag_proof Keyword.prf_asm_goal) - "wild guessing (unstructured)" + Outer_Syntax.command @{command_spec "guess"} "wild guessing (unstructured)" (Scan.optional Parse.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess_cmd))); val _ = - Outer_Syntax.command ("let", Keyword.tag_proof Keyword.prf_decl) "bind text variables" - (Parse.and_list1 (Parse.and_list1 Parse.term -- (Parse.$$$ "=" |-- Parse.term)) + Outer_Syntax.command @{command_spec "let"} "bind text variables" + (Parse.and_list1 (Parse.and_list1 Parse.term -- (@{keyword "="} |-- Parse.term)) >> (Toplevel.print oo (Toplevel.proof o Proof.let_bind_cmd))); val _ = - Outer_Syntax.command ("write", Keyword.tag_proof Keyword.prf_decl) - "add concrete syntax for constants / fixed variables" + Outer_Syntax.command @{command_spec "write"} "add concrete syntax for constants / fixed variables" (opt_mode -- Parse.and_list1 (Parse.const -- Parse_Spec.locale_mixfix) >> (fn (mode, args) => Toplevel.print o Toplevel.proof (Proof.write_cmd mode args))); val case_spec = - (Parse.$$$ "(" |-- - Parse.!!! (Parse.xname -- Scan.repeat1 (Parse.maybe Parse.binding) --| Parse.$$$ ")") || + (@{keyword "("} |-- + Parse.!!! (Parse.xname -- Scan.repeat1 (Parse.maybe Parse.binding) --| @{keyword ")"}) || Parse.xname >> rpair []) -- Parse_Spec.opt_attribs >> Parse.triple1; val _ = - Outer_Syntax.command ("case", Keyword.tag_proof Keyword.prf_asm) "invoke local context" + Outer_Syntax.command @{command_spec "case"} "invoke local context" (case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case_cmd))); (* proof structure *) val _ = - Outer_Syntax.command ("{", Keyword.tag_proof Keyword.prf_open) "begin explicit proof block" + Outer_Syntax.command @{command_spec "{"} "begin explicit proof block" (Scan.succeed (Toplevel.print o Toplevel.proof Proof.begin_block)); val _ = - Outer_Syntax.command ("}", Keyword.tag_proof Keyword.prf_close) "end explicit proof block" + Outer_Syntax.command @{command_spec "}"} "end explicit proof block" (Scan.succeed (Toplevel.print o Toplevel.proof Proof.end_block)); val _ = - Outer_Syntax.command ("next", Keyword.tag_proof Keyword.prf_block) "enter next proof block" + Outer_Syntax.command @{command_spec "next"} "enter next proof block" (Scan.succeed (Toplevel.print o Toplevel.proof Proof.next_block)); (* end proof *) val _ = - Outer_Syntax.command ("qed", Keyword.tag_proof Keyword.qed_block) "conclude (sub-)proof" + Outer_Syntax.command @{command_spec "qed"} "conclude (sub-)proof" (Scan.option Method.parse >> Isar_Cmd.qed); val _ = - Outer_Syntax.command ("by", Keyword.tag_proof Keyword.qed) "terminal backward proof" + Outer_Syntax.command @{command_spec "by"} "terminal backward proof" (Method.parse -- Scan.option Method.parse >> Isar_Cmd.terminal_proof); val _ = - Outer_Syntax.command ("..", Keyword.tag_proof Keyword.qed) "default proof" + Outer_Syntax.command @{command_spec ".."} "default proof" (Scan.succeed Isar_Cmd.default_proof); val _ = - Outer_Syntax.command (".", Keyword.tag_proof Keyword.qed) "immediate proof" + Outer_Syntax.command @{command_spec "."} "immediate proof" (Scan.succeed Isar_Cmd.immediate_proof); val _ = - Outer_Syntax.command ("done", Keyword.tag_proof Keyword.qed) "done proof" + Outer_Syntax.command @{command_spec "done"} "done proof" (Scan.succeed Isar_Cmd.done_proof); val _ = - Outer_Syntax.command ("sorry", Keyword.tag_proof Keyword.qed) - "skip proof (quick-and-dirty mode only!)" + Outer_Syntax.command @{command_spec "sorry"} "skip proof (quick-and-dirty mode only!)" (Scan.succeed Isar_Cmd.skip_proof); val _ = - Outer_Syntax.command ("oops", Keyword.tag_proof Keyword.qed_global) "forget proof" + Outer_Syntax.command @{command_spec "oops"} "forget proof" (Scan.succeed Toplevel.forget_proof); (* proof steps *) val _ = - Outer_Syntax.command ("defer", Keyword.tag_proof Keyword.prf_script) - "shuffle internal proof state" + Outer_Syntax.command @{command_spec "defer"} "shuffle internal proof state" (Scan.option Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.defer))); val _ = - Outer_Syntax.command ("prefer", Keyword.tag_proof Keyword.prf_script) - "shuffle internal proof state" + Outer_Syntax.command @{command_spec "prefer"} "shuffle internal proof state" (Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.prefer))); val _ = - Outer_Syntax.command ("apply", Keyword.tag_proof Keyword.prf_script) - "initial refinement step (unstructured)" + Outer_Syntax.command @{command_spec "apply"} "initial refinement step (unstructured)" (Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply))); val _ = - Outer_Syntax.command ("apply_end", Keyword.tag_proof Keyword.prf_script) - "terminal refinement (unstructured)" + Outer_Syntax.command @{command_spec "apply_end"} "terminal refinement (unstructured)" (Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply_end))); val _ = - Outer_Syntax.command ("proof", Keyword.tag_proof Keyword.prf_block) "backward proof" + Outer_Syntax.command @{command_spec "proof"} "backward proof" (Scan.option Method.parse >> (fn m => Toplevel.print o Toplevel.actual_proof (Proof_Node.applys (Proof.proof m)) o Toplevel.skip_proof (fn i => i + 1))); @@ -753,25 +702,23 @@ (* calculational proof commands *) val calc_args = - Scan.option (Parse.$$$ "(" |-- Parse.!!! ((Parse_Spec.xthms1 --| Parse.$$$ ")"))); + Scan.option (@{keyword "("} |-- Parse.!!! ((Parse_Spec.xthms1 --| @{keyword ")"}))); val _ = - Outer_Syntax.command ("also", Keyword.tag_proof Keyword.prf_decl) - "combine calculation and current facts" + Outer_Syntax.command @{command_spec "also"} "combine calculation and current facts" (calc_args >> (Toplevel.proofs' o Calculation.also_cmd)); val _ = - Outer_Syntax.command ("finally", Keyword.tag_proof Keyword.prf_chain) + Outer_Syntax.command @{command_spec "finally"} "combine calculation and current facts, exhibit result" (calc_args >> (Toplevel.proofs' o Calculation.finally_cmd)); val _ = - Outer_Syntax.command ("moreover", Keyword.tag_proof Keyword.prf_decl) - "augment calculation by current facts" + Outer_Syntax.command @{command_spec "moreover"} "augment calculation by current facts" (Scan.succeed (Toplevel.proof' Calculation.moreover)); val _ = - Outer_Syntax.command ("ultimately", Keyword.tag_proof Keyword.prf_chain) + Outer_Syntax.command @{command_spec "ultimately"} "augment calculation by current facts, exhibit result" (Scan.succeed (Toplevel.proof' Calculation.ultimately)); @@ -779,8 +726,7 @@ (* proof navigation *) val _ = - Outer_Syntax.command ("back", Keyword.tag_proof Keyword.prf_script) - "backtracking of proof command" + Outer_Syntax.command @{command_spec "back"} "backtracking of proof command" (Scan.succeed (Toplevel.print o Toplevel.actual_proof Proof_Node.back o Toplevel.skip_proof I)); @@ -792,7 +738,7 @@ (Position.of_properties (Position.default_properties pos props), str)); val _ = - Outer_Syntax.improper_command ("Isabelle.command", Keyword.control) "nested Isabelle command" + Outer_Syntax.improper_command @{command_spec "Isabelle.command"} "nested Isabelle command" (props_text :|-- (fn (pos, str) => (case Outer_Syntax.parse pos str of [tr] => Scan.succeed (K tr) @@ -804,166 +750,168 @@ (** diagnostic commands (for interactive mode only) **) val opt_modes = - Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) []; + Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) []; -val opt_bang = Scan.optional (Parse.$$$ "!" >> K true) false; +val opt_bang = Scan.optional (@{keyword "!"} >> K true) false; val _ = - Outer_Syntax.improper_command ("pretty_setmargin", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "pretty_setmargin"} "change default margin for pretty printing" (Parse.nat >> (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => Pretty.margin_default := n))); val _ = - Outer_Syntax.improper_command ("help", Keyword.diag) "print outer syntax commands" + Outer_Syntax.improper_command @{command_spec "help"} "print outer syntax commands" (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax)); val _ = - Outer_Syntax.improper_command ("print_commands", Keyword.diag) "print outer syntax commands" + Outer_Syntax.improper_command @{command_spec "print_commands"} "print outer syntax commands" (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax)); val _ = - Outer_Syntax.improper_command ("print_configs", Keyword.diag) "print configuration options" + Outer_Syntax.improper_command @{command_spec "print_configs"} "print configuration options" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_configs)); val _ = - Outer_Syntax.improper_command ("print_context", Keyword.diag) "print theory context name" + Outer_Syntax.improper_command @{command_spec "print_context"} "print theory context name" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_context)); val _ = - Outer_Syntax.improper_command ("print_theory", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_theory"} "print logical theory contents (verbose!)" (opt_bang >> (Toplevel.no_timing oo Isar_Cmd.print_theory)); val _ = - Outer_Syntax.improper_command ("print_syntax", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_syntax"} "print inner syntax of context (verbose!)" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_syntax)); val _ = - Outer_Syntax.improper_command ("print_abbrevs", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_abbrevs"} "print constant abbreviation of context" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_abbrevs)); val _ = - Outer_Syntax.improper_command ("print_theorems", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_theorems"} "print theorems of local theory or proof context" (opt_bang >> (Toplevel.no_timing oo Isar_Cmd.print_theorems)); val _ = - Outer_Syntax.improper_command ("print_locales", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_locales"} "print locales of this theory" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_locales)); val _ = - Outer_Syntax.improper_command ("print_classes", Keyword.diag) "print classes of this theory" + Outer_Syntax.improper_command @{command_spec "print_classes"} + "print classes of this theory" (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep (Class.print_classes o Toplevel.context_of))); val _ = - Outer_Syntax.improper_command ("print_locale", Keyword.diag) "print locale of this theory" + Outer_Syntax.improper_command @{command_spec "print_locale"} + "print locale of this theory" (opt_bang -- Parse.position Parse.xname >> (Toplevel.no_timing oo Isar_Cmd.print_locale)); val _ = - Outer_Syntax.improper_command ("print_interps", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_interps"} "print interpretations of locale for this theory or proof context" (Parse.position Parse.xname >> (Toplevel.no_timing oo Isar_Cmd.print_registrations)); val _ = - Outer_Syntax.improper_command ("print_dependencies", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_dependencies"} "print dependencies of locale expression" (opt_bang -- Parse_Spec.locale_expression true >> (Toplevel.no_timing oo Isar_Cmd.print_dependencies)); val _ = - Outer_Syntax.improper_command ("print_attributes", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_attributes"} "print attributes of this theory" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_attributes)); val _ = - Outer_Syntax.improper_command ("print_simpset", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_simpset"} "print context of Simplifier" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_simpset)); val _ = - Outer_Syntax.improper_command ("print_rules", Keyword.diag) "print intro/elim rules" + Outer_Syntax.improper_command @{command_spec "print_rules"} "print intro/elim rules" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_rules)); val _ = - Outer_Syntax.improper_command ("print_trans_rules", Keyword.diag) "print transitivity rules" + Outer_Syntax.improper_command @{command_spec "print_trans_rules"} "print transitivity rules" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_trans_rules)); val _ = - Outer_Syntax.improper_command ("print_methods", Keyword.diag) "print methods of this theory" + Outer_Syntax.improper_command @{command_spec "print_methods"} "print methods of this theory" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_methods)); val _ = - Outer_Syntax.improper_command ("print_antiquotations", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_antiquotations"} "print antiquotations (global)" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_antiquotations)); val _ = - Outer_Syntax.improper_command ("thy_deps", Keyword.diag) "visualize theory dependencies" + Outer_Syntax.improper_command @{command_spec "thy_deps"} "visualize theory dependencies" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.thy_deps)); val _ = - Outer_Syntax.improper_command ("class_deps", Keyword.diag) "visualize class dependencies" + Outer_Syntax.improper_command @{command_spec "class_deps"} "visualize class dependencies" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.class_deps)); val _ = - Outer_Syntax.improper_command ("thm_deps", Keyword.diag) "visualize theorem dependencies" + Outer_Syntax.improper_command @{command_spec "thm_deps"} "visualize theorem dependencies" (Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.thm_deps)); val _ = - Outer_Syntax.improper_command ("print_binds", Keyword.diag) "print term bindings of proof context" + Outer_Syntax.improper_command @{command_spec "print_binds"} "print term bindings of proof context" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_binds)); val _ = - Outer_Syntax.improper_command ("print_facts", Keyword.diag) "print facts of proof context" + Outer_Syntax.improper_command @{command_spec "print_facts"} "print facts of proof context" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_facts)); val _ = - Outer_Syntax.improper_command ("print_cases", Keyword.diag) "print cases of proof context" + Outer_Syntax.improper_command @{command_spec "print_cases"} "print cases of proof context" (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_cases)); val _ = - Outer_Syntax.improper_command ("print_statement", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_statement"} "print theorems as long statements" (opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.print_stmts)); val _ = - Outer_Syntax.improper_command ("thm", Keyword.diag) "print theorems" + Outer_Syntax.improper_command @{command_spec "thm"} "print theorems" (opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.print_thms)); val _ = - Outer_Syntax.improper_command ("prf", Keyword.diag) "print proof terms of theorems" + Outer_Syntax.improper_command @{command_spec "prf"} "print proof terms of theorems" (opt_modes -- Scan.option Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.print_prfs false)); val _ = - Outer_Syntax.improper_command ("full_prf", Keyword.diag) "print full proof terms of theorems" + Outer_Syntax.improper_command @{command_spec "full_prf"} "print full proof terms of theorems" (opt_modes -- Scan.option Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.print_prfs true)); val _ = - Outer_Syntax.improper_command ("prop", Keyword.diag) "read and print proposition" + Outer_Syntax.improper_command @{command_spec "prop"} "read and print proposition" (opt_modes -- Parse.term >> (Toplevel.no_timing oo Isar_Cmd.print_prop)); val _ = - Outer_Syntax.improper_command ("term", Keyword.diag) "read and print term" + Outer_Syntax.improper_command @{command_spec "term"} "read and print term" (opt_modes -- Parse.term >> (Toplevel.no_timing oo Isar_Cmd.print_term)); val _ = - Outer_Syntax.improper_command ("typ", Keyword.diag) "read and print type" + Outer_Syntax.improper_command @{command_spec "typ"} "read and print type" (opt_modes -- Parse.typ >> (Toplevel.no_timing oo Isar_Cmd.print_type)); val _ = - Outer_Syntax.improper_command ("print_codesetup", Keyword.diag) "print code generator setup" + Outer_Syntax.improper_command @{command_spec "print_codesetup"} "print code generator setup" (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep (Code.print_codesetup o Toplevel.theory_of))); val _ = - Outer_Syntax.improper_command ("unused_thms", Keyword.diag) "find unused theorems" + Outer_Syntax.improper_command @{command_spec "unused_thms"} "find unused theorems" (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 Isar_Cmd.unused_thms)); @@ -973,42 +921,42 @@ (** system commands (for interactive mode only) **) val _ = - Outer_Syntax.improper_command ("cd", Keyword.control) "change current working directory" + Outer_Syntax.improper_command @{command_spec "cd"} "change current working directory" (Parse.path >> (fn path => Toplevel.no_timing o Toplevel.imperative (fn () => File.cd path))); val _ = - Outer_Syntax.improper_command ("pwd", Keyword.diag) "print current working directory" + Outer_Syntax.improper_command @{command_spec "pwd"} "print current working directory" (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => writeln (Path.print (File.pwd ()))))); val _ = - Outer_Syntax.improper_command ("use_thy", Keyword.control) "use theory file" + Outer_Syntax.improper_command @{command_spec "use_thy"} "use theory file" (Parse.name >> (fn name => Toplevel.no_timing o Toplevel.imperative (fn () => Thy_Info.use_thy name))); val _ = - Outer_Syntax.improper_command ("remove_thy", Keyword.control) "remove theory from loader database" + Outer_Syntax.improper_command @{command_spec "remove_thy"} "remove theory from loader database" (Parse.name >> (fn name => Toplevel.no_timing o Toplevel.imperative (fn () => Thy_Info.remove_thy name))); val _ = - Outer_Syntax.improper_command ("kill_thy", Keyword.control) + Outer_Syntax.improper_command @{command_spec "kill_thy"} "kill theory -- try to remove from loader database" (Parse.name >> (fn name => Toplevel.no_timing o Toplevel.imperative (fn () => Thy_Info.kill_thy name))); val _ = - Outer_Syntax.improper_command ("display_drafts", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "display_drafts"} "display raw source files with symbols" (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo Isar_Cmd.display_drafts)); val _ = - Outer_Syntax.improper_command ("print_drafts", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "print_drafts"} "print raw source files with symbols" (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo Isar_Cmd.print_drafts)); val _ = (* FIXME tty only *) - Outer_Syntax.improper_command ("pr", Keyword.diag) "print current proof state (if present)" + Outer_Syntax.improper_command @{command_spec "pr"} "print current proof state (if present)" (opt_modes -- Scan.option Parse.nat >> (fn (modes, limit) => Toplevel.no_timing o Toplevel.keep (fn state => (case limit of NONE => () | SOME n => Goal_Display.goals_limit_default := n; @@ -1016,30 +964,109 @@ Print_Mode.with_modes modes (Toplevel.print_state true) state)))); val _ = - Outer_Syntax.improper_command ("disable_pr", Keyword.control) + Outer_Syntax.improper_command @{command_spec "disable_pr"} "disable printing of toplevel state" (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Toplevel.quiet := true))); val _ = - Outer_Syntax.improper_command ("enable_pr", Keyword.control) + Outer_Syntax.improper_command @{command_spec "enable_pr"} "enable printing of toplevel state" (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Toplevel.quiet := false))); val _ = - Outer_Syntax.improper_command ("commit", Keyword.control) + Outer_Syntax.improper_command @{command_spec "commit"} "commit current session to ML database" (Parse.opt_unit >> K (Toplevel.no_timing o Toplevel.imperative Secure.commit)); val _ = - Outer_Syntax.improper_command ("quit", Keyword.control) "quit Isabelle" + Outer_Syntax.improper_command @{command_spec "quit"} "quit Isabelle" (Parse.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative quit))); val _ = - Outer_Syntax.improper_command ("exit", Keyword.control) "exit Isar loop" + Outer_Syntax.improper_command @{command_spec "exit"} "exit Isar loop" (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state => (Context.set_thread_data (try Toplevel.generic_theory_of state); raise Runtime.TERMINATE)))); +val _ = + Outer_Syntax.improper_command @{command_spec "welcome"} "print welcome message" + (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (writeln o Session.welcome))); + + + +(** raw Isar read-eval-print loop **) + +val _ = + Outer_Syntax.improper_command @{command_spec "init_toplevel"} "init toplevel point-of-interest" + (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Isar.init)); + +val _ = + Outer_Syntax.improper_command @{command_spec "linear_undo"} "undo commands" + (Scan.optional Parse.nat 1 >> + (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => Isar.linear_undo n))); + +val _ = + Outer_Syntax.improper_command @{command_spec "undo"} "undo commands (skipping closed proofs)" + (Scan.optional Parse.nat 1 >> + (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => Isar.undo n))); + +val _ = + Outer_Syntax.improper_command @{command_spec "undos_proof"} + "undo commands (skipping closed proofs)" + (Scan.optional Parse.nat 1 >> (fn n => Toplevel.no_timing o + Toplevel.keep (fn state => + if Toplevel.is_proof state then (Isar.undo n; Isar.print ()) else raise Toplevel.UNDEF))); + +val _ = + Outer_Syntax.improper_command @{command_spec "cannot_undo"} + "partial undo -- Proof General legacy" + (Parse.name >> + (fn "end" => Toplevel.no_timing o Toplevel.imperative (fn () => Isar.undo 1) + | txt => Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt)))); + +val _ = + Outer_Syntax.improper_command @{command_spec "kill"} + "kill partial proof or theory development" + (Scan.succeed (Toplevel.no_timing o Toplevel.imperative Isar.kill)); + + + +(** extraction of programs from proofs **) + +val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") []; + +val _ = + Outer_Syntax.command @{command_spec "realizers"} + "specify realizers for primitive axioms / theorems, together with correctness proof" + (Scan.repeat1 (Parse.xname -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >> + (fn xs => Toplevel.theory (fn thy => Extraction.add_realizers + (map (fn (((a, vs), s1), s2) => (Global_Theory.get_thm thy a, (vs, s1, s2))) xs) thy))); + +val _ = + Outer_Syntax.command @{command_spec "realizability"} + "add equations characterizing realizability" + (Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_realizes_eqns)); + +val _ = + Outer_Syntax.command @{command_spec "extract_type"} + "add equations characterizing type of extracted program" + (Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_typeof_eqns)); + +val _ = + Outer_Syntax.command @{command_spec "extract"} "extract terms from proofs" + (Scan.repeat1 (Parse.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy => + Extraction.extract (map (apfst (Global_Theory.get_thm thy)) xs) thy))); + + + +(** end **) + +val _ = + Outer_Syntax.command @{command_spec "end"} "end context" + (Scan.succeed + (Toplevel.exit o Toplevel.end_local_theory o Toplevel.close_target o + Toplevel.end_proof (K Proof.end_notepad))); + end; diff -r ee9cba42d83d -r 875a4657523e src/Pure/Isar/keyword.ML --- a/src/Pure/Isar/keyword.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/Isar/keyword.ML Thu Aug 02 16:17:52 2012 +0200 @@ -43,7 +43,7 @@ val tag_ml: T -> T type spec = string * string list val spec: spec -> T - val command_spec: string * spec -> string * T + val command_spec: (string * spec) * Position.T -> (string * T) * Position.T val get_lexicons: unit -> Scan.lexicon * Scan.lexicon val is_keyword: string -> bool val command_keyword: string -> T option @@ -135,7 +135,7 @@ SOME k => k |> fold tag tags | NONE => error ("Unknown outer syntax keyword kind " ^ quote kind)); -fun command_spec (name, s) = (name, spec s); +fun command_spec ((name, s), pos) = ((name, spec s), pos); diff -r ee9cba42d83d -r 875a4657523e src/Pure/Isar/outer_syntax.ML --- a/src/Pure/Isar/outer_syntax.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/Isar/outer_syntax.ML Thu Aug 02 16:17:52 2012 +0200 @@ -13,7 +13,7 @@ val is_markup: outer_syntax -> Thy_Output.markup -> string -> bool val get_syntax: unit -> (Scan.lexicon * Scan.lexicon) * outer_syntax val check_syntax: unit -> unit - type command_spec = string * Keyword.T + type command_spec = (string * Keyword.T) * Position.T val command: command_spec -> string -> (Toplevel.transition -> Toplevel.transition) parser -> unit val markup_command: Thy_Output.markup -> command_spec -> string -> @@ -49,10 +49,17 @@ {comment: string, markup: Thy_Output.markup option, int_only: bool, - parse: (Toplevel.transition -> Toplevel.transition) parser}; + parse: (Toplevel.transition -> Toplevel.transition) parser, + pos: Position.T, + id: serial}; -fun make_command comment markup int_only parse = - Command {comment = comment, markup = markup, int_only = int_only, parse = parse}; +fun new_command comment markup int_only parse pos = + Command {comment = comment, markup = markup, int_only = int_only, parse = parse, + pos = pos, id = serial ()}; + +fun command_markup def (name, Command {pos, id, ...}) = + Markup.properties (Position.entity_properties_of def id pos) + (Isabelle_Markup.entity Isabelle_Markup.commandN name); (* parse command *) @@ -118,7 +125,7 @@ (** global outer syntax **) -type command_spec = string * Keyword.T; +type command_spec = (string * Keyword.T) * Position.T; local @@ -128,15 +135,18 @@ fun add_command (name, kind) cmd = CRITICAL (fn () => let val thy = ML_Context.the_global_context (); + val Command {pos, ...} = cmd; val _ = (case try (Thy_Header.the_keyword thy) name of SOME spec => if Option.map #1 spec = SOME (Keyword.kind_of kind) then () - else error ("Inconsistent outer syntax keyword declaration " ^ quote name) + else error ("Inconsistent outer syntax keyword declaration " ^ + quote name ^ Position.str_of pos) | NONE => if Context.theory_name thy = Context.PureN then Keyword.define (name, SOME kind) - else error ("Undeclared outer syntax command " ^ quote name)); + else error ("Undeclared outer syntax command " ^ quote name ^ Position.str_of pos)); + val _ = Position.report pos (command_markup true (name, cmd)); in Unsynchronized.change global_outer_syntax (map_commands (fn commands => (if not (Symtab.defined commands name) then () @@ -159,14 +169,14 @@ fun lookup_commands_dynamic () = lookup_commands (! global_outer_syntax); -fun command command_spec comment parse = - add_command command_spec (make_command comment NONE false parse); +fun command (spec, pos) comment parse = + add_command spec (new_command comment NONE false parse pos); -fun markup_command markup command_spec comment parse = - add_command command_spec (make_command comment (SOME markup) false parse); +fun markup_command markup (spec, pos) comment parse = + add_command spec (new_command comment (SOME markup) false parse pos); -fun improper_command command_spec comment parse = - add_command command_spec (make_command comment NONE true parse); +fun improper_command (spec, pos) comment parse = + add_command spec (new_command comment NONE true parse pos); end; @@ -265,7 +275,16 @@ let val commands = lookup_commands outer_syntax; val range_pos = Position.set_range (Token.range toks); - val _ = Position.reports (maps Thy_Syntax.reports_of_token toks); + + fun command_reports tok = + if Token.kind_of tok = Token.Command then + let val name = Token.content_of tok in + (case commands name of + NONE => [] + | SOME cmd => [(Token.position_of tok, command_markup false (name, cmd))]) + end + else []; + val _ = Position.reports (maps Thy_Syntax.reports_of_token toks @ maps command_reports toks); in (case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of [tr] => diff -r ee9cba42d83d -r 875a4657523e src/Pure/ML/ml_antiquote.ML --- a/src/Pure/ML/ml_antiquote.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/ML/ml_antiquote.ML Thu Aug 02 16:17:52 2012 +0200 @@ -190,22 +190,26 @@ fun with_keyword f = Args.theory -- Scan.lift (Parse.position Parse.string) >> (fn (thy, (name, pos)) => - (f (name, Thy_Header.the_keyword thy name) + (f ((name, Thy_Header.the_keyword thy name), pos) handle ERROR msg => error (msg ^ Position.str_of pos))); val _ = Context.>> (Context.map_theory (value (Binding.name "keyword") (with_keyword - (fn (name, NONE) => "Parse.$$$ " ^ ML_Syntax.print_string name - | (name, SOME _) => error ("Expected minor keyword " ^ quote name))) #> + (fn ((name, NONE), _) => "Parse.$$$ " ^ ML_Syntax.print_string name + | ((name, SOME _), pos) => + error ("Expected minor keyword " ^ quote name ^ Position.str_of pos))) #> value (Binding.name "command_spec") (with_keyword - (fn (name, SOME kind) => + (fn ((name, SOME kind), pos) => "Keyword.command_spec " ^ ML_Syntax.atomic - (ML_Syntax.print_pair ML_Syntax.print_string + ((ML_Syntax.print_pair (ML_Syntax.print_pair ML_Syntax.print_string - (ML_Syntax.print_list ML_Syntax.print_string)) (name, kind)) - | (name, NONE) => error ("Expected command keyword " ^ quote name))))); + (ML_Syntax.print_pair ML_Syntax.print_string + (ML_Syntax.print_list ML_Syntax.print_string))) + ML_Syntax.print_position) ((name, kind), pos)) + | ((name, NONE), pos) => + error ("Expected command keyword " ^ quote name ^ Position.str_of pos))))); end; diff -r ee9cba42d83d -r 875a4657523e src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/Proof/extraction.ML Thu Aug 02 16:17:52 2012 +0200 @@ -814,33 +814,6 @@ |> Sign.restore_naming thy end; - -(**** interface ****) - -val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") []; - -val _ = - Outer_Syntax.command ("realizers", Keyword.thy_decl) - "specify realizers for primitive axioms / theorems, together with correctness proof" - (Scan.repeat1 (Parse.xname -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >> - (fn xs => Toplevel.theory (fn thy => add_realizers - (map (fn (((a, vs), s1), s2) => (Global_Theory.get_thm thy a, (vs, s1, s2))) xs) thy))); - -val _ = - Outer_Syntax.command ("realizability", Keyword.thy_decl) - "add equations characterizing realizability" - (Scan.repeat1 Parse.string >> (Toplevel.theory o add_realizes_eqns)); - -val _ = - Outer_Syntax.command ("extract_type", Keyword.thy_decl) - "add equations characterizing type of extracted program" - (Scan.repeat1 Parse.string >> (Toplevel.theory o add_typeof_eqns)); - -val _ = - Outer_Syntax.command ("extract", Keyword.thy_decl) "extract terms from proofs" - (Scan.repeat1 (Parse.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy => - extract (map (apfst (Global_Theory.get_thm thy)) xs) thy))); - val etype_of = etype_of o add_syntax; end; diff -r ee9cba42d83d -r 875a4657523e src/Pure/ProofGeneral/proof_general_emacs.ML --- a/src/Pure/ProofGeneral/proof_general_emacs.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/ProofGeneral/proof_general_emacs.ML Thu Aug 02 16:17:52 2012 +0200 @@ -189,31 +189,37 @@ (* additional outer syntax for Isar *) val _ = - Outer_Syntax.improper_command ("ProofGeneral.pr", Keyword.diag) "print state (internal)" + Outer_Syntax.improper_command + (("ProofGeneral.pr", Keyword.diag), Position.none) "print state (internal)" (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)))); val _ = (*undo without output -- historical*) - Outer_Syntax.improper_command ("ProofGeneral.undo", Keyword.control) "(internal)" + Outer_Syntax.improper_command + (("ProofGeneral.undo", Keyword.control), Position.none) "(internal)" (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Isar.undo 1))); val _ = - Outer_Syntax.improper_command ("ProofGeneral.restart", Keyword.control) "(internal)" + Outer_Syntax.improper_command + (("ProofGeneral.restart", Keyword.control), Position.none) "(internal)" (Parse.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative restart))); val _ = - Outer_Syntax.improper_command ("ProofGeneral.kill_proof", Keyword.control) "(internal)" + Outer_Syntax.improper_command + (("ProofGeneral.kill_proof", Keyword.control), Position.none) "(internal)" (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => (Isar.kill_proof (); tell_clear_goals ())))); val _ = - Outer_Syntax.improper_command ("ProofGeneral.inform_file_processed", Keyword.control) "(internal)" + Outer_Syntax.improper_command + (("ProofGeneral.inform_file_processed", Keyword.control), Position.none) "(internal)" (Parse.name >> (fn file => Toplevel.no_timing o Toplevel.imperative (fn () => inform_file_processed file))); val _ = - Outer_Syntax.improper_command ("ProofGeneral.inform_file_retracted", Keyword.control) "(internal)" + Outer_Syntax.improper_command + (("ProofGeneral.inform_file_retracted", Keyword.control), Position.none) "(internal)" (Parse.name >> (Toplevel.no_timing oo (fn file => Toplevel.imperative (fn () => inform_file_retracted file)))); diff -r ee9cba42d83d -r 875a4657523e src/Pure/ProofGeneral/proof_general_pgip.ML --- a/src/Pure/ProofGeneral/proof_general_pgip.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/ProofGeneral/proof_general_pgip.ML Thu Aug 02 16:17:52 2012 +0200 @@ -1032,7 +1032,8 @@ (* Extra command for embedding prover-control inside document (obscure/debug usage). *) val _ = - Outer_Syntax.improper_command ("ProofGeneral.process_pgip", Keyword.control) "(internal)" + Outer_Syntax.improper_command + (("ProofGeneral.process_pgip", Keyword.control), Position.none) "(internal)" (Parse.text >> (Toplevel.no_timing oo (fn txt => Toplevel.imperative (fn () => if print_mode_active proof_general_emacsN diff -r ee9cba42d83d -r 875a4657523e src/Pure/Pure.thy --- a/src/Pure/Pure.thy Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/Pure.thy Thu Aug 02 16:17:52 2012 +0200 @@ -1,4 +1,94 @@ theory Pure + keywords + "!!" "!" "%" "(" ")" "+" "," "--" ":" "::" ";" "<" "<=" "=" "==" + "=>" "?" "[" "\" "\" "\" + "\" "\" "]" "advanced" "and" "assumes" + "attach" "begin" "binder" "constrains" "defines" "fixes" "for" + "identifier" "if" "imports" "in" "includes" "infix" "infixl" + "infixr" "is" "keywords" "notes" "obtains" "open" "output" + "overloaded" "pervasive" "shows" "structure" "unchecked" "uses" + "where" "|" + and "header" :: diag + and "chapter" :: thy_heading1 + and "section" :: thy_heading2 + and "subsection" :: thy_heading3 + and "subsubsection" :: thy_heading4 + and "text" "text_raw" :: thy_decl + and "sect" :: prf_heading2 % "proof" + and "subsect" :: prf_heading3 % "proof" + and "subsubsect" :: prf_heading4 % "proof" + and "txt" "txt_raw" :: prf_decl % "proof" + and "classes" "classrel" "default_sort" "typedecl" "type_synonym" + "nonterminal" "arities" "judgment" "consts" "syntax" "no_syntax" + "translations" "no_translations" "axioms" "defs" "definition" + "abbreviation" "type_notation" "no_type_notation" "notation" + "no_notation" "axiomatization" "theorems" "lemmas" "declare" + "hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl + and "use" "ML" :: thy_decl % "ML" + and "ML_prf" :: prf_decl % "proof" (* FIXME % "ML" ?? *) + and "ML_val" "ML_command" :: diag % "ML" + and "setup" "local_setup" "attribute_setup" "method_setup" + "declaration" "syntax_declaration" "simproc_setup" + "parse_ast_translation" "parse_translation" "print_translation" + "typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML" + and "bundle" :: thy_decl + and "include" "including" :: prf_decl + and "print_bundles" :: diag + and "context" "locale" :: thy_decl + and "sublocale" "interpretation" :: thy_schematic_goal + and "interpret" :: prf_goal % "proof" (* FIXME schematic? *) + and "class" :: thy_decl + and "subclass" :: thy_goal + and "instantiation" :: thy_decl + and "instance" :: thy_goal + and "overloading" :: thy_decl + and "code_datatype" :: thy_decl + and "theorem" "lemma" "corollary" :: thy_goal + and "schematic_theorem" "schematic_lemma" "schematic_corollary" :: thy_schematic_goal + and "notepad" :: thy_decl + and "have" "hence" :: prf_goal % "proof" + and "show" "thus" :: prf_asm_goal % "proof" + and "then" "from" "with" :: prf_chain % "proof" + and "note" "using" "unfolding" :: prf_decl % "proof" + and "fix" "assume" "presume" "def" :: prf_asm % "proof" + and "obtain" "guess" :: prf_asm_goal % "proof" + and "let" "write" :: prf_decl % "proof" + and "case" :: prf_asm % "proof" + and "{" :: prf_open % "proof" + and "}" :: prf_close % "proof" + and "next" :: prf_block % "proof" + and "qed" :: qed_block % "proof" + and "by" ".." "." "done" "sorry" :: "qed" % "proof" + and "oops" :: qed_global % "proof" + and "defer" "prefer" "apply" "apply_end" :: prf_script % "proof" + and "proof" :: prf_block % "proof" + and "also" "moreover" :: prf_decl % "proof" + and "finally" "ultimately" :: prf_chain % "proof" + and "back" :: prf_script % "proof" + and "Isabelle.command" :: control + and "pretty_setmargin" "help" "print_commands" "print_configs" + "print_context" "print_theory" "print_syntax" "print_abbrevs" + "print_theorems" "print_locales" "print_classes" "print_locale" + "print_interps" "print_dependencies" "print_attributes" + "print_simpset" "print_rules" "print_trans_rules" "print_methods" + "print_antiquotations" "thy_deps" "class_deps" "thm_deps" + "print_binds" "print_facts" "print_cases" "print_statement" "thm" + "prf" "full_prf" "prop" "term" "typ" "print_codesetup" "unused_thms" + :: diag + and "cd" :: control + and "pwd" :: diag + and "use_thy" "remove_thy" "kill_thy" :: control + and "display_drafts" "print_drafts" "pr" :: diag + and "disable_pr" "enable_pr" "commit" "quit" "exit" :: control + and "welcome" :: diag + and "init_toplevel" "linear_undo" "undo" "undos_proof" "cannot_undo" "kill" :: control + and "end" :: thy_end % "theory" + and "realizers" "realizability" "extract_type" "extract" :: thy_decl + and "find_theorems" "find_consts" :: diag + uses + "Isar/isar_syn.ML" + "Tools/find_theorems.ML" + "Tools/find_consts.ML" begin section {* Further content for the Pure theory *} diff -r ee9cba42d83d -r 875a4657523e src/Pure/ROOT --- a/src/Pure/ROOT Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/ROOT Thu Aug 02 16:17:52 2012 +0200 @@ -105,7 +105,6 @@ "Isar/expression.ML" "Isar/generic_target.ML" "Isar/isar_cmd.ML" - "Isar/isar_syn.ML" "Isar/keyword.ML" "Isar/local_defs.ML" "Isar/local_theory.ML" @@ -197,8 +196,6 @@ "Thy/thy_load.ML" "Thy/thy_output.ML" "Thy/thy_syntax.ML" - "Tools/find_consts.ML" - "Tools/find_theorems.ML" "Tools/named_thms.ML" "Tools/xml_syntax.ML" "assumption.ML" diff -r ee9cba42d83d -r 875a4657523e src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/ROOT.ML Thu Aug 02 16:17:52 2012 +0200 @@ -259,7 +259,6 @@ use "Isar/rule_insts.ML"; use "Thy/thm_deps.ML"; use "Isar/isar_cmd.ML"; -use "Isar/isar_syn.ML"; use "subgoal.ML"; @@ -283,9 +282,6 @@ use "Tools/xml_syntax.ML"; -use "Tools/find_theorems.ML"; -use "Tools/find_consts.ML"; - (* configuration for Proof General *) @@ -302,11 +298,11 @@ use "ProofGeneral/pgip_parser.ML"; -use "ProofGeneral/pgip_tests.ML"; - use "ProofGeneral/proof_general_pgip.ML"; use "ProofGeneral/proof_general_emacs.ML"; use "pure_setup.ML"; +use "ProofGeneral/pgip_tests.ML"; + diff -r ee9cba42d83d -r 875a4657523e src/Pure/System/build.scala --- a/src/Pure/System/build.scala Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/System/build.scala Thu Aug 02 16:17:52 2012 +0200 @@ -41,32 +41,50 @@ val empty: Queue = new Queue() } - final class Queue private(graph: Graph[String, Info] = Graph.string) + final class Queue private(graph: Graph[String, Option[Info]] = Graph.string) extends PartialFunction[String, Info] { - def apply(name: String): Info = graph.get_node(name) - def isDefinedAt(name: String): Boolean = graph.defined(name) + def apply(name: String): Info = graph.get_node(name).get + def isDefinedAt(name: String): Boolean = + graph.defined(name) && graph.get_node(name).isDefined def is_inner(name: String): Boolean = !graph.is_maximal(name) def is_empty: Boolean = graph.is_empty def + (name: String, info: Info): Queue = + { + val parents = info.parent.toList + + val graph1 = (graph /: (name :: parents))(_.default_node(_, None)) + if (graph1.get_node(name).isDefined) + error("Duplicate session: " + quote(name)) + new Queue( - try { graph.new_node(name, info).add_deps_acyclic(name, info.parent.toList) } + try { graph1.map_node(name, _ => Some(info)).add_deps_acyclic(name, parents) } catch { - case _: Graph.Duplicate[_] => error("Duplicate session: " + quote(name)) case exn: Graph.Cycles[_] => error(cat_lines(exn.cycles.map(cycle => "Cyclic session dependency of " + cycle.map(c => quote(c.toString)).mkString(" via ")))) }) + } def - (name: String): Queue = new Queue(graph.del_node(name)) def required(all_sessions: Boolean, session_groups: List[String], sessions: List[String]) : (List[String], Queue) = { + (for { + (name, (Some(info), _)) <- graph.entries + if info.parent.isDefined; parent = info.parent.get + if !isDefinedAt(parent) + } yield parent + " (for " + name + ")").toList match + { + case Nil => + case bad => error("Bad parent session(s): " + commas(bad)) + } + sessions.filterNot(isDefinedAt(_)) match { case Nil => case bad => error("Undefined session(s): " + commas_quote(bad)) @@ -90,8 +108,8 @@ def dequeue(skip: String => Boolean): Option[(String, Info)] = { val it = graph.entries.dropWhile( - { case (name, (_, (deps, _))) => !deps.isEmpty || skip(name) }) - if (it.hasNext) { val (name, (info, _)) = it.next; Some((name, info)) } + { case (name, (info, (deps, _))) => !deps.isEmpty || info.isEmpty || skip(name) }) + if (it.hasNext) { val (name, (info, _)) = it.next; Some((name, info.get)) } else None } @@ -169,7 +187,7 @@ private def is_pure(name: String): Boolean = name == "RAW" || name == "Pure" - private def sessions_root(options: Options, dir: Path, root: Path, queue: Session.Queue) + private def sessions_root(options: Options, dir: Path, queue: Session.Queue, root: Path) : Session.Queue = { (queue /: Parser.parse_entries(root))((queue1, entry) => @@ -183,12 +201,10 @@ } else entry.parent match { - case Some(parent_name) if queue1.isDefinedAt(parent_name) => - val full_name = - if (entry.this_name) entry.base_name - else parent_name + "-" + entry.base_name - full_name - case _ => error("Bad parent session") + case None => error("Missing parent session") + case Some(parent_name) => + if (entry.this_name) entry.base_name + else parent_name + "-" + entry.base_name } val path = @@ -219,46 +235,21 @@ }) } - private def sessions_dir(options: Options, strict: Boolean, dir: Path, queue: Session.Queue) + private def sessions_dir(options: Options, strict: Boolean, queue: Session.Queue, dir: Path) : Session.Queue = { val root = dir + ROOT - if (root.is_file) sessions_root(options, dir, root, queue) + if (root.is_file) sessions_root(options, dir, queue, root) else if (strict) error("Bad session root file: " + root.toString) else queue } - private def sessions_catalog(options: Options, dir: Path, catalog: Path, queue: Session.Queue) - : Session.Queue = - { - val dirs = - split_lines(File.read(catalog)).filterNot(line => line == "" || line.startsWith("#")) - (queue /: dirs)((queue1, dir1) => - try { - val dir2 = dir + Path.explode(dir1) - if (dir2.is_dir) sessions_dir(options, true, dir2, queue1) - else error("Bad session directory: " + dir2.toString) - } - catch { - case ERROR(msg) => - error(msg + "\nThe error(s) above occurred in session catalog " + catalog.toString) - }) - } - def find_sessions(options: Options, more_dirs: List[Path]): Session.Queue = { - var queue = Session.Queue.empty - - for (dir <- Isabelle_System.components()) { - queue = sessions_dir(options, false, dir, queue) - - val catalog = dir + SESSIONS - if (catalog.is_file) queue = sessions_catalog(options, dir, catalog, queue) - } - - for (dir <- more_dirs) queue = sessions_dir(options, true, dir, queue) - - queue + val queue1 = + (Session.Queue.empty /: Isabelle_System.components())(sessions_dir(options, false, _, _)) + val queue2 = (queue1 /: more_dirs)(sessions_dir(options, true, _, _)) + queue2 } diff -r ee9cba42d83d -r 875a4657523e src/Pure/System/isar.ML --- a/src/Pure/System/isar.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/System/isar.ML Thu Aug 02 16:17:52 2012 +0200 @@ -151,51 +151,4 @@ toplevel_loop TextIO.stdIn {init = true, welcome = true, sync = false, secure = Secure.is_secure ()}; - - -(** command syntax **) - -local - -val op >> = Scan.>>; - -in - -(* global history *) - -val _ = - Outer_Syntax.improper_command ("init_toplevel", Keyword.control) "init toplevel point-of-interest" - (Scan.succeed (Toplevel.no_timing o Toplevel.imperative init)); - -val _ = - Outer_Syntax.improper_command ("linear_undo", Keyword.control) "undo commands" - (Scan.optional Parse.nat 1 >> - (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => linear_undo n))); - -val _ = - Outer_Syntax.improper_command ("undo", Keyword.control) "undo commands (skipping closed proofs)" - (Scan.optional Parse.nat 1 >> - (fn n => Toplevel.no_timing o Toplevel.imperative (fn () => undo n))); - -val _ = - Outer_Syntax.improper_command ("undos_proof", Keyword.control) - "undo commands (skipping closed proofs)" - (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 _ = - Outer_Syntax.improper_command ("cannot_undo", Keyword.control) - "partial undo -- Proof General legacy" - (Parse.name >> - (fn "end" => Toplevel.no_timing o Toplevel.imperative (fn () => undo 1) - | txt => Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt)))); - -val _ = - Outer_Syntax.improper_command ("kill", Keyword.control) - "kill partial proof or theory development" - (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill)); - end; - -end; diff -r ee9cba42d83d -r 875a4657523e src/Pure/System/session.ML --- a/src/Pure/System/session.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/System/session.ML Thu Aug 02 16:17:52 2012 +0200 @@ -47,10 +47,6 @@ "Unofficial version of " ^ name () ^ " (" ^ Distribution.version ^ ")" ^ (if Distribution.changelog <> "" then "\nSee also " ^ Distribution.changelog else ""); -val _ = - Outer_Syntax.improper_command ("welcome", Keyword.diag) "print welcome message" - (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (writeln o welcome))); - (* add_path *) diff -r ee9cba42d83d -r 875a4657523e src/Pure/Tools/find_consts.ML --- a/src/Pure/Tools/find_consts.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/Tools/find_consts.ML Thu Aug 02 16:17:52 2012 +0200 @@ -134,7 +134,7 @@ in val _ = - Outer_Syntax.improper_command ("find_consts", Keyword.diag) "search constants by type pattern" + Outer_Syntax.improper_command @{command_spec "find_consts"} "search constants by type pattern" (Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion) >> (fn spec => Toplevel.no_timing o Toplevel.keep (fn state => find_consts (Toplevel.context_of state) spec))); diff -r ee9cba42d83d -r 875a4657523e src/Pure/Tools/find_theorems.ML --- a/src/Pure/Tools/find_theorems.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/Tools/find_theorems.ML Thu Aug 02 16:17:52 2012 +0200 @@ -615,7 +615,7 @@ val query_parser = Scan.repeat (((Scan.option Parse.minus >> is_none) -- criterion)); val _ = - Outer_Syntax.improper_command ("find_theorems", Keyword.diag) + Outer_Syntax.improper_command @{command_spec "find_theorems"} "print theorems meeting specified criteria" (options -- query_parser >> (fn ((opt_lim, rem_dups), spec) => diff -r ee9cba42d83d -r 875a4657523e src/Pure/pure_setup.ML --- a/src/Pure/pure_setup.ML Thu Aug 02 10:10:29 2012 +0200 +++ b/src/Pure/pure_setup.ML Thu Aug 02 16:17:52 2012 +0200 @@ -14,6 +14,14 @@ (* the Pure theory *) +val _ = + Outer_Syntax.command + (("theory", Keyword.tag_theory Keyword.thy_begin), Position.none) "begin theory context" + (Thy_Header.args >> (fn header => + Toplevel.print o + Toplevel.init_theory + (fn () => Thy_Info.toplevel_begin_theory (Thy_Load.get_master_path ()) header))); + Unsynchronized.setmp Multithreading.max_threads 1 use_thy "Pure"; Context.set_thread_data NONE;