| author | wenzelm | 
| Tue, 22 Oct 2024 12:45:38 +0200 | |
| changeset 81229 | e18600daa904 | 
| parent 81170 | 2d73c3287bd3 | 
| child 81533 | fb49af893986 | 
| permissions | -rw-r--r-- | 
| 48929 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 1 | (* Title: Pure/Pure.thy | 
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 2 | Author: Makarius | 
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 3 | |
| 62944 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62913diff
changeset | 4 | The Pure theory, with definitions of Isar commands and some lemmas. | 
| 48929 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 5 | *) | 
| 
05d4e5f660ae
entity markup for theory Pure, to enable hyperlinks etc.;
 wenzelm parents: 
48891diff
changeset | 6 | |
| 48638 | 7 | theory Pure | 
| 62859 | 8 | keywords | 
| 67446 | 9 | "!!" "!" "+" ":" ";" "<" "<=" "==" "=>" "?" "[" "\<comment>" "\<equiv>" "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>" | 
| 80709 
e6f026505c5b
support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
 wenzelm parents: 
80635diff
changeset | 10 | "\<subseteq>" "]" "binder" "by" "congproc" "identifier" "in" "infix" "infixl" "infixr" "is" "open" | 
| 
e6f026505c5b
support for congprocs in the Simplifier, closely following Norbert Schirmer et-al, but with only one "simproc" name space and "simproc_setup" command / ML antiquotation;
 wenzelm parents: 
80635diff
changeset | 11 | "output" "overloaded" "passive" "pervasive" "premises" "structure" "unchecked" "weak_congproc" | 
| 63441 | 12 | and "private" "qualified" :: before_command | 
| 67740 | 13 | and "assumes" "constrains" "defines" "fixes" "for" "if" "includes" "notes" "rewrites" | 
| 66194 | 14 | "obtains" "shows" "when" "where" "|" :: quasi_command | 
| 58999 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 15 | and "text" "txt" :: document_body | 
| 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 wenzelm parents: 
58928diff
changeset | 16 | and "text_raw" :: document_raw | 
| 63579 | 17 | and "default_sort" :: thy_decl | 
| 80750 | 18 | and "typedecl" "nonterminal" "judgment" "consts" "syntax" "no_syntax" "syntax_consts" | 
| 19 | "syntax_types" "translations" "no_translations" "type_notation" "no_type_notation" "notation" | |
| 20 | "no_notation" "alias" "type_alias" "declare" "hide_class" "hide_type" "hide_const" | |
| 21 | "hide_fact" :: thy_decl | |
| 69913 | 22 | and "type_synonym" "definition" "abbreviation" "lemmas" :: thy_defn | 
| 23 | and "axiomatization" :: thy_stmt | |
| 72837 | 24 | and "external_file" "bibtex_file" "ROOTS_file" :: thy_load | 
| 69381 
4c9b4e2c5460
more general command 'generate_file' for registered file types, notably Haskell;
 wenzelm parents: 
69349diff
changeset | 25 | and "generate_file" :: thy_decl | 
| 69662 | 26 | and "export_generated_files" :: diag | 
| 75686 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 27 | and "scala_build_generated_files" :: diag | 
| 70051 | 28 | and "compile_generated_files" :: diag and "external_files" "export_files" "export_prefix" | 
| 75679 
aa89255b704c
support for classpath artifacts within session structure:
 wenzelm parents: 
75660diff
changeset | 29 | and "export_classpath" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 30 | and "ML_file" "ML_file_debug" "ML_file_no_debug" :: thy_load % "ML" | 
| 60957 
574254152856
support for ML files with/without debugger information;
 wenzelm parents: 
60749diff
changeset | 31 | and "SML_file" "SML_file_debug" "SML_file_no_debug" :: thy_load % "ML" | 
| 68276 | 32 | and "SML_import" "SML_export" "ML_export" :: thy_decl % "ML" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 33 | and "ML_prf" :: prf_decl % "proof" (* FIXME % "ML" ?? *) | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 34 | and "ML_val" "ML_command" :: diag % "ML" | 
| 63579 | 35 | and "simproc_setup" :: thy_decl % "ML" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 36 | and "setup" "local_setup" "attribute_setup" "method_setup" | 
| 55762 
27a45aec67a0
suppress completion of obscure keyword, avoid confusion with plain "simp";
 wenzelm parents: 
55516diff
changeset | 37 | "declaration" "syntax_declaration" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 38 | "parse_ast_translation" "parse_translation" "print_translation" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 39 | "typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML" | 
| 81106 | 40 | and "bundle" "open_bundle" :: thy_decl_block | 
| 63282 | 41 | and "unbundle" :: thy_decl | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 42 | and "include" "including" :: prf_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 43 | and "print_bundles" :: diag | 
| 59901 | 44 | and "context" "locale" "experiment" :: thy_decl_block | 
| 51224 
c3e99efacb67
back to non-schematic 'sublocale' and 'interpretation' (despite df8fc0567a3d) for more potential parallelism;
 wenzelm parents: 
50603diff
changeset | 45 | and "interpret" :: prf_goal % "proof" | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61853diff
changeset | 46 | and "interpretation" "global_interpretation" "sublocale" :: thy_goal | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 47 | and "class" :: thy_decl_block | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 48 | and "subclass" :: thy_goal | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 49 | and "instantiation" :: thy_decl_block | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 50 | and "instance" :: thy_goal | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 51 | and "overloading" :: thy_decl_block | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 52 | and "opening" :: quasi_command | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 53 | and "code_datatype" :: thy_decl | 
| 69913 | 54 | and "theorem" "lemma" "corollary" "proposition" :: thy_goal_stmt | 
| 55 | and "schematic_goal" :: thy_goal_stmt | |
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 56 | and "notepad" :: thy_decl_block | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 57 | and "have" :: prf_goal % "proof" | 
| 63579 | 58 | and "hence" :: prf_goal % "proof" | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 59 | and "show" :: prf_asm_goal % "proof" | 
| 63579 | 60 | and "thus" :: prf_asm_goal % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 61 | and "then" "from" "with" :: prf_chain % "proof" | 
| 60371 | 62 | and "note" :: prf_decl % "proof" | 
| 63 | and "supply" :: prf_script % "proof" | |
| 64 | and "using" "unfolding" :: prf_decl % "proof" | |
| 67119 | 65 | and "fix" "assume" "presume" "define" :: prf_asm % "proof" | 
| 60448 | 66 | and "consider" :: prf_goal % "proof" | 
| 53371 
47b23c582127
more explicit indication of 'guess' as improper Isar (aka "script") element;
 wenzelm parents: 
52549diff
changeset | 67 | and "obtain" :: prf_asm_goal % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 68 | and "let" "write" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 69 | and "case" :: prf_asm % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 70 |   and "{" :: prf_open % "proof"
 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 71 | and "}" :: prf_close % "proof" | 
| 60694 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 wenzelm parents: 
60624diff
changeset | 72 | and "next" :: next_block % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 73 | and "qed" :: qed_block % "proof" | 
| 62312 
5e5a881ebc12
command '\<proof>' is an alias for 'sorry', with different typesetting;
 wenzelm parents: 
62169diff
changeset | 74 | and "by" ".." "." "sorry" "\<proof>" :: "qed" % "proof" | 
| 53571 
e58ca0311c0f
more explicit indication of 'done' as proof script element;
 wenzelm parents: 
53371diff
changeset | 75 | and "done" :: "qed_script" % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 76 | and "oops" :: qed_global % "proof" | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 77 | and "defer" "prefer" "apply" :: prf_script % "proof" | 
| 63579 | 78 | and "apply_end" :: prf_script % "proof" | 
| 60624 | 79 | and "subgoal" :: prf_script_goal % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 80 | and "proof" :: prf_block % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 81 | and "also" "moreover" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 82 | and "finally" "ultimately" :: prf_chain % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 83 | and "back" :: prf_script % "proof" | 
| 61252 | 84 | and "help" "print_commands" "print_options" "print_context" "print_theory" | 
| 85 | "print_definitions" "print_syntax" "print_abbrevs" "print_defn_rules" | |
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 86 | "print_theorems" "print_locales" "print_classes" "print_locale" | 
| 71166 | 87 | "print_interps" "print_attributes" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 88 | "print_simpset" "print_rules" "print_trans_rules" "print_methods" | 
| 56069 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 wenzelm parents: 
55762diff
changeset | 89 | "print_antiquotations" "print_ML_antiquotations" "thy_deps" | 
| 70560 | 90 | "locale_deps" "class_deps" "thm_deps" "thm_oracles" "print_term_bindings" | 
| 57415 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 wenzelm parents: 
56864diff
changeset | 91 | "print_facts" "print_cases" "print_statement" "thm" "prf" "full_prf" | 
| 78526 | 92 | "prop" "term" "typ" "print_codesetup" "print_context_tracing" "unused_thms" :: diag | 
| 67263 | 93 | and "print_state" :: diag | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 94 | and "welcome" :: diag | 
| 68505 | 95 | and "end" :: thy_end | 
| 63579 | 96 | and "realizers" :: thy_decl | 
| 97 | and "realizability" :: thy_decl | |
| 56797 | 98 | and "extract_type" "extract" :: thy_decl | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 99 | and "find_theorems" "find_consts" :: diag | 
| 57886 
7cae177c9084
support for named collections of theorems in canonical order;
 wenzelm parents: 
57506diff
changeset | 100 | and "named_theorems" :: thy_decl | 
| 70143 | 101 | abbrevs "\\tag" = "\<^marker>\<open>tag \<close>" | 
| 102 | and "===>" = "===>" (*prevent replacement of very long arrows*) | |
| 67013 
335a7dce7cb3
more uniform header syntax, in contrast to the former etc/abbrevs file-format (see 73939a9b70a3);
 wenzelm parents: 
66757diff
changeset | 103 | and "--->" = "\<midarrow>\<rightarrow>" | 
| 67724 | 104 | and "hence" "thus" "default_sort" "simproc_setup" "apply_end" "realizers" "realizability" = "" | 
| 67013 
335a7dce7cb3
more uniform header syntax, in contrast to the former etc/abbrevs file-format (see 73939a9b70a3);
 wenzelm parents: 
66757diff
changeset | 105 | and "hence" = "then have" | 
| 
335a7dce7cb3
more uniform header syntax, in contrast to the former etc/abbrevs file-format (see 73939a9b70a3);
 wenzelm parents: 
