# HG changeset patch # User wenzelm # Date 1459800827 -7200 # Node ID a8758f47f9e8adfc64e207553da4099e30899f92 # Parent 313d3b697c9a4724fc2e8d01ef6d021f6ed02b4a# Parent 3f97aa4580c67760c41ec02fe0f069f37738f54b merged diff -r 313d3b697c9a -r a8758f47f9e8 Admin/lib/Tools/makedist --- a/Admin/lib/Tools/makedist Mon Apr 04 16:52:56 2016 +0100 +++ b/Admin/lib/Tools/makedist Mon Apr 04 22:13:47 2016 +0200 @@ -159,12 +159,13 @@ perl -pi \ -e "s,val is_identified = false,val is_identified = true,g;" \ -e "s,val is_official = false,val is_official = ${IS_OFFICIAL},g;" \ - src/Pure/ROOT.ML src/Pure/ROOT.scala + src/Pure/System/distribution.ML src/Pure/System/distribution.scala perl -pi -e "s,ISABELLE_ID=\"\",ISABELLE_ID=\"$IDENT\",g" lib/scripts/getsettings perl -pi -e "s,ISABELLE_IDENTIFIER=\"\",ISABELLE_IDENTIFIER=\"$DISTNAME\",g" lib/scripts/getsettings perl -pi -e "s,{ISABELLE},$DISTNAME,g" lib/html/library_index_header.template -perl -pi -e "s,unidentified repository version,$DISTVERSION,g" src/Pure/ROOT.ML src/Pure/ROOT.scala lib/Tools/version +perl -pi -e "s,unidentified repository version,$DISTVERSION,g" \ + src/Pure/System/distribution.ML src/Pure/System/distribution.scala lib/Tools/version perl -pi -e "s,some unidentified repository version of Isabelle,$DISTVERSION,g" README mkdir -p contrib diff -r 313d3b697c9a -r a8758f47f9e8 NEWS --- a/NEWS Mon Apr 04 16:52:56 2016 +0100 +++ b/NEWS Mon Apr 04 22:13:47 2016 +0200 @@ -230,6 +230,10 @@ *** ML *** +* Low-level ML system structures (like PolyML and RunCall) are no longer +exposed to Isabelle/ML user-space. The system option ML_system_unsafe +allows to override this for special test situations. + * Antiquotation @{make_string} is available during Pure bootstrap -- with approximative output quality. diff -r 313d3b697c9a -r a8758f47f9e8 etc/options --- a/etc/options Mon Apr 04 16:52:56 2016 +0100 +++ b/etc/options Mon Apr 04 22:13:47 2016 +0200 @@ -126,6 +126,9 @@ public option ML_system_64 : bool = false -- "ML system for 64bit platform is used if possible (change requires restart)" +option ML_system_unsafe : bool = false + -- "provide access to low-level ML system structures" + section "Editor Reactivity" diff -r 313d3b697c9a -r a8758f47f9e8 src/Doc/Implementation/Integration.thy --- a/src/Doc/Implementation/Integration.thy Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Doc/Implementation/Integration.thy Mon Apr 04 22:13:47 2016 +0200 @@ -152,7 +152,6 @@ text %mlref \ \begin{mldecls} @{index_ML use_thy: "string -> unit"} \\ - @{index_ML use_thys: "string list -> unit"} \\[0.5ex] @{index_ML Thy_Info.get_theory: "string -> theory"} \\ @{index_ML Thy_Info.remove_thy: "string -> unit"} \\ @{index_ML Thy_Info.register_thy: "theory -> unit"} \\ @@ -161,14 +160,6 @@ \<^descr> @{ML use_thy}~\A\ ensures that theory \A\ is fully up-to-date wrt.\ the external file store; outdated ancestors are reloaded on demand. - \<^descr> @{ML use_thys} is similar to @{ML use_thy}, but handles several theories - simultaneously. Thus it acts like processing the import header of a theory, - without performing the merge of the result. By loading a whole sub-graph of - theories, the intrinsic parallelism can be exploited by the system to - speedup loading. - - This variant is used by default in @{tool build} @{cite "isabelle-system"}. - \<^descr> @{ML Thy_Info.get_theory}~\A\ retrieves the theory value presently associated with name \A\. Note that the result might be outdated wrt.\ the file-system content. diff -r 313d3b697c9a -r a8758f47f9e8 src/Doc/System/Environment.thy --- a/src/Doc/System/Environment.thy Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Doc/System/Environment.thy Mon Apr 04 22:13:47 2016 +0200 @@ -395,8 +395,8 @@ The user is connected to the raw ML toplevel loop: this is neither Isabelle/Isar nor Isabelle/ML within the usual formal context. The most - relevant ML commands at this stage are @{ML use}, @{ML use_thy}, @{ML - use_thys}. + relevant ML commands at this stage are @{ML use} (for ML files) and + @{ML use_thy} (for theory files). \ diff -r 313d3b697c9a -r a8758f47f9e8 src/HOL/ex/Cartouche_Examples.thy --- a/src/HOL/ex/Cartouche_Examples.thy Mon Apr 04 16:52:56 2016 +0100 +++ b/src/HOL/ex/Cartouche_Examples.thy Mon Apr 04 22:13:47 2016 +0200 @@ -274,8 +274,6 @@ text \Nested ML evaluation:\ ML \ - val ML = ML_Context.eval_source ML_Compiler.flags; - ML \ML \val a = @{thm refl}\\; ML \val b = @{thm sym}\; val c = @{thm trans} diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/General/file.ML --- a/src/Pure/General/file.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/General/file.ML Mon Apr 04 22:13:47 2016 +0200 @@ -46,21 +46,22 @@ val standard_path = Path.implode o Path.expand; val platform_path = ML_System.platform_path o standard_path; -val bash_string = - translate_string (fn ch => - let val c = ord ch in - (case ch of - "\t" => "$'\\t'" - | "\n" => "$'\\n'" - | "\f" => "$'\\f'" - | "\r" => "$'\\r'" - | _ => - if Symbol.is_ascii_letter ch orelse Symbol.is_ascii_digit ch orelse - exists_string (fn c => c = ch) "-./:_" then ch - else if c < 16 then "$'\\x0" ^ Int.fmt StringCvt.HEX c ^ "'" - else if c < 32 orelse c >= 127 then "$'\\x" ^ Int.fmt StringCvt.HEX c ^ "'" - else "\\" ^ ch) - end); +fun bash_string "" = "\"\"" + | bash_string str = + str |> translate_string (fn ch => + let val c = ord ch in + (case ch of + "\t" => "$'\\t'" + | "\n" => "$'\\n'" + | "\f" => "$'\\f'" + | "\r" => "$'\\r'" + | _ => + if Symbol.is_ascii_letter ch orelse Symbol.is_ascii_digit ch orelse + exists_string (fn c => c = ch) "-./:_" then ch + else if c < 16 then "$'\\x0" ^ Int.fmt StringCvt.HEX c ^ "'" + else if c < 32 orelse c >= 127 then "$'\\x" ^ Int.fmt StringCvt.HEX c ^ "'" + else "\\" ^ ch) + end); val bash_args = space_implode " " o map bash_string; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/General/file.scala --- a/src/Pure/General/file.scala Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/General/file.scala Mon Apr 04 22:13:47 2016 +0200 @@ -123,7 +123,8 @@ } def bash_string(s: String): String = - UTF8.bytes(s).iterator.map(bash_chr(_)).mkString + if (s == "") "\"\"" + else UTF8.bytes(s).iterator.map(bash_chr(_)).mkString def bash_args(args: List[String]): String = args.iterator.map(bash_string(_)).mkString(" ") diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Isar/calculation.ML --- a/src/Pure/Isar/calculation.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Isar/calculation.ML Mon Apr 04 22:13:47 2016 +0200 @@ -205,32 +205,4 @@ val moreover = collect false; val ultimately = collect true; - -(* outer syntax *) - -val calc_args = - Scan.option (@{keyword "("} |-- Parse.!!! ((Parse.xthms1 --| @{keyword ")"}))); - -val _ = - Outer_Syntax.command @{command_keyword also} "combine calculation and current facts" - (calc_args >> (Toplevel.proofs' o also_cmd)); - -val _ = - Outer_Syntax.command @{command_keyword finally} - "combine calculation and current facts, exhibit result" - (calc_args >> (Toplevel.proofs' o finally_cmd)); - -val _ = - Outer_Syntax.command @{command_keyword moreover} "augment calculation by current facts" - (Scan.succeed (Toplevel.proof' moreover)); - -val _ = - Outer_Syntax.command @{command_keyword ultimately} - "augment calculation by current facts, exhibit result" - (Scan.succeed (Toplevel.proof' ultimately)); - -val _ = - Outer_Syntax.command @{command_keyword print_trans_rules} "print transitivity rules" - (Scan.succeed (Toplevel.keep (print_rules o Toplevel.context_of))); - end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Isar/isar_syn.ML Mon Apr 04 22:13:47 2016 +0200 @@ -148,6 +148,12 @@ >> (fn (facts, fixes) => #2 oo Specification.theorems_cmd "" [(Attrib.empty_binding, flat facts)] fixes)); +val _ = + Outer_Syntax.local_theory @{command_keyword named_theorems} + "declare named collection of theorems" + (Parse.and_list1 (Parse.binding -- Scan.optional Parse.text "") >> + fold (fn (b, descr) => snd o Named_Theorems.declare b descr)); + (* hide names *) @@ -183,6 +189,23 @@ (* use ML text *) +local + +fun ML_file_cmd debug files = Toplevel.generic_theory (fn gthy => + let + val [{src_path, lines, digest, pos}: Token.file] = files (Context.theory_of gthy); + val provide = Resources.provide (src_path, digest); + val source = Input.source true (cat_lines lines) (pos, pos); + val flags: ML_Compiler.flags = + {SML = false, exchange = false, redirect = true, verbose = true, + debug = debug, writeln = writeln, warning = warning}; + in + gthy + |> ML_Context.exec (fn () => ML_Context.eval_source flags source) + |> Local_Theory.propagate_ml_env + |> Context.mapping provide (Local_Theory.background_theory provide) + end); + fun SML_file_cmd debug files = Toplevel.theory (fn thy => let val ([{lines, pos, ...}: Token.file], thy') = files thy; @@ -195,6 +218,22 @@ (ML_Context.exec (fn () => ML_Context.eval_source flags source)) end); +in + +val _ = + Outer_Syntax.command ("ML_file", @{here}) "read and evaluate Isabelle/ML file" + (Resources.parse_files "ML_file" >> ML_file_cmd NONE); + +val _ = + Outer_Syntax.command ("ML_file_debug", @{here}) + "read and evaluate Isabelle/ML file (with debugger information)" + (Resources.parse_files "ML_file_debug" >> ML_file_cmd (SOME true)); + +val _ = + Outer_Syntax.command ("ML_file_no_debug", @{here}) + "read and evaluate Isabelle/ML file (no debugger information)" + (Resources.parse_files "ML_file_no_debug" >> ML_file_cmd (SOME false)); + val _ = Outer_Syntax.command @{command_keyword SML_file} "read and evaluate Standard ML file" (Resources.provide_parse_files "SML_file" >> SML_file_cmd NONE); @@ -209,6 +248,8 @@ "read and evaluate Standard ML file (no debugger information)" (Resources.provide_parse_files "SML_file_no_debug" >> SML_file_cmd (SOME false)); +end; + val _ = Outer_Syntax.command @{command_keyword SML_export} "evaluate SML within Isabelle/ML environment" (Parse.ML_source >> (fn source => @@ -235,14 +276,6 @@ end)); val _ = - Outer_Syntax.command @{command_keyword ML} "ML text within theory or local theory" - (Parse.ML_source >> (fn source => - Toplevel.generic_theory - (ML_Context.exec (fn () => - ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source) #> - Local_Theory.propagate_ml_env))); - -val _ = Outer_Syntax.command @{command_keyword ML_prf} "ML text within proof" (Parse.ML_source >> (fn source => Toplevel.proof (Proof.map_context (Context.proof_map @@ -728,6 +761,34 @@ end; +(* calculation *) + +val calculation_args = + Scan.option (@{keyword "("} |-- Parse.!!! ((Parse.xthms1 --| @{keyword ")"}))); + +val _ = + Outer_Syntax.command @{command_keyword also} "combine calculation and current facts" + (calculation_args >> (Toplevel.proofs' o Calculation.also_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword finally} + "combine calculation and current facts, exhibit result" + (calculation_args >> (Toplevel.proofs' o Calculation.finally_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword moreover} "augment calculation by current facts" + (Scan.succeed (Toplevel.proof' Calculation.moreover)); + +val _ = + Outer_Syntax.command @{command_keyword ultimately} + "augment calculation by current facts, exhibit result" + (Scan.succeed (Toplevel.proof' Calculation.ultimately)); + +val _ = + Outer_Syntax.command @{command_keyword print_trans_rules} "print transitivity rules" + (Scan.succeed (Toplevel.keep (Calculation.print_rules o Toplevel.context_of))); + + (* proof navigation *) fun report_back () = @@ -936,6 +997,96 @@ Toplevel.keep (fn _ => ignore (Present.display_drafts (map Path.explode names))))); +(* deps *) + +local + val theory_bounds = + Parse.position Parse.theory_xname >> single || + (@{keyword "("} |-- Parse.enum "|" (Parse.position Parse.theory_xname) --| @{keyword ")"}); +in + +val _ = + Outer_Syntax.command @{command_keyword thy_deps} "visualize theory dependencies" + (Scan.option theory_bounds -- Scan.option theory_bounds >> + (fn args => Toplevel.keep (fn st => Thy_Deps.thy_deps_cmd (Toplevel.context_of st) args))); + +end; + +local + val class_bounds = + Parse.sort >> single || + (@{keyword "("} |-- Parse.enum "|" Parse.sort --| @{keyword ")"}); +in + +val _ = + Outer_Syntax.command @{command_keyword class_deps} "visualize class dependencies" + (Scan.option class_bounds -- Scan.option class_bounds >> (fn args => + Toplevel.keep (fn st => Class_Deps.class_deps_cmd (Toplevel.context_of st) args))); + +end; + +val _ = + Outer_Syntax.command @{command_keyword thm_deps} "visualize theorem dependencies" + (Parse.xthms1 >> (fn args => + Toplevel.keep (fn st => + Thm_Deps.thm_deps (Toplevel.theory_of st) + (Attrib.eval_thms (Toplevel.context_of st) args)))); + +local + val thy_names = + Scan.repeat1 (Scan.unless Parse.minus (Parse.position Parse.theory_xname)); +in + +val _ = + Outer_Syntax.command @{command_keyword unused_thms} "find unused theorems" + (Scan.option ((thy_names --| Parse.minus) -- Scan.option thy_names) >> (fn opt_range => + Toplevel.keep (fn st => + let + val thy = Toplevel.theory_of st; + val ctxt = Toplevel.context_of st; + fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]); + val check = Theory.check ctxt; + in + Thm_Deps.unused_thms + (case opt_range of + NONE => (Theory.parents_of thy, [thy]) + | SOME (xs, NONE) => (map check xs, [thy]) + | SOME (xs, SOME ys) => (map check xs, map check ys)) + |> map pretty_thm |> Pretty.writeln_chunks + end))); + +end; + + +(* find *) + +val _ = + Outer_Syntax.command @{command_keyword find_consts} + "find constants by name / type patterns" + (Find_Consts.query_parser >> (fn spec => + Toplevel.keep (fn st => + Pretty.writeln (Find_Consts.pretty_consts (Toplevel.context_of st) spec)))); + +local + val options = + Scan.optional + (Parse.$$$ "(" |-- + Parse.!!! (Scan.option Parse.nat -- + Scan.optional (Parse.reserved "with_dups" >> K false) true --| Parse.$$$ ")")) + (NONE, true); +in + +val _ = + Outer_Syntax.command @{command_keyword find_theorems} + "find theorems meeting specified criteria" + (options -- Find_Theorems.query_parser >> (fn ((opt_lim, rem_dups), spec) => + Toplevel.keep (fn st => + Pretty.writeln + (Find_Theorems.pretty_theorems (Find_Theorems.proof_state st) opt_lim rem_dups spec)))); + +end; + + (** extraction of programs from proofs **) diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Isar/local_theory.ML Mon Apr 04 22:13:47 2016 +0200 @@ -309,17 +309,21 @@ fun type_notation add mode raw_args lthy = let val args = map (apfst (Logic.type_map (Assumption.export_term lthy (target_of lthy)))) raw_args; + val args' = map (apsnd Mixfix.reset_pos) args; + val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args; in declaration {syntax = true, pervasive = false} - (Proof_Context.generic_type_notation add mode args) lthy + (Proof_Context.generic_type_notation add mode args') lthy end; fun notation add mode raw_args lthy = let val args = map (apfst (Assumption.export_term lthy (target_of lthy))) raw_args + val args' = map (apsnd Mixfix.reset_pos) args; + val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args; in declaration {syntax = true, pervasive = false} - (Proof_Context.generic_notation add mode args) lthy + (Proof_Context.generic_notation add mode args') lthy end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Isar/outer_syntax.ML --- a/src/Pure/Isar/outer_syntax.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Isar/outer_syntax.ML Mon Apr 04 22:13:47 2016 +0200 @@ -306,4 +306,15 @@ in error ("Bad command " ^ quote name ^ Position.here pos ^ report) end end; + +(* 'ML' command -- required for bootstrapping Isar *) + +val _ = + command ("ML", @{here}) "ML text within theory or local theory" + (Parse.ML_source >> (fn source => + Toplevel.generic_theory + (ML_Context.exec (fn () => + ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source) #> + Local_Theory.propagate_ml_env))); + end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Isar/parse.ML --- a/src/Pure/Isar/parse.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Isar/parse.ML Mon Apr 04 22:13:47 2016 +0200 @@ -85,6 +85,7 @@ val mixfix': mixfix parser val opt_mixfix: mixfix parser val opt_mixfix': mixfix parser + val syntax_mode: Syntax.mode parser val where_: string parser val const_decl: (string * string * mixfix) parser val const_binding: (binding * string * mixfix) parser @@ -356,6 +357,15 @@ end; +(* syntax mode *) + +val syntax_mode_spec = + ($$$ "output" >> K ("", false)) || name -- Scan.optional ($$$ "output" >> K false) true; + +val syntax_mode = + Scan.optional ($$$ "(" |-- !!! (syntax_mode_spec --| $$$ ")")) Syntax.mode_default; + + (* fixes *) val where_ = $$$ "where"; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/fixed_int_dummy.ML --- a/src/Pure/ML/fixed_int_dummy.ML Mon Apr 04 16:52:56 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,6 +0,0 @@ -(* Title: Pure/ML/fixed_int_dummy.ML - -FixedInt dummy that is not fixed (up to Poly/ML 5.6). -*) - -structure FixedInt = IntInf; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_antiquotation.ML --- a/src/Pure/ML/ml_antiquotation.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/ML/ml_antiquotation.ML Mon Apr 04 22:13:47 2016 +0200 @@ -42,6 +42,11 @@ (Args.context >> (ML_Pretty.make_string_fn o ML_Context.struct_name)) #> value (Binding.make ("binding", @{here})) - (Scan.lift (Parse.position Args.name) >> ML_Syntax.make_binding)); + (Scan.lift (Parse.position Args.name) >> ML_Syntax.make_binding) #> + + value (Binding.make ("cartouche", @{here})) + (Scan.lift Args.cartouche_input >> (fn source => + "Input.source true " ^ ML_Syntax.print_string (Input.text_of source) ^ " " ^ + ML_Syntax.atomic (ML_Syntax.print_range (Input.range_of source))))); end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_antiquotations.ML --- a/src/Pure/ML/ml_antiquotations.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/ML/ml_antiquotations.ML Mon Apr 04 22:13:47 2016 +0200 @@ -10,12 +10,7 @@ (* ML support *) val _ = Theory.setup - (ML_Antiquotation.value @{binding cartouche} - (Scan.lift Args.cartouche_input >> (fn source => - "Input.source true " ^ ML_Syntax.print_string (Input.text_of source) ^ " " ^ - ML_Syntax.atomic (ML_Syntax.print_range (Input.range_of source)))) #> - - ML_Antiquotation.inline @{binding undefined} + (ML_Antiquotation.inline @{binding undefined} (Scan.succeed "(raise General.Match)") #> ML_Antiquotation.inline @{binding assert} diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_context.ML --- a/src/Pure/ML/ml_context.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/ML/ml_context.ML Mon Apr 04 22:13:47 2016 +0200 @@ -237,3 +237,4 @@ end; +val ML = ML_Context.eval_source ML_Compiler.flags; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_file.ML --- a/src/Pure/ML/ml_file.ML Mon Apr 04 16:52:56 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,39 +0,0 @@ -(* Title: Pure/ML/ml_file.ML - Author: Makarius - -The 'ML_file' command. -*) - -structure ML_File: sig end = -struct - -fun ML_file_cmd debug files = Toplevel.generic_theory (fn gthy => - let - val [{src_path, lines, digest, pos}: Token.file] = files (Context.theory_of gthy); - val provide = Resources.provide (src_path, digest); - val source = Input.source true (cat_lines lines) (pos, pos); - val flags: ML_Compiler.flags = - {SML = false, exchange = false, redirect = true, verbose = true, - debug = debug, writeln = writeln, warning = warning}; - in - gthy - |> ML_Context.exec (fn () => ML_Context.eval_source flags source) - |> Local_Theory.propagate_ml_env - |> Context.mapping provide (Local_Theory.background_theory provide) - end); - -val _ = - Outer_Syntax.command ("ML_file", @{here}) "read and evaluate Isabelle/ML file" - (Resources.parse_files "ML_file" >> ML_file_cmd NONE); - -val _ = - Outer_Syntax.command ("ML_file_debug", @{here}) - "read and evaluate Isabelle/ML file (with debugger information)" - (Resources.parse_files "ML_file_debug" >> ML_file_cmd (SOME true)); - -val _ = - Outer_Syntax.command ("ML_file_no_debug", @{here}) - "read and evaluate Isabelle/ML file (no debugger information)" - (Resources.parse_files "ML_file_no_debug" >> ML_file_cmd (SOME false)); - -end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_name_space.ML --- a/src/Pure/ML/ml_name_space.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/ML/ml_name_space.ML Mon Apr 04 22:13:47 2016 +0200 @@ -1,13 +1,11 @@ (* Title: Pure/ML/ml_name_space.ML Author: Makarius -Name space for Poly/ML. +ML name space, with initial entries for strict Standard ML. *) structure ML_Name_Space = struct - val critical_structures = ["CInterface", "Foreign", "RunCall", "RuntimeCalls", "Signal"]; - open PolyML.NameSpace; type T = PolyML.NameSpace.nameSpace; @@ -66,9 +64,10 @@ val sml_type = #allType global (); val sml_fixity = #allFix global (); val sml_structure = - List.filter (fn (a, _) => a <> "PolyML" andalso a <> "ML_System" andalso - not (List.exists (fn b => a = b) critical_structures)) (#allStruct global ()); - val sml_signature = - List.filter (fn (a, _) => a <> "ML_SYSTEM" andalso a <> "ML_NAME_SPACE") (#allSig global ()); + List.filter (fn (a, _) => + List.all (fn b => a <> b) + ["PolyML", "CInterface", "Foreign", "RunCall", "RuntimeCalls", "Signal"]) + (#allStruct global ()); + val sml_signature = #allSig global (); val sml_functor = #allFunct global (); end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_pervasive.ML --- a/src/Pure/ML/ml_pervasive.ML Mon Apr 04 16:52:56 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,28 +0,0 @@ -(* Title: Pure/ML/ml_pervasive.ML - Author: Makarius - -Pervasive ML environment. -*) - -structure PolyML_Pretty = -struct - datatype context = datatype PolyML.context; - datatype pretty = datatype PolyML.pretty; -end; - -val seconds = Time.fromReal; - -val _ = - List.app ML_Name_Space.forget_val - ["isSome", "getOpt", "valOf", "foldl", "foldr", "print", "explode", "concat"]; - -val ord = SML90.ord; -val chr = SML90.chr; -val raw_explode = SML90.explode; -val implode = SML90.implode; - -val pointer_eq = PolyML.pointerEq; - -val error_depth = PolyML.error_depth; - -open Thread; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_pervasive_final.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/ML/ml_pervasive_final.ML Mon Apr 04 22:13:47 2016 +0200 @@ -0,0 +1,19 @@ +(* Title: Pure/ML/ml_pervasive_final.ML + Author: Makarius + +Pervasive ML environment: final setup. +*) + +if Options.default_bool "ML_system_unsafe" then () +else + (List.app ML_Name_Space.forget_structure + ["CInterface", "Foreign", "RunCall", "RuntimeCalls", "Signal"]; + ML \structure PolyML = struct structure IntInf = PolyML.IntInf end\); + +structure Output: OUTPUT = Output; (*seal system channels!*) + +structure Pure = struct val thy = Thy_Info.pure_theory () end; + +Proofterm.proofs := 0; + +Context.set_thread_data NONE; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ML/ml_pervasive_initial.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/ML/ml_pervasive_initial.ML Mon Apr 04 22:13:47 2016 +0200 @@ -0,0 +1,28 @@ +(* Title: Pure/ML/ml_pervasive_initial.ML + Author: Makarius + +Pervasive ML environment: initial setup. +*) + +structure PolyML_Pretty = +struct + datatype context = datatype PolyML.context; + datatype pretty = datatype PolyML.pretty; +end; + +val seconds = Time.fromReal; + +val _ = + List.app ML_Name_Space.forget_val + ["isSome", "getOpt", "valOf", "foldl", "foldr", "print", "explode", "concat"]; + +val ord = SML90.ord; +val chr = SML90.chr; +val raw_explode = SML90.explode; +val implode = SML90.implode; + +val pointer_eq = PolyML.pointerEq; + +val error_depth = PolyML.error_depth; + +open Thread; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Pure.thy --- a/src/Pure/Pure.thy Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Pure.thy Mon Apr 04 22:13:47 2016 +0200 @@ -21,9 +21,9 @@ "definition" "abbreviation" "type_notation" "no_type_notation" "notation" "no_notation" "axiomatization" "lemmas" "declare" "hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl + and "ML_file" "ML_file_debug" "ML_file_no_debug" :: thy_load % "ML" and "SML_file" "SML_file_debug" "SML_file_no_debug" :: thy_load % "ML" and "SML_import" "SML_export" :: thy_decl % "ML" - and "ML" :: thy_decl % "ML" and "ML_prf" :: prf_decl % "proof" (* FIXME % "ML" ?? *) and "ML_val" "ML_command" :: diag % "ML" and "simproc_setup" :: thy_decl % "ML" == "" @@ -93,26 +93,1222 @@ and "named_theorems" :: thy_decl begin -ML_file "ML/ml_antiquotations.ML" -ML_file "ML/ml_thms.ML" -ML_file "Tools/print_operation.ML" -ML_file "Isar/isar_syn.ML" -ML_file "Isar/calculation.ML" -ML_file "Tools/bibtex.ML" -ML_file "Tools/rail.ML" -ML_file "Tools/rule_insts.ML" -ML_file "Tools/thm_deps.ML" -ML_file "Tools/thy_deps.ML" -ML_file "Tools/class_deps.ML" -ML_file "Tools/find_theorems.ML" -ML_file "Tools/find_consts.ML" -ML_file "Tools/simplifier_trace.ML" -ML_file_no_debug "Tools/debugger.ML" -ML_file "Tools/named_theorems.ML" -ML_file "Tools/jedit.ML" +section \Isar commands\ + +subsection \Embedded ML text\ + +ML \ +local + +fun ML_file_cmd debug files = Toplevel.generic_theory (fn gthy => + let + val [{src_path, lines, digest, pos}: Token.file] = files (Context.theory_of gthy); + val provide = Resources.provide (src_path, digest); + val source = Input.source true (cat_lines lines) (pos, pos); + val flags: ML_Compiler.flags = + {SML = false, exchange = false, redirect = true, verbose = true, + debug = debug, writeln = writeln, warning = warning}; + in + gthy + |> ML_Context.exec (fn () => ML_Context.eval_source flags source) + |> Local_Theory.propagate_ml_env + |> Context.mapping provide (Local_Theory.background_theory provide) + end); + +fun SML_file_cmd debug files = Toplevel.theory (fn thy => + let + val ([{lines, pos, ...}: Token.file], thy') = files thy; + val source = Input.source true (cat_lines lines) (pos, pos); + val flags: ML_Compiler.flags = + {SML = true, exchange = false, redirect = true, verbose = true, + debug = debug, writeln = writeln, warning = warning}; + in + thy' |> Context.theory_map + (ML_Context.exec (fn () => ML_Context.eval_source flags source)) + end); + +val _ = + Outer_Syntax.command ("ML_file", @{here}) "read and evaluate Isabelle/ML file" + (Resources.parse_files "ML_file" >> ML_file_cmd NONE); + +val _ = + Outer_Syntax.command ("ML_file_debug", @{here}) + "read and evaluate Isabelle/ML file (with debugger information)" + (Resources.parse_files "ML_file_debug" >> ML_file_cmd (SOME true)); + +val _ = + Outer_Syntax.command ("ML_file_no_debug", @{here}) + "read and evaluate Isabelle/ML file (no debugger information)" + (Resources.parse_files "ML_file_no_debug" >> ML_file_cmd (SOME false)); + +val _ = + Outer_Syntax.command @{command_keyword SML_file} "read and evaluate Standard ML file" + (Resources.provide_parse_files "SML_file" >> SML_file_cmd NONE); + +val _ = + Outer_Syntax.command @{command_keyword SML_file_debug} + "read and evaluate Standard ML file (with debugger information)" + (Resources.provide_parse_files "SML_file_debug" >> SML_file_cmd (SOME true)); + +val _ = + Outer_Syntax.command @{command_keyword SML_file_no_debug} + "read and evaluate Standard ML file (no debugger information)" + (Resources.provide_parse_files "SML_file_no_debug" >> SML_file_cmd (SOME false)); + +val _ = + Outer_Syntax.command @{command_keyword SML_export} "evaluate SML within Isabelle/ML environment" + (Parse.ML_source >> (fn source => + let + val flags: ML_Compiler.flags = + {SML = true, exchange = true, redirect = false, verbose = true, + debug = NONE, writeln = writeln, warning = warning}; + in + Toplevel.theory + (Context.theory_map (ML_Context.exec (fn () => ML_Context.eval_source flags source))) + end)); + +val _ = + Outer_Syntax.command @{command_keyword SML_import} "evaluate Isabelle/ML within SML environment" + (Parse.ML_source >> (fn source => + let + val flags: ML_Compiler.flags = + {SML = false, exchange = true, redirect = false, verbose = true, + debug = NONE, writeln = writeln, warning = warning}; + in + Toplevel.generic_theory + (ML_Context.exec (fn () => ML_Context.eval_source flags source) #> + Local_Theory.propagate_ml_env) + end)); + +val _ = + Outer_Syntax.command @{command_keyword ML_prf} "ML text within proof" + (Parse.ML_source >> (fn source => + Toplevel.proof (Proof.map_context (Context.proof_map + (ML_Context.exec (fn () => + ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source))) #> + Proof.propagate_ml_env))); + +val _ = + Outer_Syntax.command @{command_keyword ML_val} "diagnostic ML text" + (Parse.ML_source >> Isar_Cmd.ml_diag true); + +val _ = + Outer_Syntax.command @{command_keyword ML_command} "diagnostic ML text (silent)" + (Parse.ML_source >> Isar_Cmd.ml_diag false); + +val _ = + Outer_Syntax.command @{command_keyword setup} "ML setup for global theory" + (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.setup)); + +val _ = + Outer_Syntax.local_theory @{command_keyword local_setup} "ML setup for local theory" + (Parse.ML_source >> Isar_Cmd.local_setup); + +val _ = + Outer_Syntax.command @{command_keyword oracle} "declare oracle" + (Parse.range Parse.name -- (@{keyword "="} |-- Parse.ML_source) >> + (fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y))); + +val _ = + Outer_Syntax.local_theory @{command_keyword attribute_setup} "define attribute in ML" + (Parse.position Parse.name -- + Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "") + >> (fn (name, (txt, cmt)) => Attrib.attribute_setup name txt cmt)); + +val _ = + Outer_Syntax.local_theory @{command_keyword method_setup} "define proof method in ML" + (Parse.position Parse.name -- + Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "") + >> (fn (name, (txt, cmt)) => Method.method_setup name txt cmt)); + +val _ = + Outer_Syntax.local_theory @{command_keyword 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 @{command_keyword 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 @{command_keyword simproc_setup} "define simproc in ML" + (Parse.position Parse.name -- + (@{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)); + +in end\ + + +subsection \Theory commands\ + +subsubsection \Sorts and types\ + +ML \ +local + +val _ = + Outer_Syntax.local_theory @{command_keyword 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)); + +val _ = + Outer_Syntax.local_theory @{command_keyword typedecl} "type declaration" + (Parse.type_args -- Parse.binding -- Parse.opt_mixfix + >> (fn ((args, a), mx) => + Typedecl.typedecl {final = true} (a, map (rpair dummyS) args, mx) #> snd)); + +val _ = + Outer_Syntax.local_theory @{command_keyword type_synonym} "declare type abbreviation" + (Parse.type_args -- Parse.binding -- + (@{keyword "="} |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix')) + >> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs)); + +in end\ + + +subsubsection \Consts\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword judgment} "declare object-logic judgment" + (Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword consts} "declare constants" + (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts_cmd)); + +in end\ + + +subsubsection \Syntax and translations\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword nonterminal} + "declare syntactic type constructors (grammar nonterminal symbols)" + (Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global)); + +val _ = + Outer_Syntax.command @{command_keyword syntax} "add raw syntax clauses" + (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl + >> (Toplevel.theory o uncurry Sign.add_syntax_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword no_syntax} "delete raw syntax clauses" + (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl + >> (Toplevel.theory o uncurry Sign.del_syntax_cmd)); + +val trans_pat = + Scan.optional + (@{keyword "("} |-- Parse.!!! (Parse.inner_syntax Parse.xname --| @{keyword ")"})) "logic" + -- Parse.inner_syntax Parse.string; + +fun trans_arrow 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 @{command_keyword translations} "add syntax translation rules" + (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.translations)); + +val _ = + Outer_Syntax.command @{command_keyword no_translations} "delete syntax translation rules" + (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.no_translations)); + +in end\ + + +subsubsection \Translation functions\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword parse_ast_translation} + "install parse ast translation functions" + (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_ast_translation)); + +val _ = + Outer_Syntax.command @{command_keyword parse_translation} + "install parse translation functions" + (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_translation)); + +val _ = + Outer_Syntax.command @{command_keyword print_translation} + "install print translation functions" + (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_translation)); + +val _ = + Outer_Syntax.command @{command_keyword typed_print_translation} + "install typed print translation functions" + (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.typed_print_translation)); + +val _ = + Outer_Syntax.command @{command_keyword print_ast_translation} + "install print ast translation functions" + (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_ast_translation)); + +in end\ + + +subsubsection \Specifications\ + +ML \ +local + +val _ = + Outer_Syntax.local_theory' @{command_keyword definition} "constant definition" + (Parse_Spec.constdef >> (fn args => #2 oo Specification.definition_cmd args)); + +val _ = + Outer_Syntax.local_theory' @{command_keyword abbreviation} "constant abbreviation" + (Parse.syntax_mode -- (Scan.option Parse_Spec.constdecl -- Parse.prop) + >> (fn (mode, args) => Specification.abbreviation_cmd mode args)); + +val _ = + Outer_Syntax.command @{command_keyword 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))); + +in end\ + + +subsubsection \Notation\ + +ML \ +local + +val _ = + Outer_Syntax.local_theory @{command_keyword type_notation} + "add concrete syntax for type constructors" + (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) + >> (fn (mode, args) => Specification.type_notation_cmd true mode args)); + +val _ = + Outer_Syntax.local_theory @{command_keyword no_type_notation} + "delete concrete syntax for type constructors" + (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) + >> (fn (mode, args) => Specification.type_notation_cmd false mode args)); + +val _ = + Outer_Syntax.local_theory @{command_keyword notation} + "add concrete syntax for constants / fixed variables" + (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) + >> (fn (mode, args) => Specification.notation_cmd true mode args)); + +val _ = + Outer_Syntax.local_theory @{command_keyword no_notation} + "delete concrete syntax for constants / fixed variables" + (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) + >> (fn (mode, args) => Specification.notation_cmd false mode args)); + +in end\ + + +subsubsection \Theorems\ + +ML \ +local + +val _ = + Outer_Syntax.local_theory' @{command_keyword lemmas} "define theorems" + (Parse_Spec.name_facts -- Parse.for_fixes >> + (fn (facts, fixes) => #2 oo Specification.theorems_cmd Thm.theoremK facts fixes)); + +val _ = + Outer_Syntax.local_theory' @{command_keyword declare} "declare theorems" + (Parse.and_list1 Parse.xthms1 -- Parse.for_fixes + >> (fn (facts, fixes) => + #2 oo Specification.theorems_cmd "" [(Attrib.empty_binding, flat facts)] fixes)); + +val _ = + Outer_Syntax.local_theory @{command_keyword named_theorems} + "declare named collection of theorems" + (Parse.and_list1 (Parse.binding -- Scan.optional Parse.text "") >> + fold (fn (b, descr) => snd o Named_Theorems.declare b descr)); + +in end\ + + +subsubsection \Hide names\ + +ML \ +local + +fun hide_names command_keyword what hide parse prep = + Outer_Syntax.command command_keyword ("hide " ^ what ^ " from name space") + ((Parse.opt_keyword "open" >> not) -- Scan.repeat1 parse >> (fn (fully, args) => + (Toplevel.theory (fn thy => + let val ctxt = Proof_Context.init_global thy + in fold (hide fully o prep ctxt) args thy end)))); + +val _ = + hide_names @{command_keyword hide_class} "classes" Sign.hide_class Parse.class + Proof_Context.read_class; + +val _ = + hide_names @{command_keyword hide_type} "types" Sign.hide_type Parse.type_const + ((#1 o dest_Type) oo Proof_Context.read_type_name {proper = true, strict = false}); + +val _ = + hide_names @{command_keyword hide_const} "consts" Sign.hide_const Parse.const + ((#1 o dest_Const) oo Proof_Context.read_const {proper = true, strict = false}); + +val _ = + hide_names @{command_keyword hide_fact} "facts" Global_Theory.hide_fact + (Parse.position Parse.xname) (Global_Theory.check_fact o Proof_Context.theory_of); + +in end\ + + +subsection \Bundled declarations\ + +ML \ +local + +val _ = + Outer_Syntax.local_theory @{command_keyword bundle} "define bundle of declarations" + ((Parse.binding --| @{keyword "="}) -- Parse.xthms1 -- Parse.for_fixes + >> (uncurry Bundle.bundle_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword include} + "include declarations from bundle in proof body" + (Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.proof o Bundle.include_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword including} + "include declarations from bundle in goal refinement" + (Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.proof o Bundle.including_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword print_bundles} + "print bundles of declarations" + (Parse.opt_bang >> (fn b => Toplevel.keep (Bundle.print_bundles b o Toplevel.context_of))); + +in end\ + + +subsection \Local theory specifications\ + +subsubsection \Specification context\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword context} "begin local theory context" + ((Parse.position Parse.xname >> (fn name => + Toplevel.begin_local_theory true (Named_Target.begin name)) || + Scan.optional Parse_Spec.includes [] -- Scan.repeat Parse_Spec.context_element + >> (fn (incls, elems) => Toplevel.open_target (#2 o Bundle.context_cmd incls elems))) + --| Parse.begin); + +val _ = + Outer_Syntax.command @{command_keyword end} "end context" + (Scan.succeed + (Toplevel.exit o Toplevel.end_local_theory o Toplevel.close_target o + Toplevel.end_proof (K Proof.end_notepad))); + +in end\ + + +subsubsection \Locales and interpretation\ + +ML \ +local + +val locale_val = + Parse_Spec.locale_expression -- + Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] || + Scan.repeat1 Parse_Spec.context_element >> pair ([], []); + +val _ = + Outer_Syntax.command @{command_keyword locale} "define named specification context" + (Parse.binding -- + Scan.optional (@{keyword "="} |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin + >> (fn ((name, (expr, elems)), begin) => + Toplevel.begin_local_theory begin + (Expression.add_locale_cmd name Binding.empty expr elems #> snd))); + +val _ = + Outer_Syntax.command @{command_keyword experiment} "open private specification context" + (Scan.repeat Parse_Spec.context_element --| Parse.begin + >> (fn elems => + Toplevel.begin_local_theory true (Experiment.experiment_cmd elems #> snd))); + +val interpretation_args = + Parse.!!! Parse_Spec.locale_expression -- + Scan.optional + (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []; + +val _ = + Outer_Syntax.command @{command_keyword interpret} + "prove interpretation of locale expression in proof context" + (interpretation_args >> (fn (expr, equations) => + Toplevel.proof (Interpretation.interpret_cmd expr equations))); + +val interpretation_args_with_defs = + Parse.!!! Parse_Spec.locale_expression -- + (Scan.optional (@{keyword "defines"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" + -- ((Parse.binding -- Parse.opt_mixfix') --| @{keyword "="} -- Parse.term))) [] -- + Scan.optional + (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []); + +val _ = + Outer_Syntax.local_theory_to_proof @{command_keyword global_interpretation} + "prove interpretation of locale expression into global theory" + (interpretation_args_with_defs >> (fn (expr, (defs, equations)) => + Interpretation.global_interpretation_cmd expr defs equations)); + +val _ = + Outer_Syntax.command @{command_keyword sublocale} + "prove sublocale relation between a locale and a locale expression" + ((Parse.position Parse.xname --| (@{keyword "\"} || @{keyword "<"}) -- + interpretation_args_with_defs >> (fn (loc, (expr, (defs, equations))) => + Toplevel.theory_to_proof (Interpretation.global_sublocale_cmd loc expr defs equations))) + || interpretation_args_with_defs >> (fn (expr, (defs, equations)) => + Toplevel.local_theory_to_proof NONE NONE (Interpretation.sublocale_cmd expr defs equations))); + +val _ = + Outer_Syntax.command @{command_keyword interpretation} + "prove interpretation of locale expression in local theory or into global theory" + (interpretation_args >> (fn (expr, equations) => + Toplevel.local_theory_to_proof NONE NONE + (Interpretation.isar_interpretation_cmd expr equations))); + +in end\ + + +subsubsection \Type classes\ + +ML \ +local + +val class_val = + Parse_Spec.class_expression -- + Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] || + Scan.repeat1 Parse_Spec.context_element >> pair []; + +val _ = + Outer_Syntax.command @{command_keyword class} "define type class" + (Parse.binding -- Scan.optional (@{keyword "="} |-- class_val) ([], []) -- Parse.opt_begin + >> (fn ((name, (supclasses, elems)), begin) => + Toplevel.begin_local_theory begin + (Class_Declaration.class_cmd name supclasses elems #> snd))); + +val _ = + Outer_Syntax.local_theory_to_proof @{command_keyword subclass} "prove a subclass relation" + (Parse.class >> Class_Declaration.subclass_cmd); + +val _ = + Outer_Syntax.command @{command_keyword instantiation} "instantiate and prove type arity" + (Parse.multi_arity --| Parse.begin + >> (fn arities => Toplevel.begin_local_theory true (Class.instantiation_cmd arities))); + +val _ = + Outer_Syntax.command @{command_keyword instance} "prove type arity or subclass relation" + ((Parse.class -- + ((@{keyword "\"} || @{keyword "<"}) |-- Parse.!!! Parse.class) >> Class.classrel_cmd || + Parse.multi_arity >> Class.instance_arity_cmd) >> Toplevel.theory_to_proof || + Scan.succeed (Toplevel.local_theory_to_proof NONE NONE (Class.instantiation_instance I))); + +in end\ + + +subsubsection \Arbitrary overloading\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword overloading} "overloaded definitions" + (Scan.repeat1 (Parse.name --| (@{keyword "=="} || @{keyword "\"}) -- Parse.term -- + Scan.optional (@{keyword "("} |-- (@{keyword "unchecked"} >> K false) --| @{keyword ")"}) true + >> Scan.triple1) --| Parse.begin + >> (fn operations => Toplevel.begin_local_theory true (Overloading.overloading_cmd operations))); + +in end\ + + +subsection \Proof commands\ + +ML \ +local + +val _ = + Outer_Syntax.local_theory_to_proof @{command_keyword notepad} "begin proof context" + (Parse.begin >> K Proof.begin_notepad); + +in end\ + + +subsubsection \Statements\ + +ML \ +local + +val structured_statement = + Parse_Spec.statement -- Parse_Spec.cond_statement -- Parse.for_fixes + >> (fn ((shows, (strict, assumes)), fixes) => (strict, fixes, assumes, shows)); + +fun theorem spec schematic descr = + Outer_Syntax.local_theory_to_proof' spec + ("state " ^ descr) + (Scan.optional (Parse_Spec.opt_thm_name ":" --| + Scan.ahead (Parse_Spec.includes >> K "" || + Parse_Spec.locale_keyword || Parse_Spec.statement_keyword)) Attrib.empty_binding -- + Scan.optional Parse_Spec.includes [] -- + Parse_Spec.general_statement >> (fn ((a, includes), (elems, concl)) => + ((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd) + Thm.theoremK NONE (K I) a includes elems concl))); + +val _ = theorem @{command_keyword theorem} false "theorem"; +val _ = theorem @{command_keyword lemma} false "lemma"; +val _ = theorem @{command_keyword corollary} false "corollary"; +val _ = theorem @{command_keyword proposition} false "proposition"; +val _ = theorem @{command_keyword schematic_goal} true "schematic goal"; -section \Basic attributes\ +val _ = + Outer_Syntax.command @{command_keyword have} "state local goal" + (structured_statement >> (fn (a, b, c, d) => + Toplevel.proof' (fn int => Proof.have_cmd a NONE (K I) b c d int #> #2))); + +val _ = + Outer_Syntax.command @{command_keyword show} "state local goal, to refine pending subgoals" + (structured_statement >> (fn (a, b, c, d) => + Toplevel.proof' (fn int => Proof.show_cmd a NONE (K I) b c d int #> #2))); + +val _ = + Outer_Syntax.command @{command_keyword hence} "old-style alias of \"then have\"" + (structured_statement >> (fn (a, b, c, d) => + Toplevel.proof' (fn int => Proof.chain #> Proof.have_cmd a NONE (K I) b c d int #> #2))); + +val _ = + Outer_Syntax.command @{command_keyword thus} "old-style alias of \"then show\"" + (structured_statement >> (fn (a, b, c, d) => + Toplevel.proof' (fn int => Proof.chain #> Proof.show_cmd a NONE (K I) b c d int #> #2))); + +in end\ + + +subsubsection \Local facts\ + +ML \ +local + +val facts = Parse.and_list1 Parse.xthms1; + +val _ = + Outer_Syntax.command @{command_keyword then} "forward chaining" + (Scan.succeed (Toplevel.proof Proof.chain)); + +val _ = + Outer_Syntax.command @{command_keyword from} "forward chaining from given facts" + (facts >> (Toplevel.proof o Proof.from_thmss_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword with} "forward chaining from given and current facts" + (facts >> (Toplevel.proof o Proof.with_thmss_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword note} "define facts" + (Parse_Spec.name_facts >> (Toplevel.proof o Proof.note_thmss_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword supply} "define facts during goal refinement (unstructured)" + (Parse_Spec.name_facts >> (Toplevel.proof o Proof.supply_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword using} "augment goal facts" + (facts >> (Toplevel.proof o Proof.using_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword unfolding} "unfold definitions in goal and facts" + (facts >> (Toplevel.proof o Proof.unfolding_cmd)); + +in end\ + + +subsubsection \Proof context\ + +ML \ +local + +val structured_statement = + Parse_Spec.statement -- Parse_Spec.if_statement' -- Parse.for_fixes + >> (fn ((shows, assumes), fixes) => (fixes, assumes, shows)); + +val _ = + Outer_Syntax.command @{command_keyword fix} "fix local variables (Skolem constants)" + (Parse.fixes >> (Toplevel.proof o Proof.fix_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword assume} "assume propositions" + (structured_statement >> (fn (a, b, c) => Toplevel.proof (Proof.assume_cmd a b c))); + +val _ = + Outer_Syntax.command @{command_keyword presume} "assume propositions, to be established later" + (structured_statement >> (fn (a, b, c) => Toplevel.proof (Proof.presume_cmd a b c))); + +val _ = + Outer_Syntax.command @{command_keyword def} "local definition (non-polymorphic)" + (Parse.and_list1 + (Parse_Spec.opt_thm_name ":" -- + ((Parse.binding -- Parse.opt_mixfix) -- + ((@{keyword "\"} || @{keyword "=="}) |-- Parse.!!! Parse.termp))) + >> (Toplevel.proof o Proof.def_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword consider} "state cases rule" + (Parse_Spec.obtains >> (Toplevel.proof' o Obtain.consider_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword obtain} "generalized elimination" + (Parse.parbinding -- Scan.optional (Parse.fixes --| Parse.where_) [] -- Parse_Spec.statement + >> (fn ((x, y), z) => Toplevel.proof' (Obtain.obtain_cmd x y z))); + +val _ = + Outer_Syntax.command @{command_keyword guess} "wild guessing (unstructured)" + (Scan.optional Parse.fixes [] >> (Toplevel.proof' o Obtain.guess_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword let} "bind text variables" + (Parse.and_list1 (Parse.and_list1 Parse.term -- (@{keyword "="} |-- Parse.term)) + >> (Toplevel.proof o Proof.let_bind_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword write} "add concrete syntax for constants / fixed variables" + (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) + >> (fn (mode, args) => Toplevel.proof (Proof.write_cmd mode args))); + +val _ = + Outer_Syntax.command @{command_keyword case} "invoke local context" + (Parse_Spec.opt_thm_name ":" -- + (@{keyword "("} |-- + Parse.!!! (Parse.position Parse.xname -- Scan.repeat (Parse.maybe Parse.binding) + --| @{keyword ")"}) || + Parse.position Parse.xname >> rpair []) >> (Toplevel.proof o Proof.case_cmd)); + +in end\ + + +subsubsection \Proof structure\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword "{"} "begin explicit proof block" + (Scan.succeed (Toplevel.proof Proof.begin_block)); + +val _ = + Outer_Syntax.command @{command_keyword "}"} "end explicit proof block" + (Scan.succeed (Toplevel.proof Proof.end_block)); + +val _ = + Outer_Syntax.command @{command_keyword next} "enter next proof block" + (Scan.succeed (Toplevel.proof Proof.next_block)); + +in end\ + + +subsubsection \End proof\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword qed} "conclude proof" + (Scan.option Method.parse >> (fn m => + (Option.map Method.report m; + Isar_Cmd.qed m))); + +val _ = + Outer_Syntax.command @{command_keyword by} "terminal backward proof" + (Method.parse -- Scan.option Method.parse >> (fn (m1, m2) => + (Method.report m1; + Option.map Method.report m2; + Isar_Cmd.terminal_proof (m1, m2)))); + +val _ = + Outer_Syntax.command @{command_keyword ".."} "default proof" + (Scan.succeed Isar_Cmd.default_proof); + +val _ = + Outer_Syntax.command @{command_keyword "."} "immediate proof" + (Scan.succeed Isar_Cmd.immediate_proof); + +val _ = + Outer_Syntax.command @{command_keyword done} "done proof" + (Scan.succeed Isar_Cmd.done_proof); + +val _ = + Outer_Syntax.command @{command_keyword sorry} "skip proof (quick-and-dirty mode only!)" + (Scan.succeed Isar_Cmd.skip_proof); + +val _ = + Outer_Syntax.command @{command_keyword "\"} "dummy proof (quick-and-dirty mode only!)" + (Scan.succeed Isar_Cmd.skip_proof); + +val _ = + Outer_Syntax.command @{command_keyword oops} "forget proof" + (Scan.succeed (Toplevel.forget_proof true)); + +in end\ + + +subsubsection \Proof steps\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword defer} "shuffle internal proof state" + (Scan.optional Parse.nat 1 >> (Toplevel.proof o Proof.defer)); + +val _ = + Outer_Syntax.command @{command_keyword prefer} "shuffle internal proof state" + (Parse.nat >> (Toplevel.proof o Proof.prefer)); + +val _ = + Outer_Syntax.command @{command_keyword apply} "initial goal refinement step (unstructured)" + (Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply m)))); + +val _ = + Outer_Syntax.command @{command_keyword apply_end} "terminal goal refinement step (unstructured)" + (Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply_end m)))); + +val _ = + Outer_Syntax.command @{command_keyword proof} "backward proof step" + (Scan.option Method.parse >> (fn m => + (Option.map Method.report m; Toplevel.proofs (Proof.proof m)))); + +in end\ + + +subsubsection \Subgoal focus\ + +ML \ +local + +val opt_fact_binding = + Scan.optional (Parse.binding -- Parse.opt_attribs || Parse.attribs >> pair Binding.empty) + Attrib.empty_binding; + +val for_params = + Scan.optional + (@{keyword "for"} |-- + Parse.!!! ((Scan.option Parse.dots >> is_some) -- + (Scan.repeat1 (Parse.position (Parse.maybe Parse.name))))) + (false, []); + +val _ = + Outer_Syntax.command @{command_keyword subgoal} + "focus on first subgoal within backward refinement" + (opt_fact_binding -- (Scan.option (@{keyword "premises"} |-- Parse.!!! opt_fact_binding)) -- + for_params >> (fn ((a, b), c) => + Toplevel.proofs (Seq.make_results o Seq.single o #2 o Subgoal.subgoal_cmd a b c))); + +in end\ + + +subsubsection \Calculation\ + +ML \ +local + +val calculation_args = + Scan.option (@{keyword "("} |-- Parse.!!! ((Parse.xthms1 --| @{keyword ")"}))); + +val _ = + Outer_Syntax.command @{command_keyword also} "combine calculation and current facts" + (calculation_args >> (Toplevel.proofs' o Calculation.also_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword finally} + "combine calculation and current facts, exhibit result" + (calculation_args >> (Toplevel.proofs' o Calculation.finally_cmd)); + +val _ = + Outer_Syntax.command @{command_keyword moreover} "augment calculation by current facts" + (Scan.succeed (Toplevel.proof' Calculation.moreover)); + +val _ = + Outer_Syntax.command @{command_keyword ultimately} + "augment calculation by current facts, exhibit result" + (Scan.succeed (Toplevel.proof' Calculation.ultimately)); + +val _ = + Outer_Syntax.command @{command_keyword print_trans_rules} "print transitivity rules" + (Scan.succeed (Toplevel.keep (Calculation.print_rules o Toplevel.context_of))); + +in end\ + + +subsubsection \Proof navigation\ + +ML \ +local + +fun report_back () = + Output.report [Markup.markup Markup.bad "Explicit backtracking"]; + +val _ = + Outer_Syntax.command @{command_keyword back} "explicit backtracking of proof command" + (Scan.succeed + (Toplevel.actual_proof (fn prf => (report_back (); Proof_Node.back prf)) o + Toplevel.skip_proof report_back)); + +in end\ + + +subsection \Diagnostic commands (for interactive mode only)\ + +ML \ +local + +val opt_modes = + Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) []; + +val _ = + Outer_Syntax.command @{command_keyword help} + "retrieve outer syntax commands according to name patterns" + (Scan.repeat Parse.name >> + (fn pats => Toplevel.keep (fn st => Outer_Syntax.help (Toplevel.theory_of st) pats))); + +val _ = + Outer_Syntax.command @{command_keyword print_commands} "print outer syntax commands" + (Scan.succeed (Toplevel.keep (Outer_Syntax.print_commands o Toplevel.theory_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_options} "print configuration options" + (Parse.opt_bang >> (fn b => Toplevel.keep (Attrib.print_options b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_context} + "print context of local theory target" + (Scan.succeed (Toplevel.keep (Pretty.writeln_chunks o Toplevel.pretty_context))); + +val _ = + Outer_Syntax.command @{command_keyword print_theory} + "print logical theory contents" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Pretty.writeln o Proof_Display.pretty_theory b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_definitions} + "print dependencies of definitional theory content" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Pretty.writeln o Proof_Display.pretty_definitions b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_syntax} + "print inner syntax of context" + (Scan.succeed (Toplevel.keep (Proof_Context.print_syntax o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_defn_rules} + "print definitional rewrite rules of context" + (Scan.succeed (Toplevel.keep (Local_Defs.print_rules o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_abbrevs} + "print constant abbreviations of context" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Proof_Context.print_abbrevs b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_theorems} + "print theorems of local theory or proof context" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Pretty.writeln o Pretty.chunks o Isar_Cmd.pretty_theorems b))); + +val _ = + Outer_Syntax.command @{command_keyword print_locales} + "print locales of this theory" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Locale.print_locales b o Toplevel.theory_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_classes} + "print classes of this theory" + (Scan.succeed (Toplevel.keep (Class.print_classes o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_locale} + "print locale of this theory" + (Parse.opt_bang -- Parse.position Parse.xname >> (fn (b, name) => + Toplevel.keep (fn state => Locale.print_locale (Toplevel.theory_of state) b name))); + +val _ = + Outer_Syntax.command @{command_keyword print_interps} + "print interpretations of locale for this theory or proof context" + (Parse.position Parse.xname >> (fn name => + Toplevel.keep (fn state => Locale.print_registrations (Toplevel.context_of state) name))); + +val _ = + Outer_Syntax.command @{command_keyword print_dependencies} + "print dependencies of locale expression" + (Parse.opt_bang -- Parse_Spec.locale_expression >> (fn (b, expr) => + Toplevel.keep (fn state => Expression.print_dependencies (Toplevel.context_of state) b expr))); + +val _ = + Outer_Syntax.command @{command_keyword print_attributes} + "print attributes of this theory" + (Parse.opt_bang >> (fn b => Toplevel.keep (Attrib.print_attributes b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_simpset} + "print context of Simplifier" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Pretty.writeln o Simplifier.pretty_simpset b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_rules} "print intro/elim rules" + (Scan.succeed (Toplevel.keep (Context_Rules.print_rules o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_methods} "print methods of this theory" + (Parse.opt_bang >> (fn b => Toplevel.keep (Method.print_methods b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_antiquotations} + "print document antiquotations" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Thy_Output.print_antiquotations b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_ML_antiquotations} + "print ML antiquotations" + (Parse.opt_bang >> (fn b => + Toplevel.keep (ML_Context.print_antiquotations b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword locale_deps} "visualize locale dependencies" + (Scan.succeed + (Toplevel.keep (Toplevel.theory_of #> (fn thy => + Locale.pretty_locale_deps thy + |> map (fn {name, parents, body} => + ((name, Graph_Display.content_node (Locale.extern thy name) [body]), parents)) + |> Graph_Display.display_graph_old)))); + +val _ = + Outer_Syntax.command @{command_keyword print_term_bindings} + "print term bindings of proof context" + (Scan.succeed + (Toplevel.keep + (Pretty.writeln_chunks o Proof_Context.pretty_term_bindings o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_facts} "print facts of proof context" + (Parse.opt_bang >> (fn b => + Toplevel.keep (Proof_Context.print_local_facts b o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_cases} "print cases of proof context" + (Scan.succeed + (Toplevel.keep (Pretty.writeln_chunks o Proof_Context.pretty_cases o Toplevel.context_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_statement} + "print theorems as long statements" + (opt_modes -- Parse.xthms1 >> Isar_Cmd.print_stmts); + +val _ = + Outer_Syntax.command @{command_keyword thm} "print theorems" + (opt_modes -- Parse.xthms1 >> Isar_Cmd.print_thms); + +val _ = + Outer_Syntax.command @{command_keyword prf} "print proof terms of theorems" + (opt_modes -- Scan.option Parse.xthms1 >> Isar_Cmd.print_prfs false); + +val _ = + Outer_Syntax.command @{command_keyword full_prf} "print full proof terms of theorems" + (opt_modes -- Scan.option Parse.xthms1 >> Isar_Cmd.print_prfs true); + +val _ = + Outer_Syntax.command @{command_keyword prop} "read and print proposition" + (opt_modes -- Parse.term >> Isar_Cmd.print_prop); + +val _ = + Outer_Syntax.command @{command_keyword term} "read and print term" + (opt_modes -- Parse.term >> Isar_Cmd.print_term); + +val _ = + Outer_Syntax.command @{command_keyword typ} "read and print type" + (opt_modes -- (Parse.typ -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort)) + >> Isar_Cmd.print_type); + +val _ = + Outer_Syntax.command @{command_keyword print_codesetup} "print code generator setup" + (Scan.succeed (Toplevel.keep (Code.print_codesetup o Toplevel.theory_of))); + +val _ = + Outer_Syntax.command @{command_keyword print_state} + "print current proof state (if present)" + (opt_modes >> (fn modes => + Toplevel.keep (Print_Mode.with_modes modes (Output.state o Toplevel.string_of_state)))); + +val _ = + Outer_Syntax.command @{command_keyword welcome} "print welcome message" + (Scan.succeed (Toplevel.keep (fn _ => writeln (Session.welcome ())))); + +val _ = + Outer_Syntax.command @{command_keyword display_drafts} + "display raw source files with symbols" + (Scan.repeat1 Parse.path >> (fn names => + Toplevel.keep (fn _ => ignore (Present.display_drafts (map Path.explode names))))); + +in end\ + + +subsection \Dependencies\ + +ML \ +local + +val theory_bounds = + Parse.position Parse.theory_xname >> single || + (@{keyword "("} |-- Parse.enum "|" (Parse.position Parse.theory_xname) --| @{keyword ")"}); + +val _ = + Outer_Syntax.command @{command_keyword thy_deps} "visualize theory dependencies" + (Scan.option theory_bounds -- Scan.option theory_bounds >> + (fn args => Toplevel.keep (fn st => Thy_Deps.thy_deps_cmd (Toplevel.context_of st) args))); + + +val class_bounds = + Parse.sort >> single || + (@{keyword "("} |-- Parse.enum "|" Parse.sort --| @{keyword ")"}); + +val _ = + Outer_Syntax.command @{command_keyword class_deps} "visualize class dependencies" + (Scan.option class_bounds -- Scan.option class_bounds >> (fn args => + Toplevel.keep (fn st => Class_Deps.class_deps_cmd (Toplevel.context_of st) args))); + + +val _ = + Outer_Syntax.command @{command_keyword thm_deps} "visualize theorem dependencies" + (Parse.xthms1 >> (fn args => + Toplevel.keep (fn st => + Thm_Deps.thm_deps (Toplevel.theory_of st) + (Attrib.eval_thms (Toplevel.context_of st) args)))); + + +val thy_names = + Scan.repeat1 (Scan.unless Parse.minus (Parse.position Parse.theory_xname)); + +val _ = + Outer_Syntax.command @{command_keyword unused_thms} "find unused theorems" + (Scan.option ((thy_names --| Parse.minus) -- Scan.option thy_names) >> (fn opt_range => + Toplevel.keep (fn st => + let + val thy = Toplevel.theory_of st; + val ctxt = Toplevel.context_of st; + fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]); + val check = Theory.check ctxt; + in + Thm_Deps.unused_thms + (case opt_range of + NONE => (Theory.parents_of thy, [thy]) + | SOME (xs, NONE) => (map check xs, [thy]) + | SOME (xs, SOME ys) => (map check xs, map check ys)) + |> map pretty_thm |> Pretty.writeln_chunks + end))); + +in end\ + + +subsubsection \Find consts and theorems\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword find_consts} + "find constants by name / type patterns" + (Find_Consts.query_parser >> (fn spec => + Toplevel.keep (fn st => + Pretty.writeln (Find_Consts.pretty_consts (Toplevel.context_of st) spec)))); + +val options = + Scan.optional + (Parse.$$$ "(" |-- + Parse.!!! (Scan.option Parse.nat -- + Scan.optional (Parse.reserved "with_dups" >> K false) true --| Parse.$$$ ")")) + (NONE, true); + +val _ = + Outer_Syntax.command @{command_keyword find_theorems} + "find theorems meeting specified criteria" + (options -- Find_Theorems.query_parser >> (fn ((opt_lim, rem_dups), spec) => + Toplevel.keep (fn st => + Pretty.writeln + (Find_Theorems.pretty_theorems (Find_Theorems.proof_state st) opt_lim rem_dups spec)))); + +in end\ + + +subsection \Code generation\ + +ML \ +local + +val _ = + Outer_Syntax.command @{command_keyword code_datatype} + "define set of code datatype constructors" + (Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd)); + +in end\ + + +subsection \Extraction of programs from proofs\ + +ML \ +local + +val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") []; + +val _ = + Outer_Syntax.command @{command_keyword 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_keyword realizability} + "add equations characterizing realizability" + (Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_realizes_eqns)); + +val _ = + Outer_Syntax.command @{command_keyword 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_keyword 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))); + +in end\ + + +section \Isar attributes\ attribute_setup tagged = \Scan.lift (Args.name -- Args.name) >> Thm.tag\ diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ROOT --- a/src/Pure/ROOT Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/ROOT Mon Apr 04 22:13:47 2016 +0200 @@ -61,6 +61,7 @@ "Isar/attrib.ML" "Isar/auto_bind.ML" "Isar/bundle.ML" + "Isar/calculation.ML" "Isar/class.ML" "Isar/class_declaration.ML" "Isar/code.ML" @@ -97,8 +98,8 @@ "Isar/typedecl.ML" "ML/exn_debugger.ML" "ML/exn_properties.ML" - "ML/fixed_int_dummy.ML" "ML/ml_antiquotation.ML" + "ML/ml_antiquotations.ML" "ML/ml_compiler.ML" "ML/ml_compiler0.ML" "ML/ml_compiler1.ML" @@ -106,18 +107,19 @@ "ML/ml_context.ML" "ML/ml_debugger.ML" "ML/ml_env.ML" - "ML/ml_file.ML" "ML/ml_heap.ML" "ML/ml_lex.ML" "ML/ml_name_space.ML" "ML/ml_options.ML" + "ML/ml_pervasive_final.ML" + "ML/ml_pervasive_initial.ML" "ML/ml_pp.ML" - "ML/ml_pervasive.ML" "ML/ml_pretty.ML" "ML/ml_profiling.ML" "ML/ml_statistics.ML" "ML/ml_syntax.ML" "ML/ml_system.ML" + "ML/ml_thms.ML" "PIDE/active.ML" "PIDE/command.ML" "PIDE/command_span.ML" @@ -152,8 +154,8 @@ "Syntax/term_position.ML" "Syntax/type_annotation.ML" "System/bash.ML" - "System/bash_windows.ML" "System/command_line.ML" + "System/distribution.ML" "System/invoke_scala.ML" "System/isabelle_process.ML" "System/isabelle_system.ML" @@ -170,9 +172,22 @@ "Thy/thy_info.ML" "Thy/thy_output.ML" "Thy/thy_syntax.ML" + "Tools/bibtex.ML" "Tools/build.ML" + "Tools/class_deps.ML" + "Tools/debugger.ML" + "Tools/find_consts.ML" + "Tools/find_theorems.ML" + "Tools/jedit.ML" + "Tools/named_theorems.ML" "Tools/named_thms.ML" "Tools/plugin.ML" + "Tools/print_operation.ML" + "Tools/rail.ML" + "Tools/rule_insts.ML" + "Tools/simplifier_trace.ML" + "Tools/thm_deps.ML" + "Tools/thy_deps.ML" "assumption.ML" "axclass.ML" "config.ML" diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ROOT.ML --- a/src/Pure/ROOT.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/ROOT.ML Mon Apr 04 22:13:47 2016 +0200 @@ -4,19 +4,9 @@ (* initial ML name space *) +use "ML/ml_name_space.ML"; +use "ML/ml_pervasive_initial.ML"; use "ML/ml_system.ML"; -use "ML/ml_name_space.ML"; -use "ML/ml_pervasive.ML"; - -if List.exists (fn (a, _) => a = "FixedInt") (#allStruct ML_Name_Space.global ()) then () -else use "ML/fixed_int_dummy.ML"; - -structure Distribution = (*filled-in by makedist*) -struct - val version = "unidentified repository version"; - val is_identified = false; - val is_official = false; -end; (* multithreading *) @@ -29,6 +19,8 @@ (* ML system *) +use "System/distribution.ML"; + use "ML/ml_heap.ML"; use "ML/ml_profiling.ML"; use "ML/ml_pretty.ML"; @@ -90,8 +82,6 @@ use "General/timing.ML"; use "General/sha1.ML"; -val _ = List.app ML_Name_Space.forget_structure ML_Name_Space.critical_structures; - use "PIDE/yxml.ML"; use "PIDE/document_id.ML"; @@ -248,6 +238,7 @@ use "Isar/element.ML"; use "Isar/obtain.ML"; use "Isar/subgoal.ML"; +use "Isar/calculation.ML"; (*local theories and targets*) use "Isar/locale.ML"; @@ -286,11 +277,10 @@ use "Proof/extraction.ML"; (*Isabelle system*) -if ML_System.platform_is_windows -then use "System/bash_windows.ML" -else use "System/bash.ML"; +use "System/bash.ML"; use "System/isabelle_system.ML"; + (*theory documents*) use "Thy/term_style.ML"; use "Isar/outer_syntax.ML"; @@ -323,28 +313,26 @@ (* miscellaneous tools and packages for Pure Isabelle *) +use "ML/ml_pp.ML"; +use "ML/ml_antiquotations.ML"; +use "ML/ml_thms.ML"; + use "Tools/build.ML"; use "Tools/named_thms.ML"; - -structure Output: OUTPUT = Output; (*seal system channels!*) - -use "ML/ml_pp.ML"; - - -(* the Pure theory *) +use "Tools/print_operation.ML"; +use "Tools/bibtex.ML"; +use "Tools/rail.ML"; +use "Tools/rule_insts.ML"; +use "Tools/thm_deps.ML"; +use "Tools/thy_deps.ML"; +use "Tools/class_deps.ML"; +use "Tools/find_theorems.ML"; +use "Tools/find_consts.ML"; +use "Tools/simplifier_trace.ML"; +use_no_debug "Tools/debugger.ML"; +use "Tools/named_theorems.ML"; +use "Tools/jedit.ML"; -use "ML/ml_file.ML"; -Runtime.toplevel_program (fn () => Thy_Info.use_thy ("Pure", Position.none)); -Context.set_thread_data NONE; -structure Pure = struct val thy = Thy_Info.pure_theory () end; - - -(* ML toplevel commands *) +use_thy "Pure"; -fun use_thys args = - Runtime.toplevel_program (fn () => Thy_Info.use_thys (map (rpair Position.none) args)); -val use_thy = use_thys o single; - -Proofterm.proofs := 0; - -structure PolyML = struct structure IntInf = PolyML.IntInf end; +use "ML/ml_pervasive_final.ML"; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/ROOT.scala --- a/src/Pure/ROOT.scala Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/ROOT.scala Mon Apr 04 22:13:47 2016 +0200 @@ -7,13 +7,6 @@ package object isabelle { - object Distribution /*filled-in by makedist*/ - { - val version = "unidentified repository version" - val is_identified = false - val is_official = false - } - val ERROR = Exn.ERROR val error = Exn.error _ val cat_error = Exn.cat_error _ diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/System/bash.ML --- a/src/Pure/System/bash.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/System/bash.ML Mon Apr 04 22:13:47 2016 +0200 @@ -1,7 +1,7 @@ (* Title: Pure/System/bash.ML Author: Makarius -GNU bash processes, with propagation of interrupts -- POSIX version. +GNU bash processes, with propagation of interrupts. *) signature BASH = @@ -9,6 +9,99 @@ val process: string -> {out: string, err: string, rc: int, terminate: unit -> unit} end; +if ML_System.platform_is_windows then ML +\ +structure Bash: BASH = +struct + +val process = uninterruptible (fn restore_attributes => fn script => + let + datatype result = Wait | Signal | Result of int; + val result = Synchronized.var "bash_result" Wait; + + val id = serial_string (); + val script_path = File.tmp_path (Path.basic ("bash_script" ^ id)); + val out_path = File.tmp_path (Path.basic ("bash_out" ^ id)); + val err_path = File.tmp_path (Path.basic ("bash_err" ^ id)); + val pid_path = File.tmp_path (Path.basic ("bash_pid" ^ id)); + + fun cleanup_files () = + (try File.rm script_path; + try File.rm out_path; + try File.rm err_path; + try File.rm pid_path); + val _ = cleanup_files (); + + val system_thread = + Standard_Thread.fork {name = "bash", stack_limit = NONE, interrupts = false} (fn () => + Multithreading.with_attributes Multithreading.private_interrupts (fn _ => + let + val _ = File.write script_path script; + val bash_script = + "bash " ^ File.bash_path script_path ^ + " > " ^ File.bash_path out_path ^ + " 2> " ^ File.bash_path err_path; + val bash_process = getenv_strict "ISABELLE_BASH_PROCESS"; + val rc = + Windows.simpleExecute ("", + quote (ML_System.platform_path bash_process) ^ " " ^ + quote (File.platform_path pid_path) ^ " \"\" bash -c " ^ quote bash_script) + |> Windows.fromStatus |> SysWord.toInt; + val res = if rc = 130 orelse rc = 512 then Signal else Result rc; + in Synchronized.change result (K res) end + handle exn => + (Synchronized.change result (fn Wait => Signal | res => res); Exn.reraise exn))); + + fun read_pid 0 = NONE + | read_pid count = + (case (Int.fromString (File.read pid_path) handle IO.Io _ => NONE) of + NONE => (OS.Process.sleep (seconds 0.1); read_pid (count - 1)) + | some => some); + + fun terminate NONE = () + | terminate (SOME pid) = + let + fun kill s = + let + val cmd = getenv_strict "CYGWIN_ROOT" ^ "\\bin\\bash.exe"; + val arg = "kill -" ^ s ^ " -" ^ string_of_int pid; + in + OS.Process.isSuccess (Windows.simpleExecute ("", quote cmd ^ " -c " ^ quote arg)) + handle OS.SysErr _ => false + end; + + fun multi_kill count s = + count = 0 orelse + (kill s; kill "0") andalso + (OS.Process.sleep (seconds 0.1); multi_kill (count - 1) s); + val _ = + multi_kill 10 "INT" andalso + multi_kill 10 "TERM" andalso + multi_kill 10 "KILL"; + in () end; + + fun cleanup () = + (Standard_Thread.interrupt_unsynchronized system_thread; + cleanup_files ()); + in + let + val _ = + restore_attributes (fn () => + Synchronized.guarded_access result (fn Wait => NONE | x => SOME ((), x))) (); + + val out = the_default "" (try File.read out_path); + val err = the_default "" (try File.read err_path); + val rc = (case Synchronized.value result of Signal => Exn.interrupt () | Result rc => rc); + val pid = read_pid 1; + val _ = cleanup (); + in {out = out, err = err, rc = rc, terminate = fn () => terminate pid} end + handle exn => (terminate (read_pid 10); cleanup (); Exn.reraise exn) + end); + +end; +\ +else ML +\ structure Bash: BASH = struct @@ -99,4 +192,4 @@ end); end; - +\; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/System/bash_windows.ML --- a/src/Pure/System/bash_windows.ML Mon Apr 04 16:52:56 2016 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,99 +0,0 @@ -(* Title: Pure/System/bash_windows.ML - Author: Makarius - -GNU bash processes, with propagation of interrupts -- Windows version. -*) - -signature BASH = -sig - val process: string -> {out: string, err: string, rc: int, terminate: unit -> unit} -end; - -structure Bash: BASH = -struct - -val process = uninterruptible (fn restore_attributes => fn script => - let - datatype result = Wait | Signal | Result of int; - val result = Synchronized.var "bash_result" Wait; - - val id = serial_string (); - val script_path = File.tmp_path (Path.basic ("bash_script" ^ id)); - val out_path = File.tmp_path (Path.basic ("bash_out" ^ id)); - val err_path = File.tmp_path (Path.basic ("bash_err" ^ id)); - val pid_path = File.tmp_path (Path.basic ("bash_pid" ^ id)); - - fun cleanup_files () = - (try File.rm script_path; - try File.rm out_path; - try File.rm err_path; - try File.rm pid_path); - val _ = cleanup_files (); - - val system_thread = - Standard_Thread.fork {name = "bash", stack_limit = NONE, interrupts = false} (fn () => - Multithreading.with_attributes Multithreading.private_interrupts (fn _ => - let - val _ = File.write script_path script; - val bash_script = - "bash " ^ File.bash_path script_path ^ - " > " ^ File.bash_path out_path ^ - " 2> " ^ File.bash_path err_path; - val bash_process = getenv_strict "ISABELLE_BASH_PROCESS"; - val rc = - Windows.simpleExecute ("", - quote (ML_System.platform_path bash_process) ^ " " ^ - quote (File.platform_path pid_path) ^ " \"\" bash -c " ^ quote bash_script) - |> Windows.fromStatus |> SysWord.toInt; - val res = if rc = 130 orelse rc = 512 then Signal else Result rc; - in Synchronized.change result (K res) end - handle exn => - (Synchronized.change result (fn Wait => Signal | res => res); Exn.reraise exn))); - - fun read_pid 0 = NONE - | read_pid count = - (case (Int.fromString (File.read pid_path) handle IO.Io _ => NONE) of - NONE => (OS.Process.sleep (seconds 0.1); read_pid (count - 1)) - | some => some); - - fun terminate NONE = () - | terminate (SOME pid) = - let - fun kill s = - let - val cmd = getenv_strict "CYGWIN_ROOT" ^ "\\bin\\bash.exe"; - val arg = "kill -" ^ s ^ " -" ^ string_of_int pid; - in - OS.Process.isSuccess (Windows.simpleExecute ("", quote cmd ^ " -c " ^ quote arg)) - handle OS.SysErr _ => false - end; - - fun multi_kill count s = - count = 0 orelse - (kill s; kill "0") andalso - (OS.Process.sleep (seconds 0.1); multi_kill (count - 1) s); - val _ = - multi_kill 10 "INT" andalso - multi_kill 10 "TERM" andalso - multi_kill 10 "KILL"; - in () end; - - fun cleanup () = - (Standard_Thread.interrupt_unsynchronized system_thread; - cleanup_files ()); - in - let - val _ = - restore_attributes (fn () => - Synchronized.guarded_access result (fn Wait => NONE | x => SOME ((), x))) (); - - val out = the_default "" (try File.read out_path); - val err = the_default "" (try File.read err_path); - val rc = (case Synchronized.value result of Signal => Exn.interrupt () | Result rc => rc); - val pid = read_pid 1; - val _ = cleanup (); - in {out = out, err = err, rc = rc, terminate = fn () => terminate pid} end - handle exn => (terminate (read_pid 10); cleanup (); Exn.reraise exn) - end); - -end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/System/distribution.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/System/distribution.ML Mon Apr 04 22:13:47 2016 +0200 @@ -0,0 +1,12 @@ +(* Title: Pure/System/distribution.ML + Author: Makarius + +The Isabelle system distribution -- filled-in by makedist. +*) + +structure Distribution = +struct + val version = "unidentified repository version"; + val is_identified = false; + val is_official = false; +end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/System/distribution.scala --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/Pure/System/distribution.scala Mon Apr 04 22:13:47 2016 +0200 @@ -0,0 +1,15 @@ +/* Title: Pure/System/distribution.scala + Author: Makarius + +The Isabelle system distribution -- filled-in by makedist. +*/ + +package isabelle + + +object Distribution +{ + val version = "unidentified repository version" + val is_identified = false + val is_official = false +} diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Thy/thy_header.ML --- a/src/Pure/Thy/thy_header.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Thy/thy_header.ML Mon Apr 04 22:13:47 2016 +0200 @@ -80,9 +80,7 @@ ((txtN, @{here}), SOME ((Keyword.document_body, []), [])), ((text_rawN, @{here}), SOME ((Keyword.document_raw, []), [])), ((theoryN, @{here}), SOME ((Keyword.thy_begin, []), ["theory"])), - (("ML_file", @{here}), SOME ((Keyword.thy_load, []), ["ML"])), - (("ML_file_debug", @{here}), SOME ((Keyword.thy_load, []), ["ML"])), - (("ML_file_no_debug", @{here}), SOME ((Keyword.thy_load, []), ["ML"]))]; + (("ML", @{here}), SOME ((Keyword.thy_decl, []), ["ML"]))]; (* theory data *) diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Thy/thy_header.scala --- a/src/Pure/Thy/thy_header.scala Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Thy/thy_header.scala Mon Apr 04 22:13:47 2016 +0200 @@ -57,9 +57,7 @@ (TXT, Some(((Keyword.DOCUMENT_BODY, Nil), Nil)), None), (TEXT_RAW, Some(((Keyword.DOCUMENT_RAW, Nil), Nil)), None), (THEORY, Some((Keyword.THY_BEGIN, Nil), List("theory")), None), - ("ML_file", Some((Keyword.THY_LOAD, Nil), List("ML")), None), - ("ML_file_debug", Some((Keyword.THY_LOAD, Nil), List("ML")), None), - ("ML_file_no_debug", Some((Keyword.THY_LOAD, Nil), List("ML")), None)) + ("ML", Some((Keyword.THY_DECL, Nil), List("ML")), None)) private val bootstrap_keywords = Keyword.Keywords.empty.add_keywords(bootstrap_header) diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Thy/thy_info.ML --- a/src/Pure/Thy/thy_info.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Thy/thy_info.ML Mon Apr 04 22:13:47 2016 +0200 @@ -388,3 +388,5 @@ fun finish () = change_thys (String_Graph.map (fn _ => fn (_, entry) => (NONE, entry))); end; + +fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy (name, Position.none)); diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Tools/class_deps.ML --- a/src/Pure/Tools/class_deps.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Tools/class_deps.ML Mon Apr 04 22:13:47 2016 +0200 @@ -37,13 +37,4 @@ val class_deps = gen_class_deps (Type.cert_sort o Proof_Context.tsig_of); val class_deps_cmd = Graph_Display.display_graph oo gen_class_deps Syntax.read_sort; -val class_bounds = - Parse.sort >> single || - (@{keyword "("} |-- Parse.enum "|" Parse.sort --| @{keyword ")"}); - -val _ = - Outer_Syntax.command @{command_keyword class_deps} "visualize class dependencies" - (Scan.option class_bounds -- Scan.option class_bounds >> (fn args => - Toplevel.keep (fn st => class_deps_cmd (Toplevel.context_of st) args))); - end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Tools/find_consts.ML --- a/src/Pure/Tools/find_consts.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Tools/find_consts.ML Mon Apr 04 22:13:47 2016 +0200 @@ -11,6 +11,8 @@ Strict of string | Loose of string | Name of string + val pretty_consts: Proof.context -> (bool * criterion) list -> Pretty.T + val query_parser: (bool * criterion) list parser val read_query: Position.T -> string -> (bool * criterion) list val find_consts : Proof.context -> (bool * criterion) list -> unit end; @@ -138,25 +140,18 @@ Parse.reserved "strict" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.typ) >> Strict || Parse.typ >> Loose; -val query = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion); - val query_keywords = Keyword.add_keywords [((":", @{here}), NONE)] Keyword.empty_keywords; in +val query_parser = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion); + fun read_query pos str = Token.explode query_keywords pos str |> filter Token.is_proper - |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query --| Scan.ahead Parse.eof))) + |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query_parser --| Scan.ahead Parse.eof))) |> #1; -val _ = - Outer_Syntax.command @{command_keyword find_consts} - "find constants by name / type patterns" - (query >> (fn spec => - Toplevel.keep (fn st => - Pretty.writeln (pretty_consts (Toplevel.context_of st) spec)))); - end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Tools/find_theorems.ML --- a/src/Pure/Tools/find_theorems.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Tools/find_theorems.ML Mon Apr 04 22:13:47 2016 +0200 @@ -15,12 +15,16 @@ rem_dups: bool, criteria: (bool * 'term criterion) list } + val query_parser: (bool * string criterion) list parser val read_query: Position.T -> string -> (bool * string criterion) list val find_theorems: Proof.context -> thm option -> int option -> bool -> (bool * term criterion) list -> int option * (Facts.ref * thm) list val find_theorems_cmd: Proof.context -> thm option -> int option -> bool -> (bool * string criterion) list -> int option * (Facts.ref * thm) list val pretty_thm: Proof.context -> Facts.ref * thm -> Pretty.T + val pretty_theorems: Proof.state -> + int option -> bool -> (bool * string criterion) list -> Pretty.T + val proof_state: Toplevel.state -> Proof.state end; structure Find_Theorems: FIND_THEOREMS = @@ -502,11 +506,6 @@ (** Isar command syntax **) -fun proof_state st = - (case try Toplevel.proof_of st of - SOME state => state - | NONE => Proof.init (Toplevel.context_of st)); - local val criterion = @@ -518,37 +517,29 @@ Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Simp || Parse.term >> Pattern; -val options = - Scan.optional - (Parse.$$$ "(" |-- - Parse.!!! (Scan.option Parse.nat -- Scan.optional (Parse.reserved "with_dups" >> K false) true - --| Parse.$$$ ")")) (NONE, true); - -val query = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion); - val query_keywords = Keyword.add_keywords [((":", @{here}), NONE)] Keyword.empty_keywords; in +val query_parser = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion); + fun read_query pos str = Token.explode query_keywords pos str |> filter Token.is_proper - |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query --| Scan.ahead Parse.eof))) + |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query_parser --| Scan.ahead Parse.eof))) |> #1; -val _ = - Outer_Syntax.command @{command_keyword find_theorems} - "find theorems meeting specified criteria" - (options -- query >> (fn ((opt_lim, rem_dups), spec) => - Toplevel.keep (fn st => - Pretty.writeln (pretty_theorems (proof_state st) opt_lim rem_dups spec)))); - end; (** PIDE query operation **) +fun proof_state st = + (case try Toplevel.proof_of st of + SOME state => state + | NONE => Proof.init (Toplevel.context_of st)); + val _ = Query_Operation.register {name = "find_theorems", pri = Task_Queue.urgent_pri} (fn {state = st, args, writeln_result, ...} => @@ -563,4 +554,3 @@ else error "Unknown context"); end; - diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Tools/ml_process.scala --- a/src/Pure/Tools/ml_process.scala Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Tools/ml_process.scala Mon Apr 04 22:13:47 2016 +0200 @@ -39,6 +39,8 @@ val eval_init = if (heaps.isEmpty) { List( + if (Isabelle_System.getenv("ML_SYSTEM") == "polyml-5.6") "structure FixedInt = IntInf" + else "", if (Platform.is_windows) "fun exit 0 = OS.Process.exit OS.Process.success" + " | exit 1 = OS.Process.exit OS.Process.failure" + diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Tools/named_theorems.ML --- a/src/Pure/Tools/named_theorems.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Tools/named_theorems.ML Mon Apr 04 22:13:47 2016 +0200 @@ -93,12 +93,6 @@ |> Attrib.local_setup binding (Attrib.add_del (add name) (del name)) description in (name, lthy') end; -val _ = - Outer_Syntax.local_theory @{command_keyword named_theorems} - "declare named collection of theorems" - (Parse.and_list1 (Parse.binding -- Scan.optional Parse.text "") >> - fold (fn (b, descr) => snd o declare b descr)); - (* ML antiquotation *) diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Tools/thm_deps.ML --- a/src/Pure/Tools/thm_deps.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Tools/thm_deps.ML Mon Apr 04 22:13:47 2016 +0200 @@ -47,12 +47,6 @@ |> Graph_Display.display_graph_old end; -val _ = - Outer_Syntax.command @{command_keyword thm_deps} "visualize theorem dependencies" - (Parse.xthms1 >> (fn args => - Toplevel.keep (fn st => - thm_deps (Toplevel.theory_of st) (Attrib.eval_thms (Toplevel.context_of st) args)))); - (* unused_thms *) @@ -107,25 +101,4 @@ else q) new_thms ([], Inttab.empty); in rev thms' end; -val thy_names = - Scan.repeat1 (Scan.unless Parse.minus (Parse.position Parse.theory_xname)); - -val _ = - Outer_Syntax.command @{command_keyword unused_thms} "find unused theorems" - (Scan.option ((thy_names --| Parse.minus) -- Scan.option thy_names) >> (fn opt_range => - Toplevel.keep (fn st => - let - val thy = Toplevel.theory_of st; - val ctxt = Toplevel.context_of st; - fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]); - val check = Theory.check ctxt; - in - unused_thms - (case opt_range of - NONE => (Theory.parents_of thy, [thy]) - | SOME (xs, NONE) => (map check xs, [thy]) - | SOME (xs, SOME ys) => (map check xs, map check ys)) - |> map pretty_thm |> Pretty.writeln_chunks - end))); - end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/Tools/thy_deps.ML --- a/src/Pure/Tools/thy_deps.ML Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/Tools/thy_deps.ML Mon Apr 04 22:13:47 2016 +0200 @@ -42,13 +42,4 @@ val thy_deps_cmd = Graph_Display.display_graph oo gen_thy_deps Theory.check; -val theory_bounds = - Parse.position Parse.theory_xname >> single || - (@{keyword "("} |-- Parse.enum "|" (Parse.position Parse.theory_xname) --| @{keyword ")"}); - -val _ = - Outer_Syntax.command @{command_keyword thy_deps} "visualize theory dependencies" - (Scan.option theory_bounds -- Scan.option theory_bounds >> - (fn args => Toplevel.keep (fn st => thy_deps_cmd (Toplevel.context_of st) args))); - end; diff -r 313d3b697c9a -r a8758f47f9e8 src/Pure/build-jars --- a/src/Pure/build-jars Mon Apr 04 16:52:56 2016 +0100 +++ b/src/Pure/build-jars Mon Apr 04 22:13:47 2016 +0200 @@ -76,6 +76,7 @@ System/bash.scala System/command_line.scala System/cygwin.scala + System/distribution.scala System/getopts.scala System/invoke_scala.scala System/isabelle_charset.scala