# HG changeset patch # User wenzelm # Date 927575833 -7200 # Node ID f342449d73cad13291a54165acc9bf5f7678b4f2 # Parent 5e82c71967898d42d377e3a4a0b588bddb5975c4 outer syntax keyword classification; no open OuterParse; diff -r 5e82c7196789 -r f342449d73ca src/HOL/Tools/datatype_package.ML --- a/src/HOL/Tools/datatype_package.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/HOL/Tools/datatype_package.ML Mon May 24 21:57:13 1999 +0200 @@ -666,11 +666,11 @@ (* outer syntax *) -local open OuterParse in +local structure P = OuterParse and K = OuterSyntax.Keyword in val datatype_decl = - Scan.option ($$$ "(" |-- name --| $$$ ")") -- type_args -- name -- opt_infix -- - ($$$ "=" |-- enum1 "|" (name -- Scan.repeat typ -- opt_mixfix)); + Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- P.type_args -- P.name -- P.opt_infix -- + (P.$$$ "=" |-- P.enum1 "|" (P.name -- Scan.repeat P.typ -- P.opt_mixfix)); fun mk_datatype args = let @@ -679,20 +679,20 @@ in #1 o add_datatype false names specs end; val datatypeP = - OuterSyntax.command "datatype" "define inductive datatypes" - (enum1 "and" datatype_decl >> (Toplevel.theory o mk_datatype)); + OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl + (P.and_list1 datatype_decl >> (Toplevel.theory o mk_datatype)); val rep_datatype_decl = - Scan.option (Scan.repeat1 name) -- - Scan.optional ($$$ "distinct" |-- !!! (and_list1 xthms1)) [] -- - Scan.optional ($$$ "inject" |-- !!! (and_list1 xthms1)) [] -- - ($$$ "induction" |-- !!! xthm); + Scan.option (Scan.repeat1 P.name) -- + Scan.optional (P.$$$ "distinct" |-- P.!!! (P.and_list1 P.xthms1)) [] -- + Scan.optional (P.$$$ "inject" |-- P.!!! (P.and_list1 P.xthms1)) [] -- + (P.$$$ "induction" |-- P.!!! P.xthm); fun mk_rep_datatype (((opt_ts, dss), iss), ind) = #1 o rep_datatype opt_ts dss iss ind; val rep_datatypeP = - OuterSyntax.command "rep_datatype" "represent existing types inductively" + OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_decl (rep_datatype_decl >> (Toplevel.theory o mk_rep_datatype)); diff -r 5e82c7196789 -r f342449d73ca src/HOL/Tools/inductive_package.ML --- a/src/HOL/Tools/inductive_package.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/HOL/Tools/inductive_package.ML Mon May 24 21:57:13 1999 +0200 @@ -704,20 +704,23 @@ (* outer syntax *) -local open OuterParse in +local structure P = OuterParse and K = OuterSyntax.Keyword in fun mk_ind coind (((sets, (atts, intrs)), monos), con_defs) = - #1 o add_inductive true coind sets atts (map triple_swap intrs) monos con_defs; + #1 o add_inductive true coind sets atts (map P.triple_swap intrs) monos con_defs; fun ind_decl coind = - Scan.repeat1 term -- - ($$$ "intrs" |-- !!! (opt_attribs -- Scan.repeat1 (opt_thm_name ":" -- term))) -- - Scan.optional ($$$ "monos" |-- !!! xthms1) [] -- - Scan.optional ($$$ "con_defs" |-- !!! xthms1) [] + Scan.repeat1 P.term -- + (P.$$$ "intrs" |-- P.!!! (P.opt_attribs -- Scan.repeat1 (P.opt_thm_name ":" -- P.term))) -- + Scan.optional (P.$$$ "monos" |-- P.!!! P.xthms1) [] -- + Scan.optional (P.$$$ "con_defs" |-- P.!!! P.xthms1) [] >> (Toplevel.theory o mk_ind coind); -val inductiveP = OuterSyntax.command "inductive" "define inductive sets" (ind_decl false); -val coinductiveP = OuterSyntax.command "coinductive" "define coinductive sets" (ind_decl true); +val inductiveP = + OuterSyntax.command "inductive" "define inductive sets" K.thy_decl (ind_decl false); + +val coinductiveP = + OuterSyntax.command "coinductive" "define coinductive sets" K.thy_decl (ind_decl true); val _ = OuterSyntax.add_keywords ["intrs", "monos", "con_defs"]; val _ = OuterSyntax.add_parsers [inductiveP, coinductiveP]; diff -r 5e82c7196789 -r f342449d73ca src/HOL/Tools/primrec_package.ML --- a/src/HOL/Tools/primrec_package.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/HOL/Tools/primrec_package.ML Mon May 24 21:57:13 1999 +0200 @@ -261,16 +261,16 @@ (* outer syntax *) -local open OuterParse in +local structure P = OuterParse and K = OuterSyntax.Keyword in val primrec_decl = - Scan.optional ($$$ "(" |-- name --| $$$ ")") "" -- - Scan.repeat1 (opt_thm_name ":" -- term); + Scan.optional (P.$$$ "(" |-- P.name --| P.$$$ ")") "" -- + Scan.repeat1 (P.opt_thm_name ":" -- P.term); val primrecP = - OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" + OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl (primrec_decl >> (fn (alt_name, eqns) => - Toplevel.theory (#1 o add_primrec alt_name (map triple_swap eqns)))); + Toplevel.theory (#1 o add_primrec alt_name (map P.triple_swap eqns)))); val _ = OuterSyntax.add_parsers [primrecP]; diff -r 5e82c7196789 -r f342449d73ca src/HOL/Tools/recdef_package.ML --- a/src/HOL/Tools/recdef_package.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/HOL/Tools/recdef_package.ML Mon May 24 21:57:13 1999 +0200 @@ -136,26 +136,26 @@ (* outer syntax *) -local open OuterParse in +local structure P = OuterParse and K = OuterSyntax.Keyword in val recdef_decl = - name -- term -- Scan.repeat1 term -- - Scan.optional ($$$ "(" |-- $$$ "congs" |-- !!! (xthms1 --| $$$ ")")) [] -- - Scan.option ($$$ "(" |-- $$$ "simpset" |-- !!! (name --| $$$ ")")) + P.name -- P.term -- Scan.repeat1 P.term -- + Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (P.xthms1 --| P.$$$ ")")) [] -- + Scan.option (P.$$$ "(" |-- P.$$$ "simpset" |-- P.!!! (P.name --| P.$$$ ")")) >> (fn ((((f, R), eqs), congs), ss) => #1 o add_recdef_x f R eqs ss congs); val recdefP = - OuterSyntax.command "recdef" "define general recursive functions (TFL)" + OuterSyntax.command "recdef" "define general recursive functions (TFL)" K.thy_decl (recdef_decl >> Toplevel.theory); val defer_recdef_decl = - name -- Scan.repeat1 term -- - Scan.optional ($$$ "(" |-- $$$ "congs" |-- !!! (xthms1 --| $$$ ")")) [] + P.name -- Scan.repeat1 P.term -- + Scan.optional (P.$$$ "(" |-- P.$$$ "congs" |-- P.!!! (P.xthms1 --| P.$$$ ")")) [] >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs); val defer_recdefP = - OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" + OuterSyntax.command "defer_recdef" "defer general recursive functions (TFL)" K.thy_decl (defer_recdef_decl >> Toplevel.theory); val _ = OuterSyntax.add_keywords ["congs", "simpset"]; diff -r 5e82c7196789 -r f342449d73ca src/HOL/Tools/record_package.ML --- a/src/HOL/Tools/record_package.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/HOL/Tools/record_package.ML Mon May 24 21:57:13 1999 +0200 @@ -878,14 +878,14 @@ (* outer syntax *) -local open OuterParse in +local structure P = OuterParse and K = OuterSyntax.Keyword in val record_decl = - type_args -- name -- ($$$ "=" |-- Scan.option (typ --| $$$ "+") - -- Scan.repeat1 (name -- ($$$ "::" |-- typ))); + P.type_args -- P.name -- (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") + -- Scan.repeat1 (P.name -- (P.$$$ "::" |-- P.typ))); val recordP = - OuterSyntax.command "record" "define extensible record" + OuterSyntax.command "record" "define extensible record" K.thy_decl (record_decl >> (fn (x, (y, z)) => Toplevel.theory (#1 o add_record x y z))); val _ = OuterSyntax.add_parsers [recordP]; diff -r 5e82c7196789 -r f342449d73ca src/HOL/Tools/typedef_package.ML --- a/src/HOL/Tools/typedef_package.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/HOL/Tools/typedef_package.ML Mon May 24 21:57:13 1999 +0200 @@ -205,24 +205,26 @@ (** outer syntax **) -local open OuterParse in +local structure P = OuterParse and K = OuterSyntax.Keyword in val typedeclP = - OuterSyntax.command "typedecl" "HOL type declaration" - (type_args -- name -- opt_mixfix >> (fn ((vs, t), mx) => + OuterSyntax.command "typedecl" "HOL type declaration" K.thy_decl + (P.type_args -- P.name -- P.opt_mixfix >> (fn ((vs, t), mx) => Toplevel.theory (add_typedecls [(t, vs, mx)]))); + val typedef_proof_decl = - Scan.option ($$$ "(" |-- name --| $$$ ")") -- - (type_args -- name) -- opt_infix -- ($$$ "=" |-- term); + Scan.option (P.$$$ "(" |-- P.name --| P.$$$ ")") -- + (P.type_args -- P.name) -- P.opt_infix -- (P.$$$ "=" |-- P.term); fun mk_typedef_proof (((opt_name, (vs, t)), mx), A) = typedef_proof (if_none opt_name (Syntax.type_name t mx)) (t, vs, mx) A; val typedefP = - OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)" + OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)" K.thy_goal (typedef_proof_decl >> (Toplevel.print oo (Toplevel.theory_to_proof o mk_typedef_proof))); + val _ = OuterSyntax.add_parsers [typedeclP, typedefP]; end; diff -r 5e82c7196789 -r f342449d73ca src/HOLCF/IOA/meta_theory/ioa_package.ML --- a/src/HOLCF/IOA/meta_theory/ioa_package.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/HOLCF/IOA/meta_theory/ioa_package.ML Mon May 24 21:57:13 1999 +0200 @@ -523,42 +523,45 @@ (* parsers *) -local open OuterParse in +local structure P = OuterParse and K = OuterSyntax.Keyword in -val actionlist = list1 term; -val inputslist = $$$ "inputs" |-- actionlist; -val outputslist = $$$ "outputs" |-- actionlist; -val internalslist = $$$ "internals" |-- actionlist; -val stateslist = $$$ "states" |-- Scan.repeat1 (name --| $$$ "::" -- typ); -val initial = $$$ "initially" |-- term; -val assign_list = list1 (name --| $$$ ":=" -- term); -val pre = $$$ "pre" |-- term; -val post = $$$ "post" |-- assign_list; +val actionlist = P.list1 P.term; +val inputslist = P.$$$ "inputs" |-- actionlist; +val outputslist = P.$$$ "outputs" |-- actionlist; +val internalslist = P.$$$ "internals" |-- actionlist; +val stateslist = P.$$$ "states" |-- Scan.repeat1 (P.name --| P.$$$ "::" -- P.typ); +val initial = P.$$$ "initially" |-- P.term; +val assign_list = P.list1 (P.name --| P.$$$ ":=" -- P.term); +val pre = P.$$$ "pre" |-- P.term; +val post = P.$$$ "post" |-- assign_list; val pre1 = (pre -- (Scan.optional post [])) >> mk_trans_of_prepost; val post1 = ((Scan.optional pre "True") -- post) >> mk_trans_of_prepost; -val transrel = ($$$ "transrel" |-- term) >> mk_trans_of_rel; -val transition = term -- (transrel || pre1 || post1); -val translist = $$$ "transitions" |-- Scan.repeat1 transition; +val transrel = (P.$$$ "transrel" |-- P.term) >> mk_trans_of_rel; +val transition = P.term -- (transrel || pre1 || post1); +val translist = P.$$$ "transitions" |-- Scan.repeat1 transition; val ioa_decl = - (name -- ($$$ "=" |-- - ($$$ "signature" |-- - ($$$ "actions" |-- - (typ -- + (P.name -- (P.$$$ "=" |-- + (P.$$$ "signature" |-- + (P.$$$ "actions" |-- + (P.typ -- (Scan.optional inputslist []) -- (Scan.optional outputslist []) -- (Scan.optional internalslist []) -- stateslist -- (Scan.optional initial "True") -- translist))))) >> mk_ioa_decl || - (name -- ($$$ "=" |-- ($$$ "compose" |-- list1 name))) >> mk_composition_decl || - (name -- ($$$ "=" |-- ($$$ "hide" |-- list1 term -- ($$$ "in" |-- name)))) >> mk_hiding_decl || - (name -- ($$$ "=" |-- ($$$ "restrict" |-- name -- ($$$ "to" |-- list1 term)))) + (P.name -- (P.$$$ "=" |-- (P.$$$ "compose" |-- P.list1 P.name))) >> mk_composition_decl || + (P.name -- (P.$$$ "=" |-- (P.$$$ "hide" |-- P.list1 P.term -- (P.$$$ "in" |-- P.name)))) + >> mk_hiding_decl || + (P.name -- (P.$$$ "=" |-- (P.$$$ "restrict" |-- P.name -- (P.$$$ "to" |-- P.list1 P.term)))) >> mk_restriction_decl || - (name -- ($$$ "=" |-- ($$$ "rename" |-- name -- ($$$ "with" |-- term)))) >> mk_rename_decl; + (P.name -- (P.$$$ "=" |-- (P.$$$ "rename" |-- P.name -- (P.$$$ "with" |-- P.term)))) + >> mk_rename_decl; -val automatonP = OuterSyntax.command "automaton" "define Lynch/Vaandrager-style I/O automaton" - (ioa_decl >> Toplevel.theory); +val automatonP = + OuterSyntax.command "automaton" "define Lynch/Vaandrager-style I/O automaton" K.thy_decl + (ioa_decl >> Toplevel.theory); end; diff -r 5e82c7196789 -r f342449d73ca src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Mon May 24 21:56:14 1999 +0200 +++ b/src/Pure/Isar/isar_syn.ML Mon May 24 21:57:13 1999 +0200 @@ -14,30 +14,30 @@ structure IsarSyn: ISAR_SYN = struct -open OuterParse; +structure P = OuterParse and K = OuterSyntax.Keyword; (** init and exit **) val theoryP = - OuterSyntax.command "theory" "begin theory" + OuterSyntax.command "theory" "begin theory" K.thy_begin (OuterSyntax.theory_header >> (Toplevel.print oo IsarThy.theory)); val end_excursionP = - OuterSyntax.command "end" "end current excursion" + OuterSyntax.command "end" "end current excursion" K.thy_end (Scan.succeed (Toplevel.print o Toplevel.exit)); val kill_excursionP = - OuterSyntax.command "kill" "kill current excursion" + OuterSyntax.command "kill" "kill current excursion" K.control (Scan.succeed (Toplevel.print o Toplevel.kill)); val contextP = - OuterSyntax.improper_command "context" "switch theory context" - (name >> (Toplevel.print oo IsarThy.context)); + OuterSyntax.improper_command "context" "switch theory context" K.thy_begin + (P.name >> (Toplevel.print oo IsarThy.context)); val update_contextP = - OuterSyntax.improper_command "update_context" "switch theory context, forcing update" - (name >> (Toplevel.print oo IsarThy.update_context)); + OuterSyntax.improper_command "update_context" "switch theory context, forcing update" K.thy_begin + (P.name >> (Toplevel.print oo IsarThy.update_context)); @@ -45,190 +45,190 @@ (* formal comments *) -val textP = OuterSyntax.command "text" "formal comments" - (comment >> (Toplevel.theory o IsarThy.add_text)); +val textP = OuterSyntax.command "text" "formal comments" K.thy_decl + (P.comment >> (Toplevel.theory o IsarThy.add_text)); -val titleP = OuterSyntax.command "title" "document title" - ((comment -- Scan.optional comment Comment.empty -- Scan.optional comment Comment.empty) +val titleP = OuterSyntax.command "title" "document title" K.thy_heading + ((P.comment -- Scan.optional P.comment Comment.empty -- Scan.optional P.comment Comment.empty) >> (fn ((x, y), z) => Toplevel.theory (IsarThy.add_title x y z))); -val chapterP = OuterSyntax.command "chapter" "chapter heading" - (comment >> (Toplevel.theory o IsarThy.add_chapter)); +val chapterP = OuterSyntax.command "chapter" "chapter heading" K.thy_heading + (P.comment >> (Toplevel.theory o IsarThy.add_chapter)); -val sectionP = OuterSyntax.command "section" "section heading" - (comment >> (Toplevel.theory o IsarThy.add_section)); +val sectionP = OuterSyntax.command "section" "section heading" K.thy_heading + (P.comment >> (Toplevel.theory o IsarThy.add_section)); -val subsectionP = OuterSyntax.command "subsection" "subsection heading" - (comment >> (Toplevel.theory o IsarThy.add_subsection)); +val subsectionP = OuterSyntax.command "subsection" "subsection heading" K.thy_heading + (P.comment >> (Toplevel.theory o IsarThy.add_subsection)); -val subsubsectionP = OuterSyntax.command "subsubsection" "subsubsection heading" - (comment >> (Toplevel.theory o IsarThy.add_subsubsection)); +val subsubsectionP = OuterSyntax.command "subsubsection" "subsubsection heading" K.thy_heading + (P.comment >> (Toplevel.theory o IsarThy.add_subsubsection)); (* classes and sorts *) val classesP = - OuterSyntax.command "classes" "declare type classes" - (Scan.repeat1 (name -- Scan.optional ($$$ "<" |-- !!! (list1 xname)) []) + OuterSyntax.command "classes" "declare type classes" K.thy_decl + (Scan.repeat1 (P.name -- Scan.optional (P.$$$ "<" |-- P.!!! (P.list1 P.xname)) []) >> (Toplevel.theory o Theory.add_classes)); val classrelP = - OuterSyntax.command "classrel" "state inclusion of type classes (axiomatic!)" - (xname -- ($$$ "<" |-- !!! xname) >> (Toplevel.theory o Theory.add_classrel o single)); + OuterSyntax.command "classrel" "state inclusion of type classes (axiomatic!)" K.thy_decl + (P.xname -- (P.$$$ "<" |-- P.!!! P.xname) >> (Toplevel.theory o Theory.add_classrel o single)); val defaultsortP = - OuterSyntax.command "defaultsort" "declare default sort" - (sort >> (Toplevel.theory o Theory.add_defsort)); + OuterSyntax.command "defaultsort" "declare default sort" K.thy_decl + (P.sort >> (Toplevel.theory o Theory.add_defsort)); (* types *) val typedeclP = - OuterSyntax.command "typedecl" "Pure type declaration" - (type_args -- name -- opt_infix + OuterSyntax.command "typedecl" "Pure type declaration" K.thy_decl + (P.type_args -- P.name -- P.opt_infix >> (fn ((args, a), mx) => Toplevel.theory (PureThy.add_typedecls [(a, args, mx)]))); val typeabbrP = - OuterSyntax.command "types" "declare type abbreviations" - (Scan.repeat1 (type_args -- name -- ($$$ "=" |-- !!! (typ -- opt_infix))) + OuterSyntax.command "types" "declare type abbreviations" K.thy_decl + (Scan.repeat1 (P.type_args -- P.name -- (P.$$$ "=" |-- P.!!! (P.typ -- P.opt_infix))) >> (Toplevel.theory o Theory.add_tyabbrs o map (fn ((args, a), (T, mx)) => (a, args, T, mx)))); val nontermP = OuterSyntax.command "nonterminals" "declare types treated as grammar nonterminal symbols" - (Scan.repeat1 name >> (Toplevel.theory o Theory.add_nonterminals)); + K.thy_decl (Scan.repeat1 P.name >> (Toplevel.theory o Theory.add_nonterminals)); val aritiesP = - OuterSyntax.command "arities" "state type arities (axiomatic!)" - (Scan.repeat1 (xname -- ($$$ "::" |-- !!! arity) >> triple2) + OuterSyntax.command "arities" "state type arities (axiomatic!)" K.thy_decl + (Scan.repeat1 (P.xname -- (P.$$$ "::" |-- P.!!! P.arity) >> P.triple2) >> (Toplevel.theory o Theory.add_arities)); (* consts and syntax *) val constsP = - OuterSyntax.command "consts" "declare constants" - (Scan.repeat1 const >> (Toplevel.theory o Theory.add_consts)); + OuterSyntax.command "consts" "declare constants" K.thy_decl + (Scan.repeat1 P.const >> (Toplevel.theory o Theory.add_consts)); val opt_mode = Scan.optional - ($$$ "(" |-- !!! (name -- Scan.optional ($$$ "output" >> K false) true --| $$$ ")")) + (P.$$$ "(" |-- P.!!! (P.name -- Scan.optional (P.$$$ "output" >> K false) true --| P.$$$ ")")) ("", true); val syntaxP = - OuterSyntax.command "syntax" "declare syntactic constants" - (opt_mode -- Scan.repeat1 const >> (Toplevel.theory o uncurry Theory.add_modesyntax)); + OuterSyntax.command "syntax" "declare syntactic constants" K.thy_decl + (opt_mode -- Scan.repeat1 P.const >> (Toplevel.theory o uncurry Theory.add_modesyntax)); (* translations *) val trans_pat = - Scan.optional ($$$ "(" |-- !!! (xname --| $$$ ")")) "logic" -- string; + Scan.optional (P.$$$ "(" |-- P.!!! (P.xname --| P.$$$ ")")) "logic" -- P.string; fun trans_arrow toks = - ($$$ "=>" >> K Syntax.ParseRule || - $$$ "<=" >> K Syntax.PrintRule || - $$$ "==" >> K Syntax.ParsePrintRule) toks; + (P.$$$ "=>" >> K Syntax.ParseRule || + P.$$$ "<=" >> K Syntax.PrintRule || + P.$$$ "==" >> K Syntax.ParsePrintRule) toks; val trans_line = - trans_pat -- !!! (trans_arrow -- trans_pat) + trans_pat -- P.!!! (trans_arrow -- trans_pat) >> (fn (left, (arr, right)) => arr (left, right)); val translationsP = - OuterSyntax.command "translations" "declare syntax translation rules" + OuterSyntax.command "translations" "declare syntax translation rules" K.thy_decl (Scan.repeat1 trans_line >> (Toplevel.theory o Theory.add_trrules)); (* axioms and definitions *) val axiomsP = - OuterSyntax.command "axioms" "state arbitrary propositions (axiomatic!)" - (Scan.repeat1 spec_name >> (Toplevel.theory o IsarThy.add_axioms)); + OuterSyntax.command "axioms" "state arbitrary propositions (axiomatic!)" K.thy_decl + (Scan.repeat1 P.spec_name >> (Toplevel.theory o IsarThy.add_axioms)); val defsP = - OuterSyntax.command "defs" "define constants" - (Scan.repeat1 spec_opt_name >> (Toplevel.theory o IsarThy.add_defs)); + OuterSyntax.command "defs" "define constants" K.thy_decl + (Scan.repeat1 P.spec_opt_name >> (Toplevel.theory o IsarThy.add_defs)); val constdefsP = - OuterSyntax.command "constdefs" "declare and define constants" - (Scan.repeat1 (const -- term) >> (Toplevel.theory o IsarThy.add_constdefs)); + OuterSyntax.command "constdefs" "declare and define constants" K.thy_decl + (Scan.repeat1 (P.const -- P.term) >> (Toplevel.theory o IsarThy.add_constdefs)); (* theorems *) -val facts = opt_thm_name "=" -- xthms1; +val facts = P.opt_thm_name "=" -- P.xthms1; val theoremsP = - OuterSyntax.command "theorems" "define theorems" + OuterSyntax.command "theorems" "define theorems" K.thy_decl (facts >> (Toplevel.theory o IsarThy.have_theorems)); val lemmasP = - OuterSyntax.command "lemmas" "define lemmas" + OuterSyntax.command "lemmas" "define lemmas" K.thy_decl (facts >> (Toplevel.theory o IsarThy.have_lemmas)); (* name space entry path *) val globalP = - OuterSyntax.command "global" "disable prefixing of theory name" + OuterSyntax.command "global" "disable prefixing of theory name" K.thy_decl (Scan.succeed (Toplevel.theory PureThy.global_path)); val localP = - OuterSyntax.command "local" "enable prefixing of theory name" + OuterSyntax.command "local" "enable prefixing of theory name" K.thy_decl (Scan.succeed (Toplevel.theory PureThy.local_path)); val pathP = - OuterSyntax.command "path" "modify name-space entry path" - (xname >> (Toplevel.theory o Theory.add_path)); + OuterSyntax.command "path" "modify name-space entry path" K.thy_decl + (P.xname >> (Toplevel.theory o Theory.add_path)); (* use ML text *) val useP = - OuterSyntax.command "use" "eval ML text from file" - (name >> IsarCmd.use); + OuterSyntax.command "use" "eval ML text from file" K.diag + (P.name >> IsarCmd.use); val mlP = - OuterSyntax.command "ML" "eval ML text" - (text >> (fn txt => IsarCmd.use_mltext txt o IsarCmd.use_mltext_theory txt)); + OuterSyntax.command "ML" "eval ML text" K.diag + (P.text >> (fn txt => IsarCmd.use_mltext txt o IsarCmd.use_mltext_theory txt)); val setupP = - OuterSyntax.command "setup" "apply ML theory transformer" - (text >> (Toplevel.theory o IsarThy.use_setup)); + OuterSyntax.command "setup" "apply ML theory transformer" K.thy_decl + (P.text >> (Toplevel.theory o IsarThy.use_setup)); (* translation functions *) val parse_ast_translationP = - OuterSyntax.command "parse_ast_translation" "install parse ast translation functions" - (text >> (Toplevel.theory o IsarThy.parse_ast_translation)); + OuterSyntax.command "parse_ast_translation" "install parse ast translation functions" K.thy_decl + (P.text >> (Toplevel.theory o IsarThy.parse_ast_translation)); val parse_translationP = - OuterSyntax.command "parse_translation" "install parse translation functions" - (text >> (Toplevel.theory o IsarThy.parse_translation)); + OuterSyntax.command "parse_translation" "install parse translation functions" K.thy_decl + (P.text >> (Toplevel.theory o IsarThy.parse_translation)); val print_translationP = - OuterSyntax.command "print_translation" "install print translation functions" - (text >> (Toplevel.theory o IsarThy.print_translation)); + OuterSyntax.command "print_translation" "install print translation functions" K.thy_decl + (P.text >> (Toplevel.theory o IsarThy.print_translation)); val typed_print_translationP = OuterSyntax.command "typed_print_translation" "install typed print translation functions" - (text >> (Toplevel.theory o IsarThy.typed_print_translation)); + K.thy_decl (P.text >> (Toplevel.theory o IsarThy.typed_print_translation)); val print_ast_translationP = - OuterSyntax.command "print_ast_translation" "install print ast translation functions" - (text >> (Toplevel.theory o IsarThy.print_ast_translation)); + OuterSyntax.command "print_ast_translation" "install print ast translation functions" K.thy_decl + (P.text >> (Toplevel.theory o IsarThy.print_ast_translation)); val token_translationP = - OuterSyntax.command "token_translation" "install token translation functions" - (text >> (Toplevel.theory o IsarThy.token_translation)); + OuterSyntax.command "token_translation" "install token translation functions" K.thy_decl + (P.text >> (Toplevel.theory o IsarThy.token_translation)); (* oracles *) val oracleP = - OuterSyntax.command "oracle" "install oracle" - (name -- text >> (Toplevel.theory o IsarThy.add_oracle)); + OuterSyntax.command "oracle" "install oracle" K.thy_decl + (P.name -- P.text >> (Toplevel.theory o IsarThy.add_oracle)); @@ -236,108 +236,108 @@ (* statements *) -val is_props = $$$ "(" |-- !!! (Scan.repeat1 ($$$ "is" |-- prop) --| $$$ ")"); -val propp = prop -- Scan.optional is_props []; -fun statement f = opt_thm_name ":" -- propp >> (fn ((x, y), z) => f x y z); +val is_props = P.$$$ "(" |-- P.!!! (Scan.repeat1 (P.$$$ "is" |-- P.prop) --| P.$$$ ")"); +val propp = P.prop -- Scan.optional is_props []; +fun statement f = P.opt_thm_name ":" -- propp >> (fn ((x, y), z) => f x y z); val theoremP = - OuterSyntax.command "theorem" "state theorem" + OuterSyntax.command "theorem" "state theorem" K.thy_goal (statement IsarThy.theorem >> (fn f => Toplevel.print o Toplevel.theory_to_proof f)); val lemmaP = - OuterSyntax.command "lemma" "state lemma" + OuterSyntax.command "lemma" "state lemma" K.thy_goal (statement IsarThy.lemma >> (fn f => Toplevel.print o Toplevel.theory_to_proof f)); val showP = - OuterSyntax.command "show" "state local goal, solving current obligation" + OuterSyntax.command "show" "state local goal, solving current obligation" K.prf_goal (statement IsarThy.show >> (fn f => Toplevel.print o Toplevel.proof f)); val haveP = - OuterSyntax.command "have" "state local goal" + OuterSyntax.command "have" "state local goal" K.prf_goal (statement IsarThy.have >> (fn f => Toplevel.print o Toplevel.proof f)); val thusP = - OuterSyntax.command "thus" "abbreviates \"then show\"" + OuterSyntax.command "thus" "abbreviates \"then show\"" K.prf_goal (statement IsarThy.thus >> (fn f => Toplevel.print o Toplevel.proof f)); val henceP = - OuterSyntax.command "hence" "abbreviates \"then have\"" + OuterSyntax.command "hence" "abbreviates \"then have\"" K.prf_goal (statement IsarThy.hence >> (fn f => Toplevel.print o Toplevel.proof f)); (* facts *) val thenP = - OuterSyntax.command "then" "forward chaining" + OuterSyntax.command "then" "forward chaining" K.prf_chain (Scan.succeed (Toplevel.print o Toplevel.proof IsarThy.chain)); val fromP = - OuterSyntax.command "from" "forward chaining from given facts" - (xthms1 >> (fn x => Toplevel.print o Toplevel.proof (IsarThy.from_facts x))); + OuterSyntax.command "from" "forward chaining from given facts" K.prf_chain + (P.xthms1 >> (fn x => Toplevel.print o Toplevel.proof (IsarThy.from_facts x))); val factsP = - OuterSyntax.command "note" "define facts" + OuterSyntax.command "note" "define facts" K.prf_decl (facts >> (Toplevel.proof o IsarThy.have_facts)); (* proof context *) val assumeP = - OuterSyntax.command "assume" "assume propositions" - (opt_thm_name ":" -- Scan.repeat1 propp >> + OuterSyntax.command "assume" "assume propositions" K.prf_decl + (P.opt_thm_name ":" -- Scan.repeat1 propp >> (fn ((x, y), z) => Toplevel.print o Toplevel.proof (IsarThy.assume x y z))); val fixP = - OuterSyntax.command "fix" "fix variables (Skolem constants)" - (Scan.repeat1 (name -- Scan.option ($$$ "::" |-- typ)) + OuterSyntax.command "fix" "fix variables (Skolem constants)" K.prf_decl + (Scan.repeat1 (P.name -- Scan.option (P.$$$ "::" |-- P.typ)) >> (fn xs => Toplevel.print o Toplevel.proof (IsarThy.fix xs))); val letP = - OuterSyntax.command "let" "bind text variables" - (and_list1 (enum1 "as" term -- ($$$ "=" |-- term)) + OuterSyntax.command "let" "bind text variables" K.prf_decl + (P.and_list1 (P.enum1 "as" P.term -- (P.$$$ "=" |-- P.term)) >> (fn bs => Toplevel.print o Toplevel.proof (IsarThy.match_bind bs))); (* proof structure *) val beginP = - OuterSyntax.command "{{" "begin explicit proof block" + OuterSyntax.command "{{" "begin explicit proof block" K.prf_block (Scan.succeed (Toplevel.print o Toplevel.proof IsarThy.begin_block)); val endP = - OuterSyntax.command "}}" "end explicit proof block" + OuterSyntax.command "}}" "end explicit proof block" K.prf_block (Scan.succeed (Toplevel.print o Toplevel.proof IsarThy.end_block)); val nextP = - OuterSyntax.command "next" "enter next proof block" + OuterSyntax.command "next" "enter next proof block" K.prf_block (Scan.succeed (Toplevel.print o Toplevel.proof IsarThy.next_block)); (* end proof *) val kill_proofP = - OuterSyntax.improper_command "kill_proof" "abort current proof" + OuterSyntax.improper_command "kill_proof" "abort current proof" K.control (Scan.succeed (Toplevel.print o Toplevel.proof_to_theory IsarThy.kill_proof)); val qed_withP = - OuterSyntax.improper_command "qed_with" "conclude proof, may patch name and attributes" - (Scan.option name -- Scan.option attribs -- Scan.option method + OuterSyntax.improper_command "qed_with" "conclude proof, may patch name and attributes" K.qed + (Scan.option P.name -- Scan.option P.attribs -- Scan.option P.method >> (uncurry IsarThy.global_qed_with)); val qedP = - OuterSyntax.command "qed" "conclude (sub-)proof" - (Scan.option method >> IsarThy.qed); + OuterSyntax.command "qed" "conclude (sub-)proof" K.qed + (Scan.option P.method >> IsarThy.qed); val terminal_proofP = - OuterSyntax.command "by" "terminal backward proof" - (method >> IsarThy.terminal_proof) + OuterSyntax.command "by" "terminal backward proof" K.qed + (P.method >> IsarThy.terminal_proof) val immediate_proofP = - OuterSyntax.command "." "immediate proof" + OuterSyntax.command "." "immediate proof" K.qed (Scan.succeed IsarThy.immediate_proof); val default_proofP = - OuterSyntax.command ".." "default proof" + OuterSyntax.command ".." "default proof" K.qed (Scan.succeed IsarThy.default_proof); @@ -345,49 +345,49 @@ val refineP = OuterSyntax.improper_command "refine" "unstructured backward proof step, ignoring facts" - (method >> (Toplevel.print oo (Toplevel.proof o IsarThy.tac))); + K.prf_script (P.method >> (Toplevel.print oo (Toplevel.proof o IsarThy.tac))); val then_refineP = OuterSyntax.improper_command "then_refine" "unstructured backward proof step, using facts" - (method >> (Toplevel.print oo (Toplevel.proof o IsarThy.then_tac))); + K.prf_script (P.method >> (Toplevel.print oo (Toplevel.proof o IsarThy.then_tac))); val proofP = - OuterSyntax.command "proof" "backward proof" - (Scan.option method >> (Toplevel.print oo (Toplevel.proof o IsarThy.proof))); + OuterSyntax.command "proof" "backward proof" K.prf_block + (Scan.option P.method >> (Toplevel.print oo (Toplevel.proof o IsarThy.proof))); (* proof history *) val clear_undoP = - OuterSyntax.improper_command "clear_undo" "clear undo information" + OuterSyntax.improper_command "clear_undo" "clear undo information" K.control (Scan.succeed (Toplevel.print o IsarCmd.clear_undo)); val undoP = - OuterSyntax.improper_command "undo" "undo last command" + OuterSyntax.improper_command "undo" "undo last command" K.control (Scan.succeed (Toplevel.print o IsarCmd.undo)); val redoP = - OuterSyntax.improper_command "redo" "redo last command" + OuterSyntax.improper_command "redo" "redo last command" K.control (Scan.succeed (Toplevel.print o IsarCmd.redo)); val undosP = - OuterSyntax.improper_command "undos" "undo last commands" - (nat >> (fn n => Toplevel.print o IsarCmd.undos n)); + OuterSyntax.improper_command "undos" "undo last commands" K.control + (P.nat >> (fn n => Toplevel.print o IsarCmd.undos n)); val backP = - OuterSyntax.improper_command "back" "backtracking of proof command" + OuterSyntax.improper_command "back" "backtracking of proof command" K.prf_script (Scan.succeed (Toplevel.print o Toplevel.proof ProofHistory.back)); val prevP = - OuterSyntax.improper_command "prev" "previous proof state" + OuterSyntax.improper_command "prev" "previous proof state" K.control (Scan.succeed (Toplevel.print o Toplevel.proof ProofHistory.prev)); val upP = - OuterSyntax.improper_command "up" "upper proof state" + OuterSyntax.improper_command "up" "upper proof state" K.control (Scan.succeed (Toplevel.print o Toplevel.proof ProofHistory.up)); val topP = - OuterSyntax.improper_command "top" "to initial proof state" + OuterSyntax.improper_command "top" "to initial proof state" K.control (Scan.succeed (Toplevel.print o Toplevel.proof ProofHistory.top)); @@ -395,101 +395,101 @@ (** diagnostic commands (for interactive mode only) **) val print_commandsP = - OuterSyntax.improper_command "help" "print outer syntax (global)" + OuterSyntax.improper_command "help" "print outer syntax (global)" K.diag (Scan.succeed (Toplevel.imperative OuterSyntax.print_outer_syntax)); val print_theoryP = - OuterSyntax.improper_command "print_theory" "print logical theory contents (verbose!)" + OuterSyntax.improper_command "print_theory" "print logical theory contents (verbose!)" K.diag (Scan.succeed IsarCmd.print_theory); val print_syntaxP = - OuterSyntax.improper_command "print_syntax" "print inner syntax of theory (verbose!)" + OuterSyntax.improper_command "print_syntax" "print inner syntax of theory (verbose!)" K.diag (Scan.succeed IsarCmd.print_syntax); val print_theoremsP = - OuterSyntax.improper_command "print_theorems" "print theorems known in this theory" + OuterSyntax.improper_command "print_theorems" "print theorems known in this theory" K.diag (Scan.succeed IsarCmd.print_theorems); val print_attributesP = - OuterSyntax.improper_command "print_attributes" "print attributes known in this theory" + OuterSyntax.improper_command "print_attributes" "print attributes known in this theory" K.diag (Scan.succeed IsarCmd.print_attributes); val print_methodsP = - OuterSyntax.improper_command "print_methods" "print methods known in this theory" + OuterSyntax.improper_command "print_methods" "print methods known in this theory" K.diag (Scan.succeed IsarCmd.print_methods); val print_bindsP = - OuterSyntax.improper_command "print_binds" "print term bindings of proof context" + OuterSyntax.improper_command "print_binds" "print term bindings of proof context" K.diag (Scan.succeed IsarCmd.print_binds); val print_lthmsP = - OuterSyntax.improper_command "print_facts" "print local theorems of proof context" + OuterSyntax.improper_command "print_facts" "print local theorems of proof context" K.diag (Scan.succeed IsarCmd.print_lthms); val print_thmsP = - OuterSyntax.improper_command "thm" "print theorems" (xthm >> IsarCmd.print_thms); + OuterSyntax.improper_command "thm" "print theorems" K.diag (P.xthm >> IsarCmd.print_thms); val print_propP = - OuterSyntax.improper_command "prop" "read and print proposition" - (term >> IsarCmd.print_prop); + OuterSyntax.improper_command "prop" "read and print proposition" K.diag + (P.term >> IsarCmd.print_prop); val print_termP = - OuterSyntax.improper_command "term" "read and print term" - (term >> IsarCmd.print_term); + OuterSyntax.improper_command "term" "read and print term" K.diag + (P.term >> IsarCmd.print_term); val print_typeP = - OuterSyntax.improper_command "typ" "read and print type" - (typ >> IsarCmd.print_type); + OuterSyntax.improper_command "typ" "read and print type" K.diag + (P.typ >> IsarCmd.print_type); (** system commands (for interactive mode only) **) val cdP = - OuterSyntax.improper_command "cd" "change current working directory" - (name >> IsarCmd.cd); + OuterSyntax.improper_command "cd" "change current working directory" K.control + (P.name >> IsarCmd.cd); val pwdP = - OuterSyntax.improper_command "pwd" "print current working directory" + OuterSyntax.improper_command "pwd" "print current working directory" K.diag (Scan.succeed IsarCmd.pwd); val use_thyP = - OuterSyntax.improper_command "use_thy" "use theory file" - (name >> IsarCmd.use_thy); + OuterSyntax.improper_command "use_thy" "use theory file" K.diag + (P.name >> IsarCmd.use_thy); val use_thy_onlyP = - OuterSyntax.improper_command "use_thy_only" "use theory file only, ignoring associated ML" - (name >> IsarCmd.use_thy_only); + OuterSyntax.improper_command "use_thy_only" "use theory file only, ignoring associated ML" K.diag + (P.name >> IsarCmd.use_thy_only); val update_thyP = - OuterSyntax.improper_command "update_thy" "update theory file" - (name >> IsarCmd.update_thy); + OuterSyntax.improper_command "update_thy" "update theory file" K.diag + (P.name >> IsarCmd.update_thy); val prP = - OuterSyntax.improper_command "pr" "print current toplevel state" + OuterSyntax.improper_command "pr" "print current toplevel state" K.diag (Scan.succeed (Toplevel.print o Toplevel.imperative (K ()))); -val opt_unit = Scan.optional ($$$ "(" -- $$$ ")" >> (K ())) (); +val opt_unit = Scan.optional (P.$$$ "(" -- P.$$$ ")" >> (K ())) (); val commitP = - OuterSyntax.improper_command "commit" "commit current session to ML database" + OuterSyntax.improper_command "commit" "commit current session to ML database" K.diag (opt_unit >> (K IsarCmd.use_commit)); val quitP = - OuterSyntax.improper_command "quit" "quit Isabelle" + OuterSyntax.improper_command "quit" "quit Isabelle" K.control (opt_unit >> (K IsarCmd.quit)); val exitP = - OuterSyntax.improper_command "exit" "exit Isar loop" + OuterSyntax.improper_command "exit" "exit Isar loop" K.control (Scan.succeed IsarCmd.exit); val restartP = - OuterSyntax.improper_command "restart" "restart Isar loop" + OuterSyntax.improper_command "restart" "restart Isar loop" K.control (Scan.succeed IsarCmd.restart); val breakP = - OuterSyntax.improper_command "break" "discontinue excursion (keep current state)" + OuterSyntax.improper_command "break" "discontinue excursion (keep current state)" K.control (Scan.succeed IsarCmd.break);