66757diff
changeset | 106 | and "thus" = "then show" | 
| 48638 | 107 | begin | 
| 15803 | 108 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 109 | section \<open>Isar commands\<close> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 110 | |
| 67281 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 111 | subsection \<open>Other files\<close> | 
| 66757 | 112 | |
| 113 | ML \<open> | |
| 67281 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 114 | local | 
| 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 115 | val _ = | 
| 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 116 | Outer_Syntax.command \<^command_keyword>\<open>external_file\<close> "formal dependency on external file" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 117 | (Resources.provide_parse_file >> (fn get_file => Toplevel.theory (#2 o get_file))); | 
| 67281 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 118 | |
| 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 119 | val _ = | 
| 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 120 | Outer_Syntax.command \<^command_keyword>\<open>bibtex_file\<close> "check bibtex database file in Prover IDE" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 121 | (Resources.provide_parse_file >> (fn get_file => | 
| 69262 
f94726501b37
more standard Resources.provide_parse_files: avoid duplicate markup reports;
 wenzelm parents: 
69059diff
changeset | 122 | Toplevel.theory (fn thy => | 
| 67281 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 123 | let | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 124 |             val ({lines, pos, ...}, thy') = get_file thy;
 | 
| 69262 
f94726501b37
more standard Resources.provide_parse_files: avoid duplicate markup reports;
 wenzelm parents: 
69059diff
changeset | 125 | val _ = Bibtex.check_database_output pos (cat_lines lines); | 
| 
f94726501b37
more standard Resources.provide_parse_files: avoid duplicate markup reports;
 wenzelm parents: 
69059diff
changeset | 126 | in thy' end))); | 
| 69381 
4c9b4e2c5460
more general command 'generate_file' for registered file types, notably Haskell;
 wenzelm parents: 
69349diff
changeset | 127 | |
| 
4c9b4e2c5460
more general command 'generate_file' for registered file types, notably Haskell;
 wenzelm parents: 
69349diff
changeset | 128 | val _ = | 
| 72837 | 129 | Outer_Syntax.command \<^command_keyword>\<open>ROOTS_file\<close> "session ROOTS file" | 
| 130 | (Resources.provide_parse_file >> (fn get_file => | |
| 131 | Toplevel.theory (fn thy => | |
| 132 | let | |
| 133 |             val ({src_path, lines, pos = pos0, ...}, thy') = get_file thy;
 | |
| 134 | val ctxt = Proof_Context.init_global thy'; | |
| 135 | val dir = Path.dir (Path.expand (Resources.master_directory thy' + src_path)); | |
| 136 | val _ = | |
| 137 | (lines, pos0) |-> fold (fn line => fn pos1 => | |
| 138 | let | |
| 74174 | 139 | val pos2 = Position.symbol_explode line pos1; | 
| 72841 | 140 | val range = Position.range (pos1, pos2); | 
| 73312 | 141 | val source = Input.source true line range; | 
| 72837 | 142 | val _ = | 
| 72838 
27a7a5499511
more accurate syntax, following Sessions.parse_roots in Scala;
 wenzelm parents: 
72837diff
changeset | 143 | if line = "" then () | 
| 
27a7a5499511
more accurate syntax, following Sessions.parse_roots in Scala;
 wenzelm parents: 
72837diff
changeset | 144 | else if String.isPrefix "#" line then | 
| 72841 | 145 | Context_Position.report ctxt (#1 range) Markup.comment | 
| 72838 
27a7a5499511
more accurate syntax, following Sessions.parse_roots in Scala;
 wenzelm parents: 
72837diff
changeset | 146 | else | 
| 73312 | 147 | (ignore (Resources.check_session_dir ctxt (SOME dir) source) | 
| 72838 
27a7a5499511
more accurate syntax, following Sessions.parse_roots in Scala;
 wenzelm parents: 
72837diff
changeset | 148 | handle ERROR msg => Output.error_message msg); | 
| 74174 | 149 | in pos2 |> Position.symbol "\n" end); | 
| 72837 | 150 | in thy' end))); | 
| 151 | ||
| 152 | val _ = | |
| 69381 
4c9b4e2c5460
more general command 'generate_file' for registered file types, notably Haskell;
 wenzelm parents: 
69349diff
changeset | 153 | Outer_Syntax.local_theory \<^command_keyword>\<open>generate_file\<close> | 
| 
4c9b4e2c5460
more general command 'generate_file' for registered file types, notably Haskell;
 wenzelm parents: 
69349diff
changeset | 154 | "generate source file, with antiquotations" | 
| 70205 | 155 | (Parse.path_binding -- (\<^keyword>\<open>=\<close> |-- Parse.embedded_input) | 
| 69383 | 156 | >> Generated_Files.generate_file_cmd); | 
| 69381 
4c9b4e2c5460
more general command 'generate_file' for registered file types, notably Haskell;
 wenzelm parents: 
69349diff
changeset | 157 | |
| 70051 | 158 | |
| 70056 | 159 | val files_in_theory = | 
| 70047 
96fe857a7a6f
clarified signature: more explicit operations for corresponding Isar commands;
 wenzelm parents: 
70009diff
changeset | 160 | (Parse.underscore >> K [] || Scan.repeat1 Parse.path_binding) -- | 
| 70054 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 161 | Scan.option (\<^keyword>\<open>(\<close> |-- Parse.!!! (\<^keyword>\<open>in\<close> |-- Parse.theory_name --| \<^keyword>\<open>)\<close>)); | 
| 70047 
96fe857a7a6f
clarified signature: more explicit operations for corresponding Isar commands;
 wenzelm parents: 
70009diff
changeset | 162 | |
| 69662 | 163 | val _ = | 
| 164 | Outer_Syntax.command \<^command_keyword>\<open>export_generated_files\<close> | |
| 70047 
96fe857a7a6f
clarified signature: more explicit operations for corresponding Isar commands;
 wenzelm parents: 
70009diff
changeset | 165 | "export generated files from given theories" | 
| 70056 | 166 | (Parse.and_list1 files_in_theory >> (fn args => | 
| 69662 | 167 | Toplevel.keep (fn st => | 
| 70054 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 168 | Generated_Files.export_generated_files_cmd (Toplevel.context_of st) args))); | 
| 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 169 | |
| 70047 
96fe857a7a6f
clarified signature: more explicit operations for corresponding Isar commands;
 wenzelm parents: 
70009diff
changeset | 170 | |
| 70054 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 171 | val base_dir = | 
| 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 172 | Scan.optional (\<^keyword>\<open>(\<close> |-- | 
| 72841 | 173 | Parse.!!! (\<^keyword>\<open>in\<close> |-- Parse.path_input --| \<^keyword>\<open>)\<close>)) (Input.string ""); | 
| 70054 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 174 | |
| 72841 | 175 | val external_files = Scan.repeat1 Parse.path_input -- base_dir; | 
| 70051 | 176 | |
| 177 | val exe = Parse.reserved "exe" >> K true || Parse.reserved "executable" >> K false; | |
| 70054 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 178 | val executable = \<^keyword>\<open>(\<close> |-- Parse.!!! (exe --| \<^keyword>\<open>)\<close>) >> SOME || Scan.succeed NONE; | 
| 70051 | 179 | |
| 180 | val export_files = Scan.repeat1 Parse.path_binding -- executable; | |
| 181 | ||
| 182 | val _ = | |
| 183 | Outer_Syntax.command \<^command_keyword>\<open>compile_generated_files\<close> | |
| 184 | "compile generated files and export results" | |
| 70056 | 185 | (Parse.and_list files_in_theory -- | 
| 70054 
a884b2967dd7
clarified export_files: Isabelle_System.copy_file_base preserves given directory sub-structure;
 wenzelm parents: 
70051diff
changeset | 186 | Scan.optional (\<^keyword>\<open>external_files\<close> |-- Parse.!!! (Parse.and_list1 external_files)) [] -- | 
| 70051 | 187 | Scan.optional (\<^keyword>\<open>export_files\<close> |-- Parse.!!! (Parse.and_list1 export_files)) [] -- | 
| 70055 
36fb663145e5
type Path.binding may be empty: check later via proper_binding;
 wenzelm parents: 
70054diff
changeset | 188 |         Scan.optional (\<^keyword>\<open>export_prefix\<close> |-- Parse.path_binding) ("", Position.none) --
 | 
| 70051 | 189 | (Parse.where_ |-- Parse.!!! Parse.ML_source) | 
| 190 | >> (fn ((((args, external), export), export_prefix), source) => | |
| 191 | Toplevel.keep (fn st => | |
| 192 | Generated_Files.compile_generated_files_cmd | |
| 193 | (Toplevel.context_of st) args external export export_prefix source))); | |
| 194 | ||
| 75660 
45d3497c0baa
support for Isabelle/Scala/Java modules in Isabelle/ML;
 wenzelm parents: 
74887diff
changeset | 195 | val _ = | 
| 75686 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 196 | Outer_Syntax.command \<^command_keyword>\<open>scala_build_generated_files\<close> | 
| 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 197 | "build and export Isabelle/Scala/Java module" | 
| 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 198 | (Parse.and_list files_in_theory -- | 
| 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 199 | Scan.optional (\<^keyword>\<open>external_files\<close> |-- Parse.!!! (Parse.and_list1 external_files)) [] | 
| 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 200 | >> (fn (args, external) => | 
| 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 201 | Toplevel.keep (fn st => | 
| 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 202 | Generated_Files.scala_build_generated_files_cmd | 
| 
42f19e398ee4
command 'scala_build_generated_files' with proper management of source dependencies;
 wenzelm parents: 
75679diff
changeset | 203 | (Toplevel.context_of st) args external))); | 
| 67281 
338fb884286b
added command 'bibtex_file' (for PIDE interaction only);
 wenzelm parents: 
67263diff
changeset | 204 | in end\<close> | 
| 66757 | 205 | |
| 76923 | 206 | external_file \<open>ROOT0.ML\<close> | 
| 207 | external_file \<open>ROOT.ML\<close> | |
| 72816 
ea4f86914cb2
support for PIDE markup for auxiliary files ("blobs");
 wenzelm parents: 
72749diff
changeset | 208 | |
| 66757 | 209 | |
| 62856 | 210 | subsection \<open>Embedded ML text\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 211 | |
| 62856 | 212 | ML \<open> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 213 | local | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 214 | |
| 67147 | 215 | val semi = Scan.option \<^keyword>\<open>;\<close>; | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 216 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 217 | val _ = | 
| 67147 | 218 | Outer_Syntax.command \<^command_keyword>\<open>ML_file\<close> "read and evaluate Isabelle/ML file" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 219 | (Resources.parse_file --| semi >> ML_File.ML NONE); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 220 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 221 | val _ = | 
| 67147 | 222 | Outer_Syntax.command \<^command_keyword>\<open>ML_file_debug\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 223 | "read and evaluate Isabelle/ML file (with debugger information)" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 224 | (Resources.parse_file --| semi >> ML_File.ML (SOME true)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 225 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 226 | val _ = | 
| 67147 | 227 | Outer_Syntax.command \<^command_keyword>\<open>ML_file_no_debug\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 228 | "read and evaluate Isabelle/ML file (no debugger information)" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 229 | (Resources.parse_file --| semi >> ML_File.ML (SOME false)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 230 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 231 | val _ = | 
| 67147 | 232 | Outer_Syntax.command \<^command_keyword>\<open>SML_file\<close> "read and evaluate Standard ML file" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 233 | (Resources.parse_file --| semi >> ML_File.SML NONE); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 234 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 235 | val _ = | 
| 67147 | 236 | Outer_Syntax.command \<^command_keyword>\<open>SML_file_debug\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 237 | "read and evaluate Standard ML file (with debugger information)" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 238 | (Resources.parse_file --| semi >> ML_File.SML (SOME true)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 239 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 240 | val _ = | 
| 67147 | 241 | Outer_Syntax.command \<^command_keyword>\<open>SML_file_no_debug\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 242 | "read and evaluate Standard ML file (no debugger information)" | 
| 72747 
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
 wenzelm parents: 
72536diff
changeset | 243 | (Resources.parse_file --| semi >> ML_File.SML (SOME false)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 244 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 245 | val _ = | 
| 67147 | 246 | Outer_Syntax.command \<^command_keyword>\<open>SML_export\<close> "evaluate SML within Isabelle/ML environment" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 247 | (Parse.ML_source >> (fn source => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 248 | let | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 249 | val flags: ML_Compiler.flags = | 
| 78728 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
78528diff
changeset | 250 |           {environment = ML_Env.SML_export, redirect = false, verbose = true, catch_all = true,
 | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 251 | debug = NONE, writeln = writeln, warning = warning}; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 252 | in | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 253 | Toplevel.theory | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 254 | (Context.theory_map (ML_Context.exec (fn () => ML_Context.eval_source flags source))) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 255 | end)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 256 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 257 | val _ = | 
| 67147 | 258 | Outer_Syntax.command \<^command_keyword>\<open>SML_import\<close> "evaluate Isabelle/ML within SML environment" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 259 | (Parse.ML_source >> (fn source => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 260 | let | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 261 | val flags: ML_Compiler.flags = | 
| 78728 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
78528diff
changeset | 262 |           {environment = ML_Env.SML_import, redirect = false, verbose = true, catch_all = true,
 | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 263 | debug = NONE, writeln = writeln, warning = warning}; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 264 | in | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 265 | Toplevel.generic_theory | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 266 | (ML_Context.exec (fn () => ML_Context.eval_source flags source) #> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 267 | Local_Theory.propagate_ml_env) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 268 | end)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 269 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 270 | val _ = | 
| 68276 | 271 |   Outer_Syntax.command ("ML_export", \<^here>)
 | 
| 272 | "ML text within theory or local theory, and export to bootstrap environment" | |
| 273 | (Parse.ML_source >> (fn source => | |
| 274 | Toplevel.generic_theory (fn context => | |
| 275 | context | |
| 68826 
deefe5837120
clarified ML_environment: ML_write_global requires "Isabelle";
 wenzelm parents: 
68820diff
changeset | 276 | |> Config.put_generic ML_Env.ML_environment ML_Env.Isabelle | 
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68813diff
changeset | 277 | |> Config.put_generic ML_Env.ML_write_global true | 
| 68276 | 278 | |> ML_Context.exec (fn () => | 
| 279 | ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source) | |
| 68827 | 280 | |> Config.restore_generic ML_Env.ML_write_global context | 
| 281 | |> Config.restore_generic ML_Env.ML_environment context | |
| 68276 | 282 | |> Local_Theory.propagate_ml_env))); | 
| 283 | ||
| 284 | val _ = | |
| 67147 | 285 | Outer_Syntax.command \<^command_keyword>\<open>ML_prf\<close> "ML text within proof" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 286 | (Parse.ML_source >> (fn source => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 287 | Toplevel.proof (Proof.map_context (Context.proof_map | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 288 | (ML_Context.exec (fn () => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 289 | ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source))) #> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 290 | Proof.propagate_ml_env))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 291 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 292 | val _ = | 
| 67147 | 293 | Outer_Syntax.command \<^command_keyword>\<open>ML_val\<close> "diagnostic ML text" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 294 | (Parse.ML_source >> Isar_Cmd.ml_diag true); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 295 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 296 | val _ = | 
| 67147 | 297 | Outer_Syntax.command \<^command_keyword>\<open>ML_command\<close> "diagnostic ML text (silent)" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 298 | (Parse.ML_source >> Isar_Cmd.ml_diag false); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 299 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 300 | val _ = | 
| 67147 | 301 | Outer_Syntax.command \<^command_keyword>\<open>setup\<close> "ML setup for global theory" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 302 | (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.setup)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 303 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 304 | val _ = | 
| 67147 | 305 | Outer_Syntax.local_theory \<^command_keyword>\<open>local_setup\<close> "ML setup for local theory" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 306 | (Parse.ML_source >> Isar_Cmd.local_setup); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 307 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 308 | val _ = | 
| 67147 | 309 | Outer_Syntax.command \<^command_keyword>\<open>oracle\<close> "declare oracle" | 
| 73787 | 310 | (Parse.range Parse.name -- Parse.!!! (\<^keyword>\<open>=\<close> |-- Parse.ML_source) >> | 
| 62856 | 311 | (fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y))); | 
| 312 | ||
| 313 | val _ = | |
| 67147 | 314 | Outer_Syntax.local_theory \<^command_keyword>\<open>attribute_setup\<close> "define attribute in ML" | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 315 | (Parse.name_position -- | 
| 74887 | 316 | Parse.!!! (\<^keyword>\<open>=\<close> |-- Parse.ML_source -- Scan.optional Parse.embedded "") | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 317 | >> (fn (name, (txt, cmt)) => Attrib.attribute_setup name txt cmt)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 318 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 319 | val _ = | 
| 67147 | 320 | Outer_Syntax.local_theory \<^command_keyword>\<open>method_setup\<close> "define proof method in ML" | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 321 | (Parse.name_position -- | 
| 74887 | 322 | Parse.!!! (\<^keyword>\<open>=\<close> |-- Parse.ML_source -- Scan.optional Parse.embedded "") | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 323 | >> (fn (name, (txt, cmt)) => Method.method_setup name txt cmt)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 324 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 325 | val _ = | 
| 67147 | 326 | Outer_Syntax.local_theory \<^command_keyword>\<open>declaration\<close> "generic ML declaration" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 327 | (Parse.opt_keyword "pervasive" -- Parse.ML_source | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 328 |       >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = false, pervasive = pervasive} txt));
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 329 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 330 | val _ = | 
| 67147 | 331 | Outer_Syntax.local_theory \<^command_keyword>\<open>syntax_declaration\<close> "generic ML syntax declaration" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 332 | (Parse.opt_keyword "pervasive" -- Parse.ML_source | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 333 |       >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = true, pervasive = pervasive} txt));
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 334 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 335 | val _ = | 
| 67147 | 336 | Outer_Syntax.local_theory \<^command_keyword>\<open>simproc_setup\<close> "define simproc in ML" | 
| 78803 | 337 | Simplifier.simproc_setup_command; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 338 | |
| 62856 | 339 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 340 | |
| 62856 | 341 | |
| 342 | subsection \<open>Theory commands\<close> | |
| 343 | ||
| 344 | subsubsection \<open>Sorts and types\<close> | |
| 345 | ||
| 346 | ML \<open> | |
| 347 | local | |
| 348 | ||
| 349 | val _ = | |
| 67147 | 350 | Outer_Syntax.local_theory \<^command_keyword>\<open>default_sort\<close> | 
| 62856 | 351 | "declare default sort for explicit type variables" | 
| 352 | (Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy)); | |
| 353 | ||
| 354 | val _ = | |
| 67147 | 355 | Outer_Syntax.local_theory \<^command_keyword>\<open>typedecl\<close> "type declaration" | 
| 62856 | 356 | (Parse.type_args -- Parse.binding -- Parse.opt_mixfix | 
| 357 | >> (fn ((args, a), mx) => | |
| 358 |           Typedecl.typedecl {final = true} (a, map (rpair dummyS) args, mx) #> snd));
 | |
| 359 | ||
| 360 | val _ = | |
| 67147 | 361 | Outer_Syntax.local_theory \<^command_keyword>\<open>type_synonym\<close> "declare type abbreviation" | 
| 62856 | 362 | (Parse.type_args -- Parse.binding -- | 
| 67147 | 363 | (\<^keyword>\<open>=\<close> |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix')) | 
| 62856 | 364 | >> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs)); | 
| 365 | ||
| 366 | in end\<close> | |
| 367 | ||
| 368 | ||
| 369 | subsubsection \<open>Consts\<close> | |
| 370 | ||
| 371 | ML \<open> | |
| 372 | local | |
| 373 | ||
| 374 | val _ = | |
| 67147 | 375 | Outer_Syntax.command \<^command_keyword>\<open>judgment\<close> "declare object-logic judgment" | 
| 62856 | 376 | (Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd)); | 
| 377 | ||
| 378 | val _ = | |
| 67147 | 379 | Outer_Syntax.command \<^command_keyword>\<open>consts\<close> "declare constants" | 
| 62856 | 380 | (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts_cmd)); | 
| 381 | ||
| 382 | in end\<close> | |
| 383 | ||
| 384 | ||
| 385 | subsubsection \<open>Syntax and translations\<close> | |
| 386 | ||
| 387 | ML \<open> | |
| 388 | local | |
| 389 | ||
| 390 | val _ = | |
| 67147 | 391 | Outer_Syntax.command \<^command_keyword>\<open>nonterminal\<close> | 
| 62856 | 392 | "declare syntactic type constructors (grammar nonterminal symbols)" | 
| 393 | (Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global)); | |
| 394 | ||
| 395 | val _ = | |
| 74333 | 396 | Outer_Syntax.local_theory \<^command_keyword>\<open>syntax\<close> "add raw syntax clauses" | 
| 62856 | 397 | (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl | 
| 74333 | 398 | >> uncurry (Local_Theory.syntax_cmd true)); | 
| 62856 | 399 | |
| 400 | val _ = | |
| 74333 | 401 | Outer_Syntax.local_theory \<^command_keyword>\<open>no_syntax\<close> "delete raw syntax clauses" | 
| 62856 | 402 | (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl | 
| 74333 | 403 | >> uncurry (Local_Theory.syntax_cmd false)); | 
| 62856 | 404 | |
| 80750 | 405 | val syntax_consts = | 
| 406 | Parse.and_list1 | |
| 81170 
2d73c3287bd3
backout somewhat pointless 5ea48342e0ae: no need to declare syntax consts for translations (e.g. constraints);
 wenzelm parents: 
81148diff
changeset | 407 | ((Scan.repeat1 Parse.name_position --| (\<^keyword>\<open>\<rightleftharpoons>\<close> || \<^keyword>\<open>==\<close>)) -- | 
| 80750 | 408 | Parse.!!! (Scan.repeat1 Parse.name_position)); | 
| 409 | ||
| 410 | val _ = | |
| 411 | Outer_Syntax.command \<^command_keyword>\<open>syntax_consts\<close> "declare syntax const dependencies" | |
| 412 | (syntax_consts >> (Toplevel.theory o Isar_Cmd.syntax_consts)); | |
| 413 | ||
| 414 | val _ = | |
| 415 | Outer_Syntax.command \<^command_keyword>\<open>syntax_types\<close> "declare syntax const dependencies (type names)" | |
| 416 | (syntax_consts >> (Toplevel.theory o Isar_Cmd.syntax_types)); | |
| 417 | ||
| 62856 | 418 | val trans_pat = | 
| 419 | Scan.optional | |
| 67147 | 420 | (\<^keyword>\<open>(\<close> |-- Parse.!!! (Parse.inner_syntax Parse.name --| \<^keyword>\<open>)\<close>)) "logic" | 
| 62856 | 421 | -- Parse.inner_syntax Parse.string; | 
| 422 | ||
| 423 | fun trans_arrow toks = | |
| 67147 | 424 | ((\<^keyword>\<open>\<rightharpoonup>\<close> || \<^keyword>\<open>=>\<close>) >> K Syntax.Parse_Rule || | 
| 425 | (\<^keyword>\<open>\<leftharpoondown>\<close> || \<^keyword>\<open><=\<close>) >> K Syntax.Print_Rule || | |
| 426 | (\<^keyword>\<open>\<rightleftharpoons>\<close> || \<^keyword>\<open>==\<close>) >> K Syntax.Parse_Print_Rule) toks; | |
| 62856 | 427 | |
| 428 | val trans_line = | |
| 429 | trans_pat -- Parse.!!! (trans_arrow -- trans_pat) | |
| 430 | >> (fn (left, (arr, right)) => arr (left, right)); | |
| 431 | ||
| 432 | val _ = | |
| 67147 | 433 | Outer_Syntax.command \<^command_keyword>\<open>translations\<close> "add syntax translation rules" | 
| 62856 | 434 | (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.translations)); | 
| 435 | ||
| 436 | val _ = | |
| 67147 | 437 | Outer_Syntax.command \<^command_keyword>\<open>no_translations\<close> "delete syntax translation rules" | 
| 62856 | 438 | (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.no_translations)); | 
| 439 | ||
| 440 | in end\<close> | |
| 441 | ||
| 442 | ||
| 443 | subsubsection \<open>Translation functions\<close> | |
| 444 | ||
| 445 | ML \<open> | |
| 446 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 447 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 448 | val _ = | 
| 67147 | 449 | Outer_Syntax.command \<^command_keyword>\<open>parse_ast_translation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 450 | "install parse ast translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 451 | (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_ast_translation)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 452 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 453 | val _ = | 
| 67147 | 454 | Outer_Syntax.command \<^command_keyword>\<open>parse_translation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 455 | "install parse translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 456 | (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_translation)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 457 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 458 | val _ = | 
| 67147 | 459 | Outer_Syntax.command \<^command_keyword>\<open>print_translation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 460 | "install print translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 461 | (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_translation)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 462 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 463 | val _ = | 
| 67147 | 464 | Outer_Syntax.command \<^command_keyword>\<open>typed_print_translation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 465 | "install typed print translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 466 | (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.typed_print_translation)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 467 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 468 | val _ = | 
| 67147 | 469 | Outer_Syntax.command \<^command_keyword>\<open>print_ast_translation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 470 | "install print ast translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 471 | (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_ast_translation)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 472 | |
| 62856 | 473 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 474 | |
| 62856 | 475 | |
| 476 | subsubsection \<open>Specifications\<close> | |
| 477 | ||
| 478 | ML \<open> | |
| 479 | local | |
| 480 | ||
| 481 | val _ = | |
| 67147 | 482 | Outer_Syntax.local_theory' \<^command_keyword>\<open>definition\<close> "constant definition" | 
| 63180 | 483 | (Scan.option Parse_Spec.constdecl -- (Parse_Spec.opt_thm_name ":" -- Parse.prop) -- | 
| 484 | Parse_Spec.if_assumes -- Parse.for_fixes >> (fn (((decl, spec), prems), params) => | |
| 485 | #2 oo Specification.definition_cmd decl params prems spec)); | |
| 62856 | 486 | |
| 487 | val _ = | |
| 67147 | 488 | Outer_Syntax.local_theory' \<^command_keyword>\<open>abbreviation\<close> "constant abbreviation" | 
| 63180 | 489 | (Parse.syntax_mode -- Scan.option Parse_Spec.constdecl -- Parse.prop -- Parse.for_fixes | 
| 490 | >> (fn (((mode, decl), spec), params) => Specification.abbreviation_cmd mode decl params spec)); | |
| 62856 | 491 | |
| 63178 | 492 | val axiomatization = | 
| 493 | Parse.and_list1 (Parse_Spec.thm_name ":" -- Parse.prop) -- | |
| 494 | Parse_Spec.if_assumes -- Parse.for_fixes >> (fn ((a, b), c) => (c, b, a)); | |
| 495 | ||
| 62856 | 496 | val _ = | 
| 67147 | 497 | Outer_Syntax.command \<^command_keyword>\<open>axiomatization\<close> "axiomatic constant specification" | 
| 63285 | 498 | (Scan.optional Parse.vars [] -- | 
| 63178 | 499 | Scan.optional (Parse.where_ |-- Parse.!!! axiomatization) ([], [], []) | 
| 500 | >> (fn (a, (b, c, d)) => Toplevel.theory (#2 o Specification.axiomatization_cmd a b c d))); | |
| 62856 | 501 | |
| 66248 | 502 | val _ = | 
| 67147 | 503 | Outer_Syntax.local_theory \<^command_keyword>\<open>alias\<close> "name-space alias for constant" | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 504 | (Parse.binding -- (Parse.!!! \<^keyword>\<open>=\<close> |-- Parse.name_position) | 
| 66248 | 505 | >> Specification.alias_cmd); | 
| 506 | ||
| 507 | val _ = | |
| 67147 | 508 | Outer_Syntax.local_theory \<^command_keyword>\<open>type_alias\<close> "name-space alias for type constructor" | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 509 | (Parse.binding -- (Parse.!!! \<^keyword>\<open>=\<close> |-- Parse.name_position) | 
| 66248 | 510 | >> Specification.type_alias_cmd); | 
| 511 | ||
| 62856 | 512 | in end\<close> | 
| 513 | ||
| 514 | ||
| 515 | subsubsection \<open>Notation\<close> | |
| 516 | ||
| 517 | ML \<open> | |
| 518 | local | |
| 519 | ||
| 520 | val _ = | |
| 67147 | 521 | Outer_Syntax.local_theory \<^command_keyword>\<open>type_notation\<close> | 
| 62856 | 522 | "add concrete syntax for type constructors" | 
| 523 | (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) | |
| 74338 | 524 | >> (fn (mode, args) => Local_Theory.type_notation_cmd true mode args)); | 
| 62856 | 525 | |
| 526 | val _ = | |
| 67147 | 527 | Outer_Syntax.local_theory \<^command_keyword>\<open>no_type_notation\<close> | 
| 62856 | 528 | "delete concrete syntax for type constructors" | 
| 529 | (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) | |
| 74338 | 530 | >> (fn (mode, args) => Local_Theory.type_notation_cmd false mode args)); | 
| 62856 | 531 | |
| 532 | val _ = | |
| 67147 | 533 | Outer_Syntax.local_theory \<^command_keyword>\<open>notation\<close> | 
| 62856 | 534 | "add concrete syntax for constants / fixed variables" | 
| 535 | (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) | |
| 74338 | 536 | >> (fn (mode, args) => Local_Theory.notation_cmd true mode args)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 537 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 538 | val _ = | 
| 67147 | 539 | Outer_Syntax.local_theory \<^command_keyword>\<open>no_notation\<close> | 
| 62856 | 540 | "delete concrete syntax for constants / fixed variables" | 
| 541 | (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) | |
| 74338 | 542 | >> (fn (mode, args) => Local_Theory.notation_cmd false mode args)); | 
| 62856 | 543 | |
| 544 | in end\<close> | |
| 545 | ||
| 546 | ||
| 547 | subsubsection \<open>Theorems\<close> | |
| 548 | ||
| 549 | ML \<open> | |
| 550 | local | |
| 551 | ||
| 63094 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 552 | val long_keyword = | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 553 | Parse_Spec.includes >> K "" || | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 554 | Parse_Spec.long_statement_keyword; | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 555 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 556 | val long_statement = | 
| 63352 | 557 | Scan.optional (Parse_Spec.opt_thm_name ":" --| Scan.ahead long_keyword) Binding.empty_atts -- | 
| 63094 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 558 | Scan.optional Parse_Spec.includes [] -- Parse_Spec.long_statement | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 559 | >> (fn ((binding, includes), (elems, concl)) => (true, binding, includes, elems, concl)); | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 560 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 561 | val short_statement = | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 562 | Parse_Spec.statement -- Parse_Spec.if_statement -- Parse.for_fixes | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 563 | >> (fn ((shows, assumes), fixes) => | 
| 63352 | 564 | (false, Binding.empty_atts, [], [Element.Fixes fixes, Element.Assumes assumes], | 
| 63094 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 565 | Element.Shows shows)); | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 566 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 567 | fun theorem spec schematic descr = | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 568 |   Outer_Syntax.local_theory_to_proof' spec ("state " ^ descr)
 | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 569 | ((long_statement || short_statement) >> (fn (long, binding, includes, elems, concl) => | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 570 | ((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd) | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 571 | long Thm.theoremK NONE (K I) binding includes elems concl))); | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 572 | |
| 67147 | 573 | val _ = theorem \<^command_keyword>\<open>theorem\<close> false "theorem"; | 
| 574 | val _ = theorem \<^command_keyword>\<open>lemma\<close> false "lemma"; | |
| 575 | val _ = theorem \<^command_keyword>\<open>corollary\<close> false "corollary"; | |
| 576 | val _ = theorem \<^command_keyword>\<open>proposition\<close> false "proposition"; | |
| 577 | val _ = theorem \<^command_keyword>\<open>schematic_goal\<close> true "schematic goal"; | |
| 63094 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 578 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 579 | in end\<close> | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 580 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 581 | ML \<open> | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 582 | local | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 583 | |
| 62856 | 584 | val _ = | 
| 67147 | 585 | Outer_Syntax.local_theory' \<^command_keyword>\<open>lemmas\<close> "define theorems" | 
| 62856 | 586 | (Parse_Spec.name_facts -- Parse.for_fixes >> | 
| 587 | (fn (facts, fixes) => #2 oo Specification.theorems_cmd Thm.theoremK facts fixes)); | |
| 588 | ||
| 589 | val _ = | |
| 67147 | 590 | Outer_Syntax.local_theory' \<^command_keyword>\<open>declare\<close> "declare theorems" | 
| 62969 | 591 | (Parse.and_list1 Parse.thms1 -- Parse.for_fixes | 
| 62856 | 592 | >> (fn (facts, fixes) => | 
| 63352 | 593 | #2 oo Specification.theorems_cmd "" [(Binding.empty_atts, flat facts)] fixes)); | 
| 62856 | 594 | |
| 595 | val _ = | |
| 67147 | 596 | Outer_Syntax.local_theory \<^command_keyword>\<open>named_theorems\<close> | 
| 62856 | 597 | "declare named collection of theorems" | 
| 74887 | 598 | (Parse.and_list1 (Parse.binding -- Scan.optional Parse.embedded "") >> | 
| 70182 
ca9dfa7ee3bd
backed out experimental b67bab2b132c, which slipped in accidentally
 haftmann parents: 
70177diff
changeset | 599 | fold (fn (b, descr) => snd o Named_Theorems.declare b descr)); | 
| 62856 | 600 | |
| 601 | in end\<close> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 602 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 603 | |
| 62856 | 604 | subsubsection \<open>Hide names\<close> | 
| 605 | ||
| 606 | ML \<open> | |
| 607 | local | |
| 608 | ||
| 609 | fun hide_names command_keyword what hide parse prep = | |
| 610 |   Outer_Syntax.command command_keyword ("hide " ^ what ^ " from name space")
 | |
| 611 | ((Parse.opt_keyword "open" >> not) -- Scan.repeat1 parse >> (fn (fully, args) => | |
| 612 | (Toplevel.theory (fn thy => | |
| 613 | let val ctxt = Proof_Context.init_global thy | |
| 614 | in fold (hide fully o prep ctxt) args thy end)))); | |
| 615 | ||
| 616 | val _ = | |
| 67147 | 617 | hide_names \<^command_keyword>\<open>hide_class\<close> "classes" Sign.hide_class Parse.class | 
| 62856 | 618 | Proof_Context.read_class; | 
| 619 | ||
| 620 | val _ = | |
| 67147 | 621 | hide_names \<^command_keyword>\<open>hide_type\<close> "types" Sign.hide_type Parse.type_const | 
| 80632 | 622 |     (dest_Type_name oo Proof_Context.read_type_name {proper = true, strict = false});
 | 
| 62856 | 623 | |
| 624 | val _ = | |
| 67147 | 625 | hide_names \<^command_keyword>\<open>hide_const\<close> "consts" Sign.hide_const Parse.const | 
| 80635 | 626 |     (dest_Const_name oo Proof_Context.read_const {proper = true, strict = false});
 | 
| 62856 | 627 | |
| 628 | val _ = | |
| 67147 | 629 | hide_names \<^command_keyword>\<open>hide_fact\<close> "facts" Global_Theory.hide_fact | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 630 | Parse.name_position (Global_Theory.check_fact o Proof_Context.theory_of); | 
| 62856 | 631 | |
| 632 | in end\<close> | |
| 633 | ||
| 634 | ||
| 635 | subsection \<open>Bundled declarations\<close> | |
| 636 | ||
| 637 | ML \<open> | |
| 638 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 639 | |
| 81106 | 640 | fun bundle_cmd open_bundle = | 
| 641 | (Parse.binding --| \<^keyword>\<open>=\<close>) -- Parse.thms1 -- Parse.for_fixes | |
| 642 |     >> (uncurry (Bundle.bundle_cmd {open_bundle = open_bundle}));
 | |
| 643 | ||
| 644 | fun bundle_begin open_bundle = | |
| 645 |   Parse.binding --| Parse.begin >> Bundle.init {open_bundle = open_bundle};
 | |
| 646 | ||
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 647 | val _ = | 
| 67147 | 648 | Outer_Syntax.maybe_begin_local_theory \<^command_keyword>\<open>bundle\<close> | 
| 81106 | 649 | "define bundle of declarations" (bundle_cmd false) (bundle_begin false); | 
| 650 | ||
| 651 | val _ = | |
| 652 | Outer_Syntax.maybe_begin_local_theory \<^command_keyword>\<open>open_bundle\<close> | |
| 653 | "define and open bundle of declarations" (bundle_cmd true) (bundle_begin true); | |
| 63270 | 654 | |
| 655 | val _ = | |
| 67147 | 656 | Outer_Syntax.local_theory \<^command_keyword>\<open>unbundle\<close> | 
| 81113 | 657 | "open bundle in local theory" (Parse_Spec.bundles >> Bundle.unbundle_cmd); | 
| 63282 | 658 | |
| 659 | val _ = | |
| 67147 | 660 | Outer_Syntax.command \<^command_keyword>\<open>include\<close> | 
| 81113 | 661 | "open bundle in proof body" (Parse_Spec.bundles >> (Toplevel.proof o Bundle.include_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 662 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 663 | val _ = | 
| 67147 | 664 | Outer_Syntax.command \<^command_keyword>\<open>including\<close> | 
| 81113 | 665 | "open bundle in goal refinement" (Parse_Spec.bundles >> (Toplevel.proof o Bundle.including_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 666 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 667 | val _ = | 
| 67147 | 668 | Outer_Syntax.command \<^command_keyword>\<open>print_bundles\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 669 | "print bundles of declarations" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 670 | (Parse.opt_bang >> (fn b => Toplevel.keep (Bundle.print_bundles b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 671 | |
| 62856 | 672 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 673 | |
| 62856 | 674 | |
| 675 | subsection \<open>Local theory specifications\<close> | |
| 676 | ||
| 677 | subsubsection \<open>Specification context\<close> | |
| 678 | ||
| 679 | ML \<open> | |
| 680 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 681 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 682 | val _ = | 
| 67147 | 683 | Outer_Syntax.command \<^command_keyword>\<open>context\<close> "begin local theory context" | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 684 | (((Parse.name_position -- Scan.optional Parse_Spec.opening []) | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 685 | >> (fn (name, incls) => Toplevel.begin_main_target true (Target_Context.context_begin_named_cmd incls name)) || | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 686 | Scan.optional Parse_Spec.includes [] -- Scan.repeat Parse_Spec.context_element | 
| 72453 | 687 | >> (fn (incls, elems) => Toplevel.begin_nested_target (Target_Context.context_begin_nested_cmd incls elems))) | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 688 | --| Parse.begin); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 689 | |
| 62856 | 690 | val _ = | 
| 67147 | 691 | Outer_Syntax.command \<^command_keyword>\<open>end\<close> "end context" | 
| 62856 | 692 | (Scan.succeed | 
| 72434 | 693 | (Toplevel.exit o Toplevel.end_main_target o Toplevel.end_nested_target o | 
| 62856 | 694 | Toplevel.end_proof (K Proof.end_notepad))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 695 | |
| 62856 | 696 | in end\<close> | 
| 697 | ||
| 698 | ||
| 699 | subsubsection \<open>Locales and interpretation\<close> | |
| 700 | ||
| 701 | ML \<open> | |
| 702 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 703 | |
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
72453diff
changeset | 704 | val locale_context_elements = | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 705 | Scan.repeat1 Parse_Spec.context_element; | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
72453diff
changeset | 706 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 707 | val locale_val = | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 708 | ((Parse_Spec.locale_expression -- Scan.optional Parse_Spec.opening []) | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 709 | || Parse_Spec.opening >> pair ([], [])) | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 710 | -- Scan.optional (\<^keyword>\<open>+\<close> |-- Parse.!!! locale_context_elements) [] | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 711 | || locale_context_elements >> pair (([], []), []); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 712 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 713 | val _ = | 
| 67147 | 714 | Outer_Syntax.command \<^command_keyword>\<open>locale\<close> "define named specification context" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 715 | (Parse.binding -- | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 716 | Scan.optional (\<^keyword>\<open>=\<close> |-- Parse.!!! locale_val) ((([], []), []), []) -- Parse.opt_begin | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 717 | >> (fn ((name, ((expr, includes), elems)), begin) => | 
| 72434 | 718 | Toplevel.begin_main_target begin | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
72453diff
changeset | 719 | (Expression.add_locale_cmd name Binding.empty includes expr elems #> snd))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 720 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 721 | val _ = | 
| 67147 | 722 | Outer_Syntax.command \<^command_keyword>\<open>experiment\<close> "open private specification context" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 723 | (Scan.repeat Parse_Spec.context_element --| Parse.begin | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 724 | >> (fn elems => | 
| 72434 | 725 | Toplevel.begin_main_target true (Experiment.experiment_cmd elems #> snd))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 726 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 727 | val _ = | 
| 67147 | 728 | Outer_Syntax.command \<^command_keyword>\<open>interpret\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 729 | "prove interpretation of locale expression in proof context" | 
| 67740 | 730 | (Parse.!!! Parse_Spec.locale_expression >> (fn expr => | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 731 | Toplevel.proof (Interpretation.interpret_cmd expr))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 732 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 733 | val interpretation_args_with_defs = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 734 | Parse.!!! Parse_Spec.locale_expression -- | 
| 67147 | 735 | (Scan.optional (\<^keyword>\<open>defines\<close> |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 736 | -- ((Parse.binding -- Parse.opt_mixfix') --| \<^keyword>\<open>=\<close> -- Parse.term))) ([])); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 737 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 738 | val _ = | 
| 67147 | 739 | Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>global_interpretation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 740 | "prove interpretation of locale expression into global theory" | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 741 | (interpretation_args_with_defs >> (fn (expr, defs) => | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 742 | Interpretation.global_interpretation_cmd expr defs)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 743 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 744 | val _ = | 
| 67147 | 745 | Outer_Syntax.command \<^command_keyword>\<open>sublocale\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 746 | "prove sublocale relation between a locale and a locale expression" | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 747 | ((Parse.name_position --| (\<^keyword>\<open>\<subseteq>\<close> || \<^keyword>\<open><\<close>) -- | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 748 | interpretation_args_with_defs >> (fn (loc, (expr, defs)) => | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 749 | Toplevel.theory_to_proof (Interpretation.global_sublocale_cmd loc expr defs))) | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 750 | || interpretation_args_with_defs >> (fn (expr, defs) => | 
| 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 751 | Toplevel.local_theory_to_proof NONE NONE (Interpretation.sublocale_cmd expr defs))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 752 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 753 | val _ = | 
| 67147 | 754 | Outer_Syntax.command \<^command_keyword>\<open>interpretation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 755 | "prove interpretation of locale expression in local theory or into global theory" | 
| 67740 | 756 | (Parse.!!! Parse_Spec.locale_expression >> (fn expr => | 
| 62856 | 757 | Toplevel.local_theory_to_proof NONE NONE | 
| 67777 
2d3c1091527b
Drop rewrite rule arguments of sublocale and interpretation implementations.
 ballarin parents: 
67764diff
changeset | 758 | (Interpretation.isar_interpretation_cmd expr))); | 
| 62856 | 759 | |
| 760 | in end\<close> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 761 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 762 | |
| 62856 | 763 | subsubsection \<open>Type classes\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 764 | |
| 62856 | 765 | ML \<open> | 
| 766 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 767 | |
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
72453diff
changeset | 768 | val class_context_elements = | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 769 | Scan.repeat1 Parse_Spec.context_element; | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
72453diff
changeset | 770 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 771 | val class_val = | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 772 | ((Parse_Spec.class_expression -- Scan.optional Parse_Spec.opening []) | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 773 | || Parse_Spec.opening >> pair []) | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 774 | -- Scan.optional (\<^keyword>\<open>+\<close> |-- Parse.!!! class_context_elements) [] || | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 775 | class_context_elements >> pair ([], []); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 776 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 777 | val _ = | 
| 67147 | 778 | Outer_Syntax.command \<^command_keyword>\<open>class\<close> "define type class" | 
| 72739 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 779 | (Parse.binding -- Scan.optional (\<^keyword>\<open>=\<close> |-- class_val) (([], []), []) -- Parse.opt_begin | 
| 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 haftmann parents: 
72536diff
changeset | 780 | >> (fn ((name, ((supclasses, includes), elems)), begin) => | 
| 72434 | 781 | Toplevel.begin_main_target begin | 
| 72536 
589645894305
bundle mixins for locale and class specifications
 haftmann parents: 
72453diff
changeset | 782 | (Class_Declaration.class_cmd name includes supclasses elems #> snd))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 783 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 784 | val _ = | 
| 67147 | 785 | Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>subclass\<close> "prove a subclass relation" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 786 | (Parse.class >> Class_Declaration.subclass_cmd); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 787 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 788 | val _ = | 
| 67147 | 789 | Outer_Syntax.command \<^command_keyword>\<open>instantiation\<close> "instantiate and prove type arity" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 790 | (Parse.multi_arity --| Parse.begin | 
| 72434 | 791 | >> (fn arities => Toplevel.begin_main_target true (Class.instantiation_cmd arities))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 792 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 793 | val _ = | 
| 67147 | 794 | Outer_Syntax.command \<^command_keyword>\<open>instance\<close> "prove type arity or subclass relation" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 795 | ((Parse.class -- | 
| 67147 | 796 | ((\<^keyword>\<open>\<subseteq>\<close> || \<^keyword>\<open><\<close>) |-- Parse.!!! Parse.class) >> Class.classrel_cmd || | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 797 | Parse.multi_arity >> Class.instance_arity_cmd) >> Toplevel.theory_to_proof || | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 798 | Scan.succeed (Toplevel.local_theory_to_proof NONE NONE (Class.instantiation_instance I))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 799 | |
| 62856 | 800 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 801 | |
| 62856 | 802 | |
| 803 | subsubsection \<open>Arbitrary overloading\<close> | |
| 804 | ||
| 805 | ML \<open> | |
| 806 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 807 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 808 | val _ = | 
| 67147 | 809 | Outer_Syntax.command \<^command_keyword>\<open>overloading\<close> "overloaded definitions" | 
| 810 | (Scan.repeat1 (Parse.name --| (\<^keyword>\<open>==\<close> || \<^keyword>\<open>\<equiv>\<close>) -- Parse.term -- | |
| 811 | Scan.optional (\<^keyword>\<open>(\<close> |-- (\<^keyword>\<open>unchecked\<close> >> K false) --| \<^keyword>\<open>)\<close>) true | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 812 | >> Scan.triple1) --| Parse.begin | 
| 72434 | 813 | >> (fn operations => Toplevel.begin_main_target true (Overloading.overloading_cmd operations))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 814 | |
| 62856 | 815 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 816 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 817 | |
| 62856 | 818 | subsection \<open>Proof commands\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 819 | |
| 62856 | 820 | ML \<open> | 
| 821 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 822 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 823 | val _ = | 
| 67147 | 824 | Outer_Syntax.local_theory_to_proof \<^command_keyword>\<open>notepad\<close> "begin proof context" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 825 | (Parse.begin >> K Proof.begin_notepad); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 826 | |
| 62856 | 827 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 828 | |
| 62856 | 829 | |
| 830 | subsubsection \<open>Statements\<close> | |
| 831 | ||
| 832 | ML \<open> | |
| 833 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 834 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 835 | val structured_statement = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 836 | Parse_Spec.statement -- Parse_Spec.cond_statement -- Parse.for_fixes | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 837 | >> (fn ((shows, (strict, assumes)), fixes) => (strict, fixes, assumes, shows)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 838 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 839 | val _ = | 
| 67147 | 840 | Outer_Syntax.command \<^command_keyword>\<open>have\<close> "state local goal" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 841 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 842 | Toplevel.proof' (fn int => Proof.have_cmd a NONE (K I) b c d int #> #2))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 843 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 844 | val _ = | 
| 67147 | 845 | Outer_Syntax.command \<^command_keyword>\<open>show\<close> "state local goal, to refine pending subgoals" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 846 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 847 | Toplevel.proof' (fn int => Proof.show_cmd a NONE (K I) b c d int #> #2))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 848 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 849 | val _ = | 
| 67147 | 850 | Outer_Syntax.command \<^command_keyword>\<open>hence\<close> "old-style alias of \"then have\"" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 851 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 852 | Toplevel.proof' (fn int => Proof.chain #> Proof.have_cmd a NONE (K I) b c d int #> #2))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 853 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 854 | val _ = | 
| 67147 | 855 | Outer_Syntax.command \<^command_keyword>\<open>thus\<close> "old-style alias of \"then show\"" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 856 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 857 | Toplevel.proof' (fn int => Proof.chain #> Proof.show_cmd a NONE (K I) b c d int #> #2))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 858 | |
| 62856 | 859 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 860 | |
| 62856 | 861 | |
| 862 | subsubsection \<open>Local facts\<close> | |
| 863 | ||
| 864 | ML \<open> | |
| 865 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 866 | |
| 62969 | 867 | val facts = Parse.and_list1 Parse.thms1; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 868 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 869 | val _ = | 
| 67147 | 870 | Outer_Syntax.command \<^command_keyword>\<open>then\<close> "forward chaining" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 871 | (Scan.succeed (Toplevel.proof Proof.chain)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 872 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 873 | val _ = | 
| 67147 | 874 | Outer_Syntax.command \<^command_keyword>\<open>from\<close> "forward chaining from given facts" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 875 | (facts >> (Toplevel.proof o Proof.from_thmss_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 876 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 877 | val _ = | 
| 67147 | 878 | Outer_Syntax.command \<^command_keyword>\<open>with\<close> "forward chaining from given and current facts" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 879 | (facts >> (Toplevel.proof o Proof.with_thmss_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 880 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 881 | val _ = | 
| 67147 | 882 | Outer_Syntax.command \<^command_keyword>\<open>note\<close> "define facts" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 883 | (Parse_Spec.name_facts >> (Toplevel.proof o Proof.note_thmss_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 884 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 885 | val _ = | 
| 67147 | 886 | Outer_Syntax.command \<^command_keyword>\<open>supply\<close> "define facts during goal refinement (unstructured)" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 887 | (Parse_Spec.name_facts >> (Toplevel.proof o Proof.supply_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 888 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 889 | val _ = | 
| 67147 | 890 | Outer_Syntax.command \<^command_keyword>\<open>using\<close> "augment goal facts" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 891 | (facts >> (Toplevel.proof o Proof.using_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 892 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 893 | val _ = | 
| 67147 | 894 | Outer_Syntax.command \<^command_keyword>\<open>unfolding\<close> "unfold definitions in goal and facts" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 895 | (facts >> (Toplevel.proof o Proof.unfolding_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 896 | |
| 62856 | 897 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 898 | |
| 62856 | 899 | |
| 900 | subsubsection \<open>Proof context\<close> | |
| 901 | ||
| 902 | ML \<open> | |
| 903 | local | |
| 904 | ||
| 905 | val structured_statement = | |
| 906 | Parse_Spec.statement -- Parse_Spec.if_statement' -- Parse.for_fixes | |
| 907 | >> (fn ((shows, assumes), fixes) => (fixes, assumes, shows)); | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 908 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 909 | val _ = | 
| 67147 | 910 | Outer_Syntax.command \<^command_keyword>\<open>fix\<close> "fix local variables (Skolem constants)" | 
| 63285 | 911 | (Parse.vars >> (Toplevel.proof o Proof.fix_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 912 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 913 | val _ = | 
| 67147 | 914 | Outer_Syntax.command \<^command_keyword>\<open>assume\<close> "assume propositions" | 
| 62856 | 915 | (structured_statement >> (fn (a, b, c) => Toplevel.proof (Proof.assume_cmd a b c))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 916 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 917 | val _ = | 
| 67147 | 918 | Outer_Syntax.command \<^command_keyword>\<open>presume\<close> "assume propositions, to be established later" | 
| 62856 | 919 | (structured_statement >> (fn (a, b, c) => Toplevel.proof (Proof.presume_cmd a b c))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 920 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 921 | val _ = | 
| 67147 | 922 | Outer_Syntax.command \<^command_keyword>\<open>define\<close> "local definition (non-polymorphic)" | 
| 63285 | 923 | ((Parse.vars --| Parse.where_) -- Parse_Spec.statement -- Parse.for_fixes | 
| 63039 | 924 | >> (fn ((a, b), c) => Toplevel.proof (Proof.define_cmd a c b))); | 
| 925 | ||
| 926 | val _ = | |
| 67147 | 927 | Outer_Syntax.command \<^command_keyword>\<open>consider\<close> "state cases rule" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 928 | (Parse_Spec.obtains >> (Toplevel.proof' o Obtain.consider_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 929 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 930 | val _ = | 
| 67147 | 931 | Outer_Syntax.command \<^command_keyword>\<open>obtain\<close> "generalized elimination" | 
| 63285 | 932 | (Parse.parbinding -- Scan.optional (Parse.vars --| Parse.where_) [] -- structured_statement | 
| 63059 
3f577308551e
'obtain' supports structured statements (similar to 'define');
 wenzelm parents: 
63043diff
changeset | 933 | >> (fn ((a, b), (c, d, e)) => Toplevel.proof' (Obtain.obtain_cmd a b c d e))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 934 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 935 | val _ = | 
| 67147 | 936 | Outer_Syntax.command \<^command_keyword>\<open>let\<close> "bind text variables" | 
| 937 | (Parse.and_list1 (Parse.and_list1 Parse.term -- (\<^keyword>\<open>=\<close> |-- Parse.term)) | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 938 | >> (Toplevel.proof o Proof.let_bind_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 939 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 940 | val _ = | 
| 67147 | 941 | Outer_Syntax.command \<^command_keyword>\<open>write\<close> "add concrete syntax for constants / fixed variables" | 
| 62856 | 942 | (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 943 | >> (fn (mode, args) => Toplevel.proof (Proof.write_cmd mode args))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 944 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 945 | val _ = | 
| 67147 | 946 | Outer_Syntax.command \<^command_keyword>\<open>case\<close> "invoke local context" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 947 | (Parse_Spec.opt_thm_name ":" -- | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 948 | (\<^keyword>\<open>(\<close> |-- Parse.!!! (Parse.name_position -- Scan.repeat (Parse.maybe Parse.binding) | 
| 67147 | 949 | --| \<^keyword>\<open>)\<close>) || | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 950 | Parse.name_position >> rpair []) >> (Toplevel.proof o Proof.case_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 951 | |
| 62856 | 952 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 953 | |
| 62856 | 954 | |
| 955 | subsubsection \<open>Proof structure\<close> | |
| 956 | ||
| 957 | ML \<open> | |
| 958 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 959 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 960 | val _ = | 
| 67147 | 961 |   Outer_Syntax.command \<^command_keyword>\<open>{\<close> "begin explicit proof block"
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 962 | (Scan.succeed (Toplevel.proof Proof.begin_block)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 963 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 964 | val _ = | 
| 67147 | 965 | Outer_Syntax.command \<^command_keyword>\<open>}\<close> "end explicit proof block" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 966 | (Scan.succeed (Toplevel.proof Proof.end_block)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 967 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 968 | val _ = | 
| 67147 | 969 | Outer_Syntax.command \<^command_keyword>\<open>next\<close> "enter next proof block" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 970 | (Scan.succeed (Toplevel.proof Proof.next_block)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 971 | |
| 62856 | 972 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 973 | |
| 62856 | 974 | |
| 975 | subsubsection \<open>End proof\<close> | |
| 976 | ||
| 977 | ML \<open> | |
| 978 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 979 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 980 | val _ = | 
| 67147 | 981 | Outer_Syntax.command \<^command_keyword>\<open>qed\<close> "conclude proof" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 982 | (Scan.option Method.parse >> (fn m => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 983 | (Option.map Method.report m; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 984 | Isar_Cmd.qed m))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 985 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 986 | val _ = | 
| 67147 | 987 | Outer_Syntax.command \<^command_keyword>\<open>by\<close> "terminal backward proof" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 988 | (Method.parse -- Scan.option Method.parse >> (fn (m1, m2) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 989 | (Method.report m1; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 990 | Option.map Method.report m2; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 991 | Isar_Cmd.terminal_proof (m1, m2)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 992 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 993 | val _ = | 
| 67147 | 994 | Outer_Syntax.command \<^command_keyword>\<open>..\<close> "default proof" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 995 | (Scan.succeed Isar_Cmd.default_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 996 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 997 | val _ = | 
| 67147 | 998 | Outer_Syntax.command \<^command_keyword>\<open>.\<close> "immediate proof" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 999 | (Scan.succeed Isar_Cmd.immediate_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1000 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1001 | val _ = | 
| 67147 | 1002 | Outer_Syntax.command \<^command_keyword>\<open>done\<close> "done proof" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1003 | (Scan.succeed Isar_Cmd.done_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1004 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1005 | val _ = | 
| 67147 | 1006 | Outer_Syntax.command \<^command_keyword>\<open>sorry\<close> "skip proof (quick-and-dirty mode only!)" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1007 | (Scan.succeed Isar_Cmd.skip_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1008 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1009 | val _ = | 
| 67147 | 1010 | Outer_Syntax.command \<^command_keyword>\<open>\<proof>\<close> "dummy proof (quick-and-dirty mode only!)" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1011 | (Scan.succeed Isar_Cmd.skip_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1012 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1013 | val _ = | 
| 67147 | 1014 | Outer_Syntax.command \<^command_keyword>\<open>oops\<close> "forget proof" | 
| 68876 
cefaac3d24ff
no reset_proof for notepad: begin/end structure takes precedence over goal/proof structure;
 wenzelm parents: 
68827diff
changeset | 1015 | (Scan.succeed Toplevel.forget_proof); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1016 | |
| 62856 | 1017 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1018 | |
| 62856 | 1019 | |
| 1020 | subsubsection \<open>Proof steps\<close> | |
| 1021 | ||
| 1022 | ML \<open> | |
| 1023 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1024 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1025 | val _ = | 
| 67147 | 1026 | Outer_Syntax.command \<^command_keyword>\<open>defer\<close> "shuffle internal proof state" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1027 | (Scan.optional Parse.nat 1 >> (Toplevel.proof o Proof.defer)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1028 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1029 | val _ = | 
| 67147 | 1030 | Outer_Syntax.command \<^command_keyword>\<open>prefer\<close> "shuffle internal proof state" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1031 | (Parse.nat >> (Toplevel.proof o Proof.prefer)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1032 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1033 | val _ = | 
| 67147 | 1034 | Outer_Syntax.command \<^command_keyword>\<open>apply\<close> "initial goal refinement step (unstructured)" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1035 | (Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply m)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1036 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1037 | val _ = | 
| 67147 | 1038 | Outer_Syntax.command \<^command_keyword>\<open>apply_end\<close> "terminal goal refinement step (unstructured)" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1039 | (Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply_end m)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1040 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1041 | val _ = | 
| 67147 | 1042 | Outer_Syntax.command \<^command_keyword>\<open>proof\<close> "backward proof step" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1043 | (Scan.option Method.parse >> (fn m => | 
| 63513 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1044 | (Option.map Method.report m; | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1045 | Toplevel.proof (fn state => | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1046 | let | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1047 | val state' = state |> Proof.proof m |> Seq.the_result ""; | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1048 | val _ = | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1049 | Output.information | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1050 | (Proof_Context.print_cases_proof (Proof.context_of state) (Proof.context_of state')); | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 1051 | in state' end)))) | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1052 | |
| 62856 | 1053 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1054 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1055 | |
| 62856 | 1056 | subsubsection \<open>Subgoal focus\<close> | 
| 1057 | ||
| 1058 | ML \<open> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1059 | local | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1060 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1061 | val opt_fact_binding = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1062 | Scan.optional (Parse.binding -- Parse.opt_attribs || Parse.attribs >> pair Binding.empty) | 
| 63352 | 1063 | Binding.empty_atts; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1064 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1065 | val for_params = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1066 | Scan.optional | 
| 67147 | 1067 | (\<^keyword>\<open>for\<close> |-- | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1068 | Parse.!!! ((Scan.option Parse.dots >> is_some) -- | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 1069 | (Scan.repeat1 (Parse.maybe_position Parse.name_position)))) | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1070 | (false, []); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1071 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1072 | val _ = | 
| 67147 | 1073 | Outer_Syntax.command \<^command_keyword>\<open>subgoal\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1074 | "focus on first subgoal within backward refinement" | 
| 67147 | 1075 | (opt_fact_binding -- (Scan.option (\<^keyword>\<open>premises\<close> |-- Parse.!!! opt_fact_binding)) -- | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1076 | for_params >> (fn ((a, b), c) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1077 | Toplevel.proofs (Seq.make_results o Seq.single o #2 o Subgoal.subgoal_cmd a b c))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1078 | |
| 62856 | 1079 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1080 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1081 | |
| 62856 | 1082 | subsubsection \<open>Calculation\<close> | 
| 1083 | ||
| 1084 | ML \<open> | |
| 1085 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1086 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1087 | val calculation_args = | 
| 67147 | 1088 | Scan.option (\<^keyword>\<open>(\<close> |-- Parse.!!! ((Parse.thms1 --| \<^keyword>\<open>)\<close>))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1089 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1090 | val _ = | 
| 67147 | 1091 | Outer_Syntax.command \<^command_keyword>\<open>also\<close> "combine calculation and current facts" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1092 | (calculation_args >> (Toplevel.proofs' o Calculation.also_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1093 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1094 | val _ = | 
| 67147 | 1095 | Outer_Syntax.command \<^command_keyword>\<open>finally\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1096 | "combine calculation and current facts, exhibit result" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1097 | (calculation_args >> (Toplevel.proofs' o Calculation.finally_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1098 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1099 | val _ = | 
| 67147 | 1100 | Outer_Syntax.command \<^command_keyword>\<open>moreover\<close> "augment calculation by current facts" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1101 | (Scan.succeed (Toplevel.proof' Calculation.moreover)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1102 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1103 | val _ = | 
| 67147 | 1104 | Outer_Syntax.command \<^command_keyword>\<open>ultimately\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1105 | "augment calculation by current facts, exhibit result" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1106 | (Scan.succeed (Toplevel.proof' Calculation.ultimately)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1107 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1108 | val _ = | 
| 67147 | 1109 | Outer_Syntax.command \<^command_keyword>\<open>print_trans_rules\<close> "print transitivity rules" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1110 | (Scan.succeed (Toplevel.keep (Calculation.print_rules o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1111 | |
| 62856 | 1112 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1113 | |
| 62856 | 1114 | |
| 1115 | subsubsection \<open>Proof navigation\<close> | |
| 1116 | ||
| 1117 | ML \<open> | |
| 1118 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1119 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1120 | fun report_back () = | 
| 64677 
8dc24130e8fe
more uniform treatment of "bad" like other messages (with serial number);
 wenzelm parents: 
64595diff
changeset | 1121 | Output.report [Markup.markup (Markup.bad ()) "Explicit backtracking"]; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1122 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1123 | val _ = | 
| 67147 | 1124 | Outer_Syntax.command \<^command_keyword>\<open>back\<close> "explicit backtracking of proof command" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1125 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1126 | (Toplevel.actual_proof (fn prf => (report_back (); Proof_Node.back prf)) o | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1127 | Toplevel.skip_proof report_back)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1128 | |
| 62856 | 1129 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1130 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1131 | |
| 62856 | 1132 | subsection \<open>Diagnostic commands (for interactive mode only)\<close> | 
| 1133 | ||
| 1134 | ML \<open> | |
| 1135 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1136 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1137 | val opt_modes = | 
| 67147 | 1138 | Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) []; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1139 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1140 | val _ = | 
| 67147 | 1141 | Outer_Syntax.command \<^command_keyword>\<open>help\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1142 | "retrieve outer syntax commands according to name patterns" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1143 | (Scan.repeat Parse.name >> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1144 | (fn pats => Toplevel.keep (fn st => Outer_Syntax.help (Toplevel.theory_of st) pats))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1145 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1146 | val _ = | 
| 67147 | 1147 | Outer_Syntax.command \<^command_keyword>\<open>print_commands\<close> "print outer syntax commands" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1148 | (Scan.succeed (Toplevel.keep (Outer_Syntax.print_commands o Toplevel.theory_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1149 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1150 | val _ = | 
| 67147 | 1151 | Outer_Syntax.command \<^command_keyword>\<open>print_options\<close> "print configuration options" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1152 | (Parse.opt_bang >> (fn b => Toplevel.keep (Attrib.print_options b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1153 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1154 | val _ = | 
| 67147 | 1155 | Outer_Syntax.command \<^command_keyword>\<open>print_context\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1156 | "print context of local theory target" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1157 | (Scan.succeed (Toplevel.keep (Pretty.writeln_chunks o Toplevel.pretty_context))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1158 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1159 | val _ = | 
| 67147 | 1160 | Outer_Syntax.command \<^command_keyword>\<open>print_theory\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1161 | "print logical theory contents" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1162 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1163 | Toplevel.keep (Pretty.writeln o Proof_Display.pretty_theory b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1164 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1165 | val _ = | 
| 67147 | 1166 | Outer_Syntax.command \<^command_keyword>\<open>print_definitions\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1167 | "print dependencies of definitional theory content" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1168 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1169 | Toplevel.keep (Pretty.writeln o Proof_Display.pretty_definitions b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1170 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1171 | val _ = | 
| 67147 | 1172 | Outer_Syntax.command \<^command_keyword>\<open>print_syntax\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1173 | "print inner syntax of context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1174 | (Scan.succeed (Toplevel.keep (Proof_Context.print_syntax o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1175 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1176 | val _ = | 
| 67147 | 1177 | Outer_Syntax.command \<^command_keyword>\<open>print_defn_rules\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1178 | "print definitional rewrite rules of context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1179 | (Scan.succeed (Toplevel.keep (Local_Defs.print_rules o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1180 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1181 | val _ = | 
| 67147 | 1182 | Outer_Syntax.command \<^command_keyword>\<open>print_abbrevs\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1183 | "print constant abbreviations of context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1184 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1185 | Toplevel.keep (Proof_Context.print_abbrevs b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1186 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1187 | val _ = | 
| 67147 | 1188 | Outer_Syntax.command \<^command_keyword>\<open>print_theorems\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1189 | "print theorems of local theory or proof context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1190 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1191 | Toplevel.keep (Pretty.writeln o Pretty.chunks o Isar_Cmd.pretty_theorems b))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1192 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1193 | val _ = | 
| 67147 | 1194 | Outer_Syntax.command \<^command_keyword>\<open>print_locales\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1195 | "print locales of this theory" | 
| 69059 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1196 | (Parse.opt_bang >> (fn verbose => | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1197 | Toplevel.keep (fn state => | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1198 | let val thy = Toplevel.theory_of state | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1199 | in Pretty.writeln (Locale.pretty_locales thy verbose) end))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1200 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1201 | val _ = | 
| 67147 | 1202 | Outer_Syntax.command \<^command_keyword>\<open>print_classes\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1203 | "print classes of this theory" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1204 | (Scan.succeed (Toplevel.keep (Class.print_classes o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1205 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1206 | val _ = | 
| 67147 | 1207 | Outer_Syntax.command \<^command_keyword>\<open>print_locale\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1208 | "print locale of this theory" | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 1209 | (Parse.opt_bang -- Parse.name_position >> (fn (show_facts, raw_name) => | 
| 69057 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
68876diff
changeset | 1210 | Toplevel.keep (fn state => | 
| 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
68876diff
changeset | 1211 | let | 
| 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
68876diff
changeset | 1212 | val thy = Toplevel.theory_of state; | 
| 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
68876diff
changeset | 1213 | val name = Locale.check thy raw_name; | 
| 
56c6378ebaea
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
68876diff
changeset | 1214 | in Pretty.writeln (Locale.pretty_locale thy show_facts name) end))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1215 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1216 | val _ = | 
| 67147 | 1217 | Outer_Syntax.command \<^command_keyword>\<open>print_interps\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1218 | "print interpretations of locale for this theory or proof context" | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 1219 | (Parse.name_position >> (fn raw_name => | 
| 69059 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1220 | Toplevel.keep (fn state => | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1221 | let | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1222 | val ctxt = Toplevel.context_of state; | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1223 | val thy = Toplevel.theory_of state; | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1224 | val name = Locale.check thy raw_name; | 
| 
70f9826753f6
tuned signature: prefer value-oriented pretty-printing;
 wenzelm parents: 
69057diff
changeset | 1225 | in Pretty.writeln (Locale.pretty_registrations ctxt name) end))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1226 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1227 | val _ = | 
| 67147 | 1228 | Outer_Syntax.command \<^command_keyword>\<open>print_attributes\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1229 | "print attributes of this theory" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1230 | (Parse.opt_bang >> (fn b => Toplevel.keep (Attrib.print_attributes b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1231 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1232 | val _ = | 
| 67147 | 1233 | Outer_Syntax.command \<^command_keyword>\<open>print_simpset\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1234 | "print context of Simplifier" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1235 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1236 | Toplevel.keep (Pretty.writeln o Simplifier.pretty_simpset b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1237 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1238 | val _ = | 
| 67147 | 1239 | Outer_Syntax.command \<^command_keyword>\<open>print_rules\<close> "print intro/elim rules" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1240 | (Scan.succeed (Toplevel.keep (Context_Rules.print_rules o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1241 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1242 | val _ = | 
| 67147 | 1243 | Outer_Syntax.command \<^command_keyword>\<open>print_methods\<close> "print methods of this theory" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1244 | (Parse.opt_bang >> (fn b => Toplevel.keep (Method.print_methods b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1245 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1246 | val _ = | 
| 67147 | 1247 | Outer_Syntax.command \<^command_keyword>\<open>print_antiquotations\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1248 | "print document antiquotations" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1249 | (Parse.opt_bang >> (fn b => | 
| 67386 | 1250 | Toplevel.keep (Document_Antiquotation.print_antiquotations b o Toplevel.context_of))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1251 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1252 | val _ = | 
| 67147 | 1253 | Outer_Syntax.command \<^command_keyword>\<open>print_ML_antiquotations\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1254 | "print ML antiquotations" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1255 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1256 | Toplevel.keep (ML_Context.print_antiquotations b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1257 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1258 | val _ = | 
| 67147 | 1259 | Outer_Syntax.command \<^command_keyword>\<open>locale_deps\<close> "visualize locale dependencies" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1260 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1261 | (Toplevel.keep (Toplevel.theory_of #> (fn thy => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1262 | Locale.pretty_locale_deps thy | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1263 |         |> map (fn {name, parents, body} =>
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1264 | ((name, Graph_Display.content_node (Locale.extern thy name) [body]), parents)) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1265 | |> Graph_Display.display_graph_old)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1266 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1267 | val _ = | 
| 67147 | 1268 | Outer_Syntax.command \<^command_keyword>\<open>print_term_bindings\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1269 | "print term bindings of proof context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1270 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1271 | (Toplevel.keep | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1272 | (Pretty.writeln_chunks o Proof_Context.pretty_term_bindings o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1273 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1274 | val _ = | 
| 67147 | 1275 | Outer_Syntax.command \<^command_keyword>\<open>print_facts\<close> "print facts of proof context" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1276 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1277 | Toplevel.keep (Proof_Context.print_local_facts b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1278 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1279 | val _ = | 
| 67147 | 1280 | Outer_Syntax.command \<^command_keyword>\<open>print_cases\<close> "print cases of proof context" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1281 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1282 | (Toplevel.keep (Pretty.writeln_chunks o Proof_Context.pretty_cases o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1283 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1284 | val _ = | 
| 67147 | 1285 | Outer_Syntax.command \<^command_keyword>\<open>print_statement\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1286 | "print theorems as long statements" | 
| 62969 | 1287 | (opt_modes -- Parse.thms1 >> Isar_Cmd.print_stmts); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1288 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1289 | val _ = | 
| 67147 | 1290 | Outer_Syntax.command \<^command_keyword>\<open>thm\<close> "print theorems" | 
| 62969 | 1291 | (opt_modes -- Parse.thms1 >> Isar_Cmd.print_thms); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1292 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1293 | val _ = | 
| 67147 | 1294 | Outer_Syntax.command \<^command_keyword>\<open>prf\<close> "print proof terms of theorems" | 
| 62969 | 1295 | (opt_modes -- Scan.option Parse.thms1 >> Isar_Cmd.print_prfs false); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1296 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1297 | val _ = | 
| 67147 | 1298 | Outer_Syntax.command \<^command_keyword>\<open>full_prf\<close> "print full proof terms of theorems" | 
| 62969 | 1299 | (opt_modes -- Scan.option Parse.thms1 >> Isar_Cmd.print_prfs true); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1300 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1301 | val _ = | 
| 67147 | 1302 | Outer_Syntax.command \<^command_keyword>\<open>prop\<close> "read and print proposition" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1303 | (opt_modes -- Parse.term >> Isar_Cmd.print_prop); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1304 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1305 | val _ = | 
| 67147 | 1306 | Outer_Syntax.command \<^command_keyword>\<open>term\<close> "read and print term" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1307 | (opt_modes -- Parse.term >> Isar_Cmd.print_term); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1308 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1309 | val _ = | 
| 67147 | 1310 | Outer_Syntax.command \<^command_keyword>\<open>typ\<close> "read and print type" | 
| 1311 | (opt_modes -- (Parse.typ -- Scan.option (\<^keyword>\<open>::\<close> |-- Parse.!!! Parse.sort)) | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1312 | >> Isar_Cmd.print_type); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1313 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1314 | val _ = | 
| 67147 | 1315 | Outer_Syntax.command \<^command_keyword>\<open>print_codesetup\<close> "print code generator setup" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1316 | (Scan.succeed (Toplevel.keep (Code.print_codesetup o Toplevel.theory_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1317 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1318 | val _ = | 
| 78526 | 1319 | Outer_Syntax.command \<^command_keyword>\<open>print_context_tracing\<close> | 
| 1320 | "print result of context tracing from ML heap" | |
| 78528 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1321 | (Scan.repeat Parse.name_position >> (fn raw_names => Toplevel.keep (fn st => | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1322 | let | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1323 | val pred = | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1324 | if null raw_names then K true | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1325 | else | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1326 | let | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1327 | val ctxt = Toplevel.context_of st; | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1328 | val insert = Symset.insert o Context.theory_long_name o Thy_Info.check_theory ctxt; | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1329 | val names = Symset.build (fold insert raw_names); | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1330 | in Symset.member names o Context.theory_long_name o Context.theory_of end; | 
| 
3d6dbf215559
clarified command arguments: optionally restrict to given theories (from theory loader);
 wenzelm parents: 
78526diff
changeset | 1331 | in Session.print_context_tracing pred end))); | 
| 78526 | 1332 | |
| 1333 | val _ = | |
| 67147 | 1334 | Outer_Syntax.command \<^command_keyword>\<open>print_state\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1335 | "print current proof state (if present)" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1336 | (opt_modes >> (fn modes => | 
| 76077 
0f48e873e187
clarified message channel for 'print_state' (NB: the command was originally for TTY or Proof General);
 wenzelm parents: 
75687diff
changeset | 1337 | Toplevel.keep (Print_Mode.with_modes modes (Output.writeln o Toplevel.string_of_state)))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1338 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1339 | val _ = | 
| 67147 | 1340 | Outer_Syntax.command \<^command_keyword>\<open>welcome\<close> "print welcome message" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1341 | (Scan.succeed (Toplevel.keep (fn _ => writeln (Session.welcome ())))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1342 | |
| 62856 | 1343 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1344 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1345 | |
| 62856 | 1346 | subsection \<open>Dependencies\<close> | 
| 1347 | ||
| 1348 | ML \<open> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1349 | local | 
| 62856 | 1350 | |
| 1351 | val theory_bounds = | |
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 1352 | Parse.theory_name >> single || | 
| 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 1353 | (\<^keyword>\<open>(\<close> |-- Parse.enum "|" Parse.theory_name --| \<^keyword>\<open>)\<close>); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1354 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1355 | val _ = | 
| 67147 | 1356 | Outer_Syntax.command \<^command_keyword>\<open>thy_deps\<close> "visualize theory dependencies" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1357 | (Scan.option theory_bounds -- Scan.option theory_bounds >> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1358 | (fn args => Toplevel.keep (fn st => Thy_Deps.thy_deps_cmd (Toplevel.context_of st) args))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1359 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1360 | |
| 62856 | 1361 | val class_bounds = | 
| 1362 | Parse.sort >> single || | |
| 67147 | 1363 | (\<^keyword>\<open>(\<close> |-- Parse.enum "|" Parse.sort --| \<^keyword>\<open>)\<close>); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1364 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1365 | val _ = | 
| 67147 | 1366 | Outer_Syntax.command \<^command_keyword>\<open>class_deps\<close> "visualize class dependencies" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1367 | (Scan.option class_bounds -- Scan.option class_bounds >> (fn args => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1368 | Toplevel.keep (fn st => Class_Deps.class_deps_cmd (Toplevel.context_of st) args))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1369 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1370 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1371 | val _ = | 
| 70605 | 1372 | Outer_Syntax.command \<^command_keyword>\<open>thm_deps\<close> | 
| 1373 | "print theorem dependencies (immediate non-transitive)" | |
| 62969 | 1374 | (Parse.thms1 >> (fn args => | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1375 | Toplevel.keep (fn st => | 
| 70570 | 1376 | let | 
| 1377 | val thy = Toplevel.theory_of st; | |
| 1378 | val ctxt = Toplevel.context_of st; | |
| 70605 | 1379 | in Pretty.writeln (Thm_Deps.pretty_thm_deps thy (Attrib.eval_thms ctxt args)) end))); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1380 | |
| 70560 | 1381 | val _ = | 
| 1382 | Outer_Syntax.command \<^command_keyword>\<open>thm_oracles\<close> | |
| 1383 | "print all oracles used in theorems (full graph of transitive dependencies)" | |
| 1384 | (Parse.thms1 >> (fn args => | |
| 1385 | Toplevel.keep (fn st => | |
| 1386 | let | |
| 1387 | val ctxt = Toplevel.context_of st; | |
| 1388 | val thms = Attrib.eval_thms ctxt args; | |
| 1389 | in Pretty.writeln (Thm_Deps.pretty_thm_oracles ctxt thms) end))); | |
| 62856 | 1390 | |
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
69262diff
changeset | 1391 | val thy_names = Scan.repeat1 (Scan.unless Parse.minus Parse.theory_name); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1392 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1393 | val _ = | 
| 67147 | 1394 | Outer_Syntax.command \<^command_keyword>\<open>unused_thms\<close> "find unused theorems" | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1395 | (Scan.option ((thy_names --| Parse.minus) -- Scan.option thy_names) >> (fn opt_range => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1396 | Toplevel.keep (fn st => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1397 | let | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1398 | val thy = Toplevel.theory_of st; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1399 | val ctxt = Toplevel.context_of st; | 
| 80299 
a397fd0c451a
more accurate output of Thm_Name.T wrt. facts name space;
 wenzelm parents: 
80295diff
changeset | 1400 | fun pretty_thm (a, th) = | 
| 
a397fd0c451a
more accurate output of Thm_Name.T wrt. facts name space;
 wenzelm parents: 
80295diff
changeset | 1401 | Pretty.block [Pretty.quote (Proof_Context.pretty_thm_name ctxt a), | 
| 
a397fd0c451a
more accurate output of Thm_Name.T wrt. facts name space;
 wenzelm parents: 
80295diff
changeset | 1402 | Pretty.str ":", Pretty.brk 1, Thm.pretty_thm ctxt th]; | 
| 68482 | 1403 |             val check = Theory.check {long = false} ctxt;
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1404 | in | 
| 70570 | 1405 | Thm_Deps.unused_thms_cmd | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1406 | (case opt_range of | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1407 | NONE => (Theory.parents_of thy, [thy]) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1408 | | SOME (xs, NONE) => (map check xs, [thy]) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1409 | | SOME (xs, SOME ys) => (map check xs, map check ys)) | 
| 80299 
a397fd0c451a
more accurate output of Thm_Name.T wrt. facts name space;
 wenzelm parents: 
80295diff
changeset | 1410 | |> map pretty_thm |> Pretty.writeln_chunks | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1411 | end))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1412 | |
| 62856 | 1413 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1414 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1415 | |
| 62856 | 1416 | subsubsection \<open>Find consts and theorems\<close> | 
| 1417 | ||
| 1418 | ML \<open> | |
| 1419 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1420 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1421 | val _ = | 
| 67147 | 1422 | Outer_Syntax.command \<^command_keyword>\<open>find_consts\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1423 | "find constants by name / type patterns" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1424 | (Find_Consts.query_parser >> (fn spec => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1425 | Toplevel.keep (fn st => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1426 | Pretty.writeln (Find_Consts.pretty_consts (Toplevel.context_of st) spec)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1427 | |
| 62856 | 1428 | val options = | 
| 1429 | Scan.optional | |
| 1430 |     (Parse.$$$ "(" |--
 | |
| 1431 | Parse.!!! (Scan.option Parse.nat -- | |
| 1432 | Scan.optional (Parse.reserved "with_dups" >> K false) true --| Parse.$$$ ")")) | |
| 1433 | (NONE, true); | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1434 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1435 | val _ = | 
| 67147 | 1436 | Outer_Syntax.command \<^command_keyword>\<open>find_theorems\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1437 | "find theorems meeting specified criteria" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1438 | (options -- Find_Theorems.query_parser >> (fn ((opt_lim, rem_dups), spec) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1439 | Toplevel.keep (fn st => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1440 | Pretty.writeln | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1441 | (Find_Theorems.pretty_theorems (Find_Theorems.proof_state st) opt_lim rem_dups spec)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1442 | |
| 62856 | 1443 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1444 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1445 | |
| 62856 | 1446 | subsection \<open>Code generation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1447 | |
| 62856 | 1448 | ML \<open> | 
| 1449 | local | |
| 1450 | ||
| 1451 | val _ = | |
| 67147 | 1452 | Outer_Syntax.command \<^command_keyword>\<open>code_datatype\<close> | 
| 62856 | 1453 | "define set of code datatype constructors" | 
| 66251 
cd935b7cb3fb
proper concept of code declaration wrt. atomicity and Isar declarations
 haftmann parents: 
66248diff
changeset | 1454 | (Scan.repeat1 Parse.term >> (Toplevel.theory o Code.declare_datatype_cmd)); | 
| 62856 | 1455 | |
| 1456 | in end\<close> | |
| 1457 | ||
| 1458 | ||
| 1459 | subsection \<open>Extraction of programs from proofs\<close> | |
| 1460 | ||
| 1461 | ML \<open> | |
| 1462 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1463 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1464 | val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") [];
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1465 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1466 | val _ = | 
| 67147 | 1467 | Outer_Syntax.command \<^command_keyword>\<open>realizers\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1468 | "specify realizers for primitive axioms / theorems, together with correctness proof" | 
| 62969 | 1469 | (Scan.repeat1 (Parse.name -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1470 | (fn xs => Toplevel.theory (fn thy => Extraction.add_realizers | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1471 | (map (fn (((a, vs), s1), s2) => (Global_Theory.get_thm thy a, (vs, s1, s2))) xs) thy))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1472 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1473 | val _ = | 
| 67147 | 1474 | Outer_Syntax.command \<^command_keyword>\<open>realizability\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1475 | "add equations characterizing realizability" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1476 | (Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_realizes_eqns)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1477 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1478 | val _ = | 
| 67147 | 1479 | Outer_Syntax.command \<^command_keyword>\<open>extract_type\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1480 | "add equations characterizing type of extracted program" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1481 | (Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_typeof_eqns)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1482 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1483 | val _ = | 
| 67147 | 1484 | Outer_Syntax.command \<^command_keyword>\<open>extract\<close> "extract terms from proofs" | 
| 62969 | 1485 | (Scan.repeat1 (Parse.name -- parse_vars) >> (fn xs => Toplevel.theory (fn thy => | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1486 | Extraction.extract (map (apfst (Global_Theory.get_thm thy)) xs) thy))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1487 | |
| 62856 | 1488 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1489 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1490 | |
| 62944 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62913diff
changeset | 1491 | section \<open>Auxiliary lemmas\<close> | 
| 20627 | 1492 | |
| 58611 | 1493 | subsection \<open>Meta-level connectives in assumptions\<close> | 
| 15803 | 1494 | |
| 1495 | lemma meta_mp: | |
| 58612 | 1496 | assumes "PROP P \<Longrightarrow> PROP Q" and "PROP P" | 
| 15803 | 1497 | shows "PROP Q" | 
| 58612 | 1498 | by (rule \<open>PROP P \<Longrightarrow> PROP Q\<close> [OF \<open>PROP P\<close>]) | 
| 15803 | 1499 | |
| 23432 | 1500 | lemmas meta_impE = meta_mp [elim_format] | 
| 1501 | ||
| 15803 | 1502 | lemma meta_spec: | 
| 58612 | 1503 | assumes "\<And>x. PROP P x" | 
| 26958 | 1504 | shows "PROP P x" | 
| 58612 | 1505 | by (rule \<open>\<And>x. PROP P x\<close>) | 
| 15803 | 1506 | |
| 1507 | lemmas meta_allE = meta_spec [elim_format] | |
| 1508 | ||
| 26570 | 1509 | lemma swap_params: | 
| 58612 | 1510 | "(\<And>x y. PROP P x y) \<equiv> (\<And>y x. PROP P x y)" .. | 
| 26570 | 1511 | |
| 71512 
fe93a863d946
infrastructure for extraction of equations x = t from premises beneath meta-all
 haftmann parents: 
71166diff
changeset | 1512 | lemma equal_allI: | 
| 
fe93a863d946
infrastructure for extraction of equations x = t from premises beneath meta-all
 haftmann parents: 
71166diff
changeset | 1513 | \<open>(\<And>x. PROP P x) \<equiv> (\<And>x. PROP Q x)\<close> if \<open>\<And>x. PROP P x \<equiv> PROP Q x\<close> | 
| 
fe93a863d946
infrastructure for extraction of equations x = t from premises beneath meta-all
 haftmann parents: 
71166diff
changeset | 1514 | by (simp only: that) | 
| 
fe93a863d946
infrastructure for extraction of equations x = t from premises beneath meta-all
 haftmann parents: 
71166diff
changeset | 1515 | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1516 | |
| 58611 | 1517 | subsection \<open>Meta-level conjunction\<close> | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1518 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1519 | lemma all_conjunction: | 
| 58612 | 1520 | "(\<And>x. PROP A x &&& PROP B x) \<equiv> ((\<And>x. PROP A x) &&& (\<And>x. PROP B x))" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1521 | proof | 
| 58612 | 1522 | assume conj: "\<And>x. PROP A x &&& PROP B x" | 
| 1523 | show "(\<And>x. PROP A x) &&& (\<And>x. PROP B x)" | |
| 19121 | 1524 | proof - | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1525 | fix x | 
| 26958 | 1526 | from conj show "PROP A x" by (rule conjunctionD1) | 
| 1527 | from conj show "PROP B x" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1528 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1529 | next | 
| 58612 | 1530 | assume conj: "(\<And>x. PROP A x) &&& (\<And>x. PROP B x)" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1531 | fix x | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 1532 | show "PROP A x &&& PROP B x" | 
| 19121 | 1533 | proof - | 
| 26958 | 1534 | show "PROP A x" by (rule conj [THEN conjunctionD1, rule_format]) | 
| 1535 | show "PROP B x" by (rule conj [THEN conjunctionD2, rule_format]) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1536 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1537 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1538 | |
| 19121 | 1539 | lemma imp_conjunction: | 
| 58612 | 1540 | "(PROP A \<Longrightarrow> PROP B &&& PROP C) \<equiv> ((PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C))" | 
| 18836 | 1541 | proof | 
| 58612 | 1542 | assume conj: "PROP A \<Longrightarrow> PROP B &&& PROP C" | 
| 1543 | show "(PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C)" | |
| 19121 | 1544 | proof - | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1545 | assume "PROP A" | 
| 58611 | 1546 | from conj [OF \<open>PROP A\<close>] show "PROP B" by (rule conjunctionD1) | 
| 1547 | from conj [OF \<open>PROP A\<close>] show "PROP C" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1548 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1549 | next | 
| 58612 | 1550 | assume conj: "(PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C)" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1551 | assume "PROP A" | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 1552 | show "PROP B &&& PROP C" | 
| 19121 | 1553 | proof - | 
| 58611 | 1554 | from \<open>PROP A\<close> show "PROP B" by (rule conj [THEN conjunctionD1]) | 
| 1555 | from \<open>PROP A\<close> show "PROP C" by (rule conj [THEN conjunctionD2]) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1556 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1557 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1558 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1559 | lemma conjunction_imp: | 
| 58612 | 1560 | "(PROP A &&& PROP B \<Longrightarrow> PROP C) \<equiv> (PROP A \<Longrightarrow> PROP B \<Longrightarrow> PROP C)" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1561 | proof | 
| 58612 | 1562 | assume r: "PROP A &&& PROP B \<Longrightarrow> PROP C" | 
| 22933 | 1563 | assume ab: "PROP A" "PROP B" | 
| 1564 | show "PROP C" | |
| 1565 | proof (rule r) | |
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 1566 | from ab show "PROP A &&& PROP B" . | 
| 22933 | 1567 | qed | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1568 | next | 
| 58612 | 1569 | assume r: "PROP A \<Longrightarrow> PROP B \<Longrightarrow> PROP C" | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 1570 | assume conj: "PROP A &&& PROP B" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1571 | show "PROP C" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1572 | proof (rule r) | 
| 19121 | 1573 | from conj show "PROP A" by (rule conjunctionD1) | 
| 1574 | from conj show "PROP B" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1575 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1576 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1577 | |
| 81136 | 1578 | |
| 1579 | section \<open>Misc\<close> | |
| 1580 | ||
| 1581 | bundle constrain_space_syntax \<comment> \<open>type constraints with spacing\<close> | |
| 1582 | begin | |
| 1583 | no_syntax (output) | |
| 81148 | 1584 | "_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" (\<open>_::_\<close> [4, 0] 3) | 
| 1585 | "_constrain" :: "prop' \<Rightarrow> type \<Rightarrow> prop'" (\<open>_::_\<close> [4, 0] 3) | |
| 81136 | 1586 | syntax (output) | 
| 81148 | 1587 | "_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" (\<open>_ :: _\<close> [4, 0] 3) | 
| 1588 | "_constrain" :: "prop' \<Rightarrow> type \<Rightarrow> prop'" (\<open>_ :: _\<close> [4, 0] 3) | |
| 81136 | 1589 | end | 
| 1590 | ||
| 1591 | ||
| 68816 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68813diff
changeset | 1592 | declare [[ML_write_global = false]] | 
| 
5a53724fe247
support named ML environments, notably "Isabelle", "SML";
 wenzelm parents: 
68813diff
changeset | 1593 | |
| 78728 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
78528diff
changeset | 1594 | ML_command \<open>\<^assert> (not (can ML_command \<open>() handle _ => ()\<close>))\<close> | 
| 78740 
45ff003d337c
discontinue obsolete "Interrupt" constructor (NB: catch-all pattern produces ML compiler error);
 wenzelm parents: 
78728diff
changeset | 1595 | ML_command \<open>\<^assert> (not (can ML_command \<open>() handle Interrupt => ()\<close>))\<close> | 
| 78728 
72631efa3821
explicitly reject 'handle' with catch-all patterns;
 wenzelm parents: 
78528diff
changeset | 1596 | |
| 48638 | 1597 | end |