# HG changeset patch # User wenzelm # Date 1343903814 -7200 # Node ID 91281e9472d8a21ce180a1deaf3913c458842c9f # Parent 33f00ce23e63daf327a58ca3ca596589e9daea6c more official command specifications, including source position; diff -r 33f00ce23e63 -r 91281e9472d8 src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/Isar/isar_syn.ML Thu Aug 02 12:36:54 2012 +0200 @@ -989,6 +989,75 @@ (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 **) diff -r 33f00ce23e63 -r 91281e9472d8 src/Pure/Isar/keyword.ML --- a/src/Pure/Isar/keyword.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/Isar/keyword.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/Isar/outer_syntax.ML --- a/src/Pure/Isar/outer_syntax.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/Isar/outer_syntax.ML Thu Aug 02 12:36:54 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 -> @@ -118,25 +118,26 @@ (** global outer syntax **) -type command_spec = string * Keyword.T; +type command_spec = (string * Keyword.T) * Position.T; local (*synchronized wrt. Keywords*) val global_outer_syntax = Unsynchronized.ref empty_outer_syntax; -fun add_command (name, kind) cmd = CRITICAL (fn () => +fun add_command ((name, kind), pos) cmd = CRITICAL (fn () => let val thy = ML_Context.the_global_context (); 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)); in Unsynchronized.change global_outer_syntax (map_commands (fn commands => (if not (Symtab.defined commands name) then () diff -r 33f00ce23e63 -r 91281e9472d8 src/Pure/ML/ml_antiquote.ML --- a/src/Pure/ML/ml_antiquote.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/ML/ml_antiquote.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/Proof/extraction.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/ProofGeneral/proof_general_emacs.ML --- a/src/Pure/ProofGeneral/proof_general_emacs.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/ProofGeneral/proof_general_emacs.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/ProofGeneral/proof_general_pgip.ML --- a/src/Pure/ProofGeneral/proof_general_pgip.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/ProofGeneral/proof_general_pgip.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/Pure.thy --- a/src/Pure/Pure.thy Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/Pure.thy Thu Aug 02 12:36:54 2012 +0200 @@ -80,8 +80,15 @@ 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" - uses "Isar/isar_syn.ML" + 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 33f00ce23e63 -r 91281e9472d8 src/Pure/ROOT --- a/src/Pure/ROOT Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/ROOT Thu Aug 02 12:36:54 2012 +0200 @@ -196,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 33f00ce23e63 -r 91281e9472d8 src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/ROOT.ML Thu Aug 02 12:36:54 2012 +0200 @@ -282,9 +282,6 @@ use "Tools/xml_syntax.ML"; -use "Tools/find_theorems.ML"; -use "Tools/find_consts.ML"; - (* configuration for Proof General *) diff -r 33f00ce23e63 -r 91281e9472d8 src/Pure/System/isar.ML --- a/src/Pure/System/isar.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/System/isar.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/System/session.ML --- a/src/Pure/System/session.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/System/session.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/Tools/find_consts.ML --- a/src/Pure/Tools/find_consts.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/Tools/find_consts.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/Tools/find_theorems.ML --- a/src/Pure/Tools/find_theorems.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/Tools/find_theorems.ML Thu Aug 02 12:36:54 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 33f00ce23e63 -r 91281e9472d8 src/Pure/pure_setup.ML --- a/src/Pure/pure_setup.ML Thu Aug 02 11:32:23 2012 +0200 +++ b/src/Pure/pure_setup.ML Thu Aug 02 12:36:54 2012 +0200 @@ -15,7 +15,8 @@ (* the Pure theory *) val _ = - Outer_Syntax.command ("theory", Keyword.tag_theory Keyword.thy_begin) "begin theory context" + 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