| author | wenzelm | 
| Wed, 10 Aug 2016 22:05:36 +0200 | |
| changeset 63654 | f90e3926e627 | 
| parent 63579 | 73939a9b70a3 | 
| child 63808 | e8462a4349fc | 
| 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 | 
| 63579 | 9 | "!!" "!" "+" "--" ":" ";" "<" "<=" "==" "=>" "?" "[" "\<comment>" "\<equiv>" "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>" | 
| 63434 | 10 | "\<subseteq>" "]" "attach" "binder" "for" "if" "in" "infix" "infixl" "infixr" "is" | 
| 63451 | 11 | "open" "output" "overloaded" "pervasive" "premises" "structure" "unchecked" "when" | 
| 63441 | 12 | and "private" "qualified" :: before_command | 
| 63434 | 13 | and "assumes" "constrains" "defines" "fixes" "includes" "notes" "rewrites" | 
| 63451 | 14 | "obtains" "shows" "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 | 
| 57506 | 18 | and "typedecl" "type_synonym" "nonterminal" "judgment" | 
| 62169 | 19 | "consts" "syntax" "no_syntax" "translations" "no_translations" | 
| 55385 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 wenzelm parents: 
55152diff
changeset | 20 | "definition" "abbreviation" "type_notation" "no_type_notation" "notation" | 
| 61337 | 21 | "no_notation" "axiomatization" "lemmas" "declare" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 22 | "hide_class" "hide_type" "hide_const" "hide_fact" :: thy_decl | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 23 | 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 | 24 | and "SML_file" "SML_file_debug" "SML_file_no_debug" :: thy_load % "ML" | 
| 56618 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 wenzelm parents: 
56275diff
changeset | 25 | and "SML_import" "SML_export" :: thy_decl % "ML" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 26 | and "ML_prf" :: prf_decl % "proof" (* FIXME % "ML" ?? *) | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 27 | and "ML_val" "ML_command" :: diag % "ML" | 
| 63579 | 28 | and "simproc_setup" :: thy_decl % "ML" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 29 | and "setup" "local_setup" "attribute_setup" "method_setup" | 
| 55762 
27a45aec67a0
suppress completion of obscure keyword, avoid confusion with plain "simp";
 wenzelm parents: 
55516diff
changeset | 30 | "declaration" "syntax_declaration" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 31 | "parse_ast_translation" "parse_translation" "print_translation" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 32 | "typed_print_translation" "print_ast_translation" "oracle" :: thy_decl % "ML" | 
| 63273 | 33 | and "bundle" :: thy_decl_block | 
| 63282 | 34 | and "unbundle" :: thy_decl | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 35 | and "include" "including" :: prf_decl | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 36 | and "print_bundles" :: diag | 
| 59901 | 37 | 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 | 38 | and "interpret" :: prf_goal % "proof" | 
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61853diff
changeset | 39 | and "interpretation" "global_interpretation" "sublocale" :: thy_goal | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 40 | and "class" :: thy_decl_block | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 41 | and "subclass" :: thy_goal | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 42 | and "instantiation" :: thy_decl_block | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 43 | and "instance" :: thy_goal | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 44 | and "overloading" :: thy_decl_block | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 45 | and "code_datatype" :: thy_decl | 
| 61338 | 46 | and "theorem" "lemma" "corollary" "proposition" :: thy_goal | 
| 61337 | 47 | and "schematic_goal" :: thy_goal | 
| 58800 
bfed1c26caed
explicit keyword category for commands that may start a block;
 wenzelm parents: 
58660diff
changeset | 48 | and "notepad" :: thy_decl_block | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 49 | and "have" :: prf_goal % "proof" | 
| 63579 | 50 | and "hence" :: prf_goal % "proof" | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 51 | and "show" :: prf_asm_goal % "proof" | 
| 63579 | 52 | and "thus" :: prf_asm_goal % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 53 | and "then" "from" "with" :: prf_chain % "proof" | 
| 60371 | 54 | and "note" :: prf_decl % "proof" | 
| 55 | and "supply" :: prf_script % "proof" | |
| 56 | and "using" "unfolding" :: prf_decl % "proof" | |
| 63039 | 57 | and "fix" "assume" "presume" "define" "def" :: prf_asm % "proof" | 
| 60448 | 58 | and "consider" :: prf_goal % "proof" | 
| 53371 
47b23c582127
more explicit indication of 'guess' as improper Isar (aka "script") element;
 wenzelm parents: 
52549diff
changeset | 59 | and "obtain" :: prf_asm_goal % "proof" | 
| 60624 | 60 | and "guess" :: prf_script_asm_goal % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 61 | and "let" "write" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 62 | and "case" :: prf_asm % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 63 |   and "{" :: prf_open % "proof"
 | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 64 | and "}" :: prf_close % "proof" | 
| 60694 
b3fa4a8cdb5f
clarified text folds: proof ... qed counts as extra block;
 wenzelm parents: 
60624diff
changeset | 65 | and "next" :: next_block % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 66 | and "qed" :: qed_block % "proof" | 
| 62312 
5e5a881ebc12
command '\<proof>' is an alias for 'sorry', with different typesetting;
 wenzelm parents: 
62169diff
changeset | 67 | and "by" ".." "." "sorry" "\<proof>" :: "qed" % "proof" | 
| 53571 
e58ca0311c0f
more explicit indication of 'done' as proof script element;
 wenzelm parents: 
53371diff
changeset | 68 | and "done" :: "qed_script" % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 69 | and "oops" :: qed_global % "proof" | 
| 50128 
599c935aac82
alternative completion for outer syntax keywords;
 wenzelm parents: 
49569diff
changeset | 70 | and "defer" "prefer" "apply" :: prf_script % "proof" | 
| 63579 | 71 | and "apply_end" :: prf_script % "proof" | 
| 60624 | 72 | and "subgoal" :: prf_script_goal % "proof" | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 73 | and "proof" :: prf_block % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 74 | and "also" "moreover" :: prf_decl % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 75 | and "finally" "ultimately" :: prf_chain % "proof" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 76 | and "back" :: prf_script % "proof" | 
| 61252 | 77 | and "help" "print_commands" "print_options" "print_context" "print_theory" | 
| 78 | "print_definitions" "print_syntax" "print_abbrevs" "print_defn_rules" | |
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 79 | "print_theorems" "print_locales" "print_classes" "print_locale" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 80 | "print_interps" "print_dependencies" "print_attributes" | 
| 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 81 | "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 | 82 | "print_antiquotations" "print_ML_antiquotations" "thy_deps" | 
| 58845 | 83 | "locale_deps" "class_deps" "thm_deps" "print_term_bindings" | 
| 57415 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 wenzelm parents: 
56864diff
changeset | 84 | "print_facts" "print_cases" "print_statement" "thm" "prf" "full_prf" | 
| 
e721124f1b1e
command 'print_term_bindings' supersedes 'print_binds';
 wenzelm parents: 
56864diff
changeset | 85 | "prop" "term" "typ" "print_codesetup" "unused_thms" :: diag | 
| 58845 | 86 | and "display_drafts" "print_state" :: diag | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 87 | and "welcome" :: diag | 
| 48641 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 wenzelm parents: 
48638diff
changeset | 88 | and "end" :: thy_end % "theory" | 
| 63579 | 89 | and "realizers" :: thy_decl | 
| 90 | and "realizability" :: thy_decl | |
| 56797 | 91 | and "extract_type" "extract" :: thy_decl | 
| 48646 
91281e9472d8
more official command specifications, including source position;
 wenzelm parents: 
48641diff
changeset | 92 | and "find_theorems" "find_consts" :: diag | 
| 57886 
7cae177c9084
support for named collections of theorems in canonical order;
 wenzelm parents: 
57506diff
changeset | 93 | and "named_theorems" :: thy_decl | 
| 63579 | 94 | abbrevs | 
| 95 | "default_sort" = "" | |
| 96 | "simproc_setup" = "" | |
| 97 | "hence" = "then have" | |
| 98 | "thus" = "then show" | |
| 99 | "apply_end" = "" | |
| 100 | "realizers" = "" | |
| 101 | "realizability" = "" | |
| 48638 | 102 | begin | 
| 15803 | 103 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 104 | section \<open>Isar commands\<close> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 105 | |
| 62856 | 106 | subsection \<open>Embedded ML text\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 107 | |
| 62856 | 108 | ML \<open> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 109 | local | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 110 | |
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 111 | val semi = Scan.option @{keyword ";"};
 | 
| 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 112 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 113 | val _ = | 
| 62867 | 114 |   Outer_Syntax.command @{command_keyword ML_file} "read and evaluate Isabelle/ML file"
 | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 115 | (Resources.parse_files "ML_file" --| semi >> ML_File.ML NONE); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 116 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 117 | val _ = | 
| 62867 | 118 |   Outer_Syntax.command @{command_keyword ML_file_debug}
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 119 | "read and evaluate Isabelle/ML file (with debugger information)" | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 120 | (Resources.parse_files "ML_file_debug" --| semi >> ML_File.ML (SOME true)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 121 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 122 | val _ = | 
| 62867 | 123 |   Outer_Syntax.command @{command_keyword ML_file_no_debug}
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 124 | "read and evaluate Isabelle/ML file (no debugger information)" | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 125 | (Resources.parse_files "ML_file_no_debug" --| semi >> ML_File.ML (SOME false)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 126 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 127 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 128 |   Outer_Syntax.command @{command_keyword SML_file} "read and evaluate Standard ML file"
 | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 129 | (Resources.parse_files "SML_file" --| semi >> ML_File.SML NONE); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 130 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 131 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 132 |   Outer_Syntax.command @{command_keyword SML_file_debug}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 133 | "read and evaluate Standard ML file (with debugger information)" | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 134 | (Resources.parse_files "SML_file_debug" --| semi >> ML_File.SML (SOME true)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 135 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 136 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 137 |   Outer_Syntax.command @{command_keyword SML_file_no_debug}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 138 | "read and evaluate Standard ML file (no debugger information)" | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 139 | (Resources.parse_files "SML_file_no_debug" --| semi >> ML_File.SML (SOME false)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 140 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 141 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 142 |   Outer_Syntax.command @{command_keyword SML_export} "evaluate SML within Isabelle/ML environment"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 143 | (Parse.ML_source >> (fn source => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 144 | let | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 145 | val flags: ML_Compiler.flags = | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 146 |           {SML = true, exchange = true, redirect = false, verbose = true,
 | 
| 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 147 | debug = NONE, writeln = writeln, warning = warning}; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 148 | in | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 149 | Toplevel.theory | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 150 | (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 | 151 | end)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 152 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 153 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 154 |   Outer_Syntax.command @{command_keyword SML_import} "evaluate Isabelle/ML within SML environment"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 155 | (Parse.ML_source >> (fn source => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 156 | let | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 157 | val flags: ML_Compiler.flags = | 
| 62902 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 158 |           {SML = false, exchange = true, redirect = false, verbose = true,
 | 
| 
3c0f53eae166
more conventional theory syntax for ML bootstrap, with 'ML_file' instead of 'use';
 wenzelm parents: 
62898diff
changeset | 159 | debug = NONE, writeln = writeln, warning = warning}; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 160 | in | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 161 | Toplevel.generic_theory | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 162 | (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 | 163 | Local_Theory.propagate_ml_env) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 164 | end)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 165 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 166 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 167 |   Outer_Syntax.command @{command_keyword ML_prf} "ML text within proof"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 168 | (Parse.ML_source >> (fn source => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 169 | Toplevel.proof (Proof.map_context (Context.proof_map | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 170 | (ML_Context.exec (fn () => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 171 | 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 | 172 | Proof.propagate_ml_env))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 173 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 174 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 175 |   Outer_Syntax.command @{command_keyword ML_val} "diagnostic ML text"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 176 | (Parse.ML_source >> Isar_Cmd.ml_diag true); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 177 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 178 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 179 |   Outer_Syntax.command @{command_keyword ML_command} "diagnostic ML text (silent)"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 180 | (Parse.ML_source >> Isar_Cmd.ml_diag false); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 181 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 182 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 183 |   Outer_Syntax.command @{command_keyword setup} "ML setup for global theory"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 184 | (Parse.ML_source >> (Toplevel.theory o Isar_Cmd.setup)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 185 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 186 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 187 |   Outer_Syntax.local_theory @{command_keyword local_setup} "ML setup for local theory"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 188 | (Parse.ML_source >> Isar_Cmd.local_setup); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 189 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 190 | val _ = | 
| 62856 | 191 |   Outer_Syntax.command @{command_keyword oracle} "declare oracle"
 | 
| 192 |     (Parse.range Parse.name -- (@{keyword "="} |-- Parse.ML_source) >>
 | |
| 193 | (fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y))); | |
| 194 | ||
| 195 | val _ = | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 196 |   Outer_Syntax.local_theory @{command_keyword attribute_setup} "define attribute in ML"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 197 | (Parse.position Parse.name -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 198 |         Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 199 | >> (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 | 200 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 201 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 202 |   Outer_Syntax.local_theory @{command_keyword method_setup} "define proof method in ML"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 203 | (Parse.position Parse.name -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 204 |         Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 205 | >> (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 | 206 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 207 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 208 |   Outer_Syntax.local_theory @{command_keyword declaration} "generic ML declaration"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 209 | (Parse.opt_keyword "pervasive" -- Parse.ML_source | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 210 |       >> (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 | 211 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 212 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 213 |   Outer_Syntax.local_theory @{command_keyword syntax_declaration} "generic ML syntax declaration"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 214 | (Parse.opt_keyword "pervasive" -- Parse.ML_source | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 215 |       >> (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 | 216 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 217 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 218 |   Outer_Syntax.local_theory @{command_keyword simproc_setup} "define simproc in ML"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 219 | (Parse.position Parse.name -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 220 |       (@{keyword "("} |-- Parse.enum1 "|" Parse.term --| @{keyword ")"} --| @{keyword "="}) --
 | 
| 62913 | 221 | Parse.ML_source >> (fn ((a, b), c) => Isar_Cmd.simproc_setup a b c)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 222 | |
| 62856 | 223 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 224 | |
| 62856 | 225 | |
| 226 | subsection \<open>Theory commands\<close> | |
| 227 | ||
| 228 | subsubsection \<open>Sorts and types\<close> | |
| 229 | ||
| 230 | ML \<open> | |
| 231 | local | |
| 232 | ||
| 233 | val _ = | |
| 234 |   Outer_Syntax.local_theory @{command_keyword default_sort}
 | |
| 235 | "declare default sort for explicit type variables" | |
| 236 | (Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy)); | |
| 237 | ||
| 238 | val _ = | |
| 239 |   Outer_Syntax.local_theory @{command_keyword typedecl} "type declaration"
 | |
| 240 | (Parse.type_args -- Parse.binding -- Parse.opt_mixfix | |
| 241 | >> (fn ((args, a), mx) => | |
| 242 |           Typedecl.typedecl {final = true} (a, map (rpair dummyS) args, mx) #> snd));
 | |
| 243 | ||
| 244 | val _ = | |
| 245 |   Outer_Syntax.local_theory @{command_keyword type_synonym} "declare type abbreviation"
 | |
| 246 | (Parse.type_args -- Parse.binding -- | |
| 247 |       (@{keyword "="} |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix'))
 | |
| 248 | >> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs)); | |
| 249 | ||
| 250 | in end\<close> | |
| 251 | ||
| 252 | ||
| 253 | subsubsection \<open>Consts\<close> | |
| 254 | ||
| 255 | ML \<open> | |
| 256 | local | |
| 257 | ||
| 258 | val _ = | |
| 259 |   Outer_Syntax.command @{command_keyword judgment} "declare object-logic judgment"
 | |
| 260 | (Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd)); | |
| 261 | ||
| 262 | val _ = | |
| 263 |   Outer_Syntax.command @{command_keyword consts} "declare constants"
 | |
| 264 | (Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts_cmd)); | |
| 265 | ||
| 266 | in end\<close> | |
| 267 | ||
| 268 | ||
| 269 | subsubsection \<open>Syntax and translations\<close> | |
| 270 | ||
| 271 | ML \<open> | |
| 272 | local | |
| 273 | ||
| 274 | val _ = | |
| 275 |   Outer_Syntax.command @{command_keyword nonterminal}
 | |
| 276 | "declare syntactic type constructors (grammar nonterminal symbols)" | |
| 277 | (Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global)); | |
| 278 | ||
| 279 | val _ = | |
| 280 |   Outer_Syntax.command @{command_keyword syntax} "add raw syntax clauses"
 | |
| 281 | (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl | |
| 282 | >> (Toplevel.theory o uncurry Sign.add_syntax_cmd)); | |
| 283 | ||
| 284 | val _ = | |
| 285 |   Outer_Syntax.command @{command_keyword no_syntax} "delete raw syntax clauses"
 | |
| 286 | (Parse.syntax_mode -- Scan.repeat1 Parse.const_decl | |
| 287 | >> (Toplevel.theory o uncurry Sign.del_syntax_cmd)); | |
| 288 | ||
| 289 | val trans_pat = | |
| 290 | Scan.optional | |
| 62969 | 291 |     (@{keyword "("} |-- Parse.!!! (Parse.inner_syntax Parse.name --| @{keyword ")"})) "logic"
 | 
| 62856 | 292 | -- Parse.inner_syntax Parse.string; | 
| 293 | ||
| 294 | fun trans_arrow toks = | |
| 295 |   ((@{keyword "\<rightharpoonup>"} || @{keyword "=>"}) >> K Syntax.Parse_Rule ||
 | |
| 296 |     (@{keyword "\<leftharpoondown>"} || @{keyword "<="}) >> K Syntax.Print_Rule ||
 | |
| 297 |     (@{keyword "\<rightleftharpoons>"} || @{keyword "=="}) >> K Syntax.Parse_Print_Rule) toks;
 | |
| 298 | ||
| 299 | val trans_line = | |
| 300 | trans_pat -- Parse.!!! (trans_arrow -- trans_pat) | |
| 301 | >> (fn (left, (arr, right)) => arr (left, right)); | |
| 302 | ||
| 303 | val _ = | |
| 304 |   Outer_Syntax.command @{command_keyword translations} "add syntax translation rules"
 | |
| 305 | (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.translations)); | |
| 306 | ||
| 307 | val _ = | |
| 308 |   Outer_Syntax.command @{command_keyword no_translations} "delete syntax translation rules"
 | |
| 309 | (Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.no_translations)); | |
| 310 | ||
| 311 | in end\<close> | |
| 312 | ||
| 313 | ||
| 314 | subsubsection \<open>Translation functions\<close> | |
| 315 | ||
| 316 | ML \<open> | |
| 317 | local | |
| 62849 
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 _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 320 |   Outer_Syntax.command @{command_keyword parse_ast_translation}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 321 | "install parse ast translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 322 | (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 | 323 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 324 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 325 |   Outer_Syntax.command @{command_keyword parse_translation}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 326 | "install parse translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 327 | (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 | 328 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 329 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 330 |   Outer_Syntax.command @{command_keyword print_translation}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 331 | "install print translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 332 | (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 | 333 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 334 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 335 |   Outer_Syntax.command @{command_keyword typed_print_translation}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 336 | "install typed print translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 337 | (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 | 338 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 339 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 340 |   Outer_Syntax.command @{command_keyword print_ast_translation}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 341 | "install print ast translation functions" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 342 | (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 | 343 | |
| 62856 | 344 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 345 | |
| 62856 | 346 | |
| 347 | subsubsection \<open>Specifications\<close> | |
| 348 | ||
| 349 | ML \<open> | |
| 350 | local | |
| 351 | ||
| 352 | val _ = | |
| 353 |   Outer_Syntax.local_theory' @{command_keyword definition} "constant definition"
 | |
| 63180 | 354 | (Scan.option Parse_Spec.constdecl -- (Parse_Spec.opt_thm_name ":" -- Parse.prop) -- | 
| 355 | Parse_Spec.if_assumes -- Parse.for_fixes >> (fn (((decl, spec), prems), params) => | |
| 356 | #2 oo Specification.definition_cmd decl params prems spec)); | |
| 62856 | 357 | |
| 358 | val _ = | |
| 359 |   Outer_Syntax.local_theory' @{command_keyword abbreviation} "constant abbreviation"
 | |
| 63180 | 360 | (Parse.syntax_mode -- Scan.option Parse_Spec.constdecl -- Parse.prop -- Parse.for_fixes | 
| 361 | >> (fn (((mode, decl), spec), params) => Specification.abbreviation_cmd mode decl params spec)); | |
| 62856 | 362 | |
| 63178 | 363 | val axiomatization = | 
| 364 | Parse.and_list1 (Parse_Spec.thm_name ":" -- Parse.prop) -- | |
| 365 | Parse_Spec.if_assumes -- Parse.for_fixes >> (fn ((a, b), c) => (c, b, a)); | |
| 366 | ||
| 62856 | 367 | val _ = | 
| 368 |   Outer_Syntax.command @{command_keyword axiomatization} "axiomatic constant specification"
 | |
| 63285 | 369 | (Scan.optional Parse.vars [] -- | 
| 63178 | 370 | Scan.optional (Parse.where_ |-- Parse.!!! axiomatization) ([], [], []) | 
| 371 | >> (fn (a, (b, c, d)) => Toplevel.theory (#2 o Specification.axiomatization_cmd a b c d))); | |
| 62856 | 372 | |
| 373 | in end\<close> | |
| 374 | ||
| 375 | ||
| 376 | subsubsection \<open>Notation\<close> | |
| 377 | ||
| 378 | ML \<open> | |
| 379 | local | |
| 380 | ||
| 381 | val _ = | |
| 382 |   Outer_Syntax.local_theory @{command_keyword type_notation}
 | |
| 383 | "add concrete syntax for type constructors" | |
| 384 | (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) | |
| 385 | >> (fn (mode, args) => Specification.type_notation_cmd true mode args)); | |
| 386 | ||
| 387 | val _ = | |
| 388 |   Outer_Syntax.local_theory @{command_keyword no_type_notation}
 | |
| 389 | "delete concrete syntax for type constructors" | |
| 390 | (Parse.syntax_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix) | |
| 391 | >> (fn (mode, args) => Specification.type_notation_cmd false mode args)); | |
| 392 | ||
| 393 | val _ = | |
| 394 |   Outer_Syntax.local_theory @{command_keyword notation}
 | |
| 395 | "add concrete syntax for constants / fixed variables" | |
| 396 | (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) | |
| 397 | >> (fn (mode, args) => Specification.notation_cmd true mode args)); | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 398 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 399 | val _ = | 
| 62856 | 400 |   Outer_Syntax.local_theory @{command_keyword no_notation}
 | 
| 401 | "delete concrete syntax for constants / fixed variables" | |
| 402 | (Parse.syntax_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix) | |
| 403 | >> (fn (mode, args) => Specification.notation_cmd false mode args)); | |
| 404 | ||
| 405 | in end\<close> | |
| 406 | ||
| 407 | ||
| 408 | subsubsection \<open>Theorems\<close> | |
| 409 | ||
| 410 | ML \<open> | |
| 411 | local | |
| 412 | ||
| 63094 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 413 | val long_keyword = | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 414 | Parse_Spec.includes >> K "" || | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 415 | Parse_Spec.long_statement_keyword; | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 416 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 417 | val long_statement = | 
| 63352 | 418 | 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 | 419 | Scan.optional Parse_Spec.includes [] -- Parse_Spec.long_statement | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 420 | >> (fn ((binding, includes), (elems, concl)) => (true, binding, includes, elems, concl)); | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 421 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 422 | val short_statement = | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 423 | Parse_Spec.statement -- Parse_Spec.if_statement -- Parse.for_fixes | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 424 | >> (fn ((shows, assumes), fixes) => | 
| 63352 | 425 | (false, Binding.empty_atts, [], [Element.Fixes fixes, Element.Assumes assumes], | 
| 63094 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 426 | Element.Shows shows)); | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 427 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 428 | fun theorem spec schematic descr = | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 429 |   Outer_Syntax.local_theory_to_proof' spec ("state " ^ descr)
 | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 430 | ((long_statement || short_statement) >> (fn (long, binding, includes, elems, concl) => | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 431 | ((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd) | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 432 | long Thm.theoremK NONE (K I) binding includes elems concl))); | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 433 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 434 | val _ = theorem @{command_keyword theorem} false "theorem";
 | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 435 | val _ = theorem @{command_keyword lemma} false "lemma";
 | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 436 | val _ = theorem @{command_keyword corollary} false "corollary";
 | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 437 | val _ = theorem @{command_keyword proposition} false "proposition";
 | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 438 | val _ = theorem @{command_keyword schematic_goal} true "schematic goal";
 | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 439 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 440 | in end\<close> | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 441 | |
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 442 | ML \<open> | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 443 | local | 
| 
056ea294c256
toplevel theorem statements support 'if'/'for' eigen-context;
 wenzelm parents: 
63064diff
changeset | 444 | |
| 62856 | 445 | val _ = | 
| 446 |   Outer_Syntax.local_theory' @{command_keyword lemmas} "define theorems"
 | |
| 447 | (Parse_Spec.name_facts -- Parse.for_fixes >> | |
| 448 | (fn (facts, fixes) => #2 oo Specification.theorems_cmd Thm.theoremK facts fixes)); | |
| 449 | ||
| 450 | val _ = | |
| 451 |   Outer_Syntax.local_theory' @{command_keyword declare} "declare theorems"
 | |
| 62969 | 452 | (Parse.and_list1 Parse.thms1 -- Parse.for_fixes | 
| 62856 | 453 | >> (fn (facts, fixes) => | 
| 63352 | 454 | #2 oo Specification.theorems_cmd "" [(Binding.empty_atts, flat facts)] fixes)); | 
| 62856 | 455 | |
| 456 | val _ = | |
| 457 |   Outer_Syntax.local_theory @{command_keyword named_theorems}
 | |
| 458 | "declare named collection of theorems" | |
| 459 | (Parse.and_list1 (Parse.binding -- Scan.optional Parse.text "") >> | |
| 460 | fold (fn (b, descr) => snd o Named_Theorems.declare b descr)); | |
| 461 | ||
| 462 | in end\<close> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 463 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 464 | |
| 62856 | 465 | subsubsection \<open>Hide names\<close> | 
| 466 | ||
| 467 | ML \<open> | |
| 468 | local | |
| 469 | ||
| 470 | fun hide_names command_keyword what hide parse prep = | |
| 471 |   Outer_Syntax.command command_keyword ("hide " ^ what ^ " from name space")
 | |
| 472 | ((Parse.opt_keyword "open" >> not) -- Scan.repeat1 parse >> (fn (fully, args) => | |
| 473 | (Toplevel.theory (fn thy => | |
| 474 | let val ctxt = Proof_Context.init_global thy | |
| 475 | in fold (hide fully o prep ctxt) args thy end)))); | |
| 476 | ||
| 477 | val _ = | |
| 478 |   hide_names @{command_keyword hide_class} "classes" Sign.hide_class Parse.class
 | |
| 479 | Proof_Context.read_class; | |
| 480 | ||
| 481 | val _ = | |
| 482 |   hide_names @{command_keyword hide_type} "types" Sign.hide_type Parse.type_const
 | |
| 483 |     ((#1 o dest_Type) oo Proof_Context.read_type_name {proper = true, strict = false});
 | |
| 484 | ||
| 485 | val _ = | |
| 486 |   hide_names @{command_keyword hide_const} "consts" Sign.hide_const Parse.const
 | |
| 487 |     ((#1 o dest_Const) oo Proof_Context.read_const {proper = true, strict = false});
 | |
| 488 | ||
| 489 | val _ = | |
| 490 |   hide_names @{command_keyword hide_fact} "facts" Global_Theory.hide_fact
 | |
| 62969 | 491 | (Parse.position Parse.name) (Global_Theory.check_fact o Proof_Context.theory_of); | 
| 62856 | 492 | |
| 493 | in end\<close> | |
| 494 | ||
| 495 | ||
| 496 | subsection \<open>Bundled declarations\<close> | |
| 497 | ||
| 498 | ML \<open> | |
| 499 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 500 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 501 | val _ = | 
| 63273 | 502 |   Outer_Syntax.maybe_begin_local_theory @{command_keyword bundle}
 | 
| 503 | "define bundle of declarations" | |
| 62969 | 504 |     ((Parse.binding --| @{keyword "="}) -- Parse.thms1 -- Parse.for_fixes
 | 
| 63273 | 505 | >> (uncurry Bundle.bundle_cmd)) | 
| 506 | (Parse.binding --| Parse.begin >> Bundle.init); | |
| 63270 | 507 | |
| 508 | val _ = | |
| 63282 | 509 |   Outer_Syntax.local_theory @{command_keyword unbundle}
 | 
| 510 | "activate declarations from bundle in local theory" | |
| 511 | (Scan.repeat1 (Parse.position Parse.name) >> Bundle.unbundle_cmd); | |
| 512 | ||
| 513 | val _ = | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 514 |   Outer_Syntax.command @{command_keyword include}
 | 
| 63282 | 515 | "activate declarations from bundle in proof body" | 
| 62969 | 516 | (Scan.repeat1 (Parse.position Parse.name) >> (Toplevel.proof o Bundle.include_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 517 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 518 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 519 |   Outer_Syntax.command @{command_keyword including}
 | 
| 63282 | 520 | "activate declarations from bundle in goal refinement" | 
| 62969 | 521 | (Scan.repeat1 (Parse.position Parse.name) >> (Toplevel.proof o Bundle.including_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 522 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 523 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 524 |   Outer_Syntax.command @{command_keyword print_bundles}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 525 | "print bundles of declarations" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 526 | (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 | 527 | |
| 62856 | 528 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 529 | |
| 62856 | 530 | |
| 531 | subsection \<open>Local theory specifications\<close> | |
| 532 | ||
| 533 | subsubsection \<open>Specification context\<close> | |
| 534 | ||
| 535 | ML \<open> | |
| 536 | local | |
| 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 _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 539 |   Outer_Syntax.command @{command_keyword context} "begin local theory context"
 | 
| 62969 | 540 | ((Parse.position Parse.name >> (fn name => | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 541 | Toplevel.begin_local_theory true (Named_Target.begin name)) || | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 542 | Scan.optional Parse_Spec.includes [] -- Scan.repeat Parse_Spec.context_element | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 543 | >> (fn (incls, elems) => Toplevel.open_target (#2 o Bundle.context_cmd incls elems))) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 544 | --| Parse.begin); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 545 | |
| 62856 | 546 | val _ = | 
| 547 |   Outer_Syntax.command @{command_keyword end} "end context"
 | |
| 548 | (Scan.succeed | |
| 549 | (Toplevel.exit o Toplevel.end_local_theory o Toplevel.close_target o | |
| 550 | Toplevel.end_proof (K Proof.end_notepad))); | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 551 | |
| 62856 | 552 | in end\<close> | 
| 553 | ||
| 554 | ||
| 555 | subsubsection \<open>Locales and interpretation\<close> | |
| 556 | ||
| 557 | ML \<open> | |
| 558 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 559 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 560 | val locale_val = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 561 | Parse_Spec.locale_expression -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 562 |     Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 563 | Scan.repeat1 Parse_Spec.context_element >> pair ([], []); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 564 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 565 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 566 |   Outer_Syntax.command @{command_keyword locale} "define named specification context"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 567 | (Parse.binding -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 568 |       Scan.optional (@{keyword "="} |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 569 | >> (fn ((name, (expr, elems)), begin) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 570 | Toplevel.begin_local_theory begin | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 571 | (Expression.add_locale_cmd name Binding.empty expr elems #> snd))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 572 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 573 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 574 |   Outer_Syntax.command @{command_keyword experiment} "open private specification context"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 575 | (Scan.repeat Parse_Spec.context_element --| Parse.begin | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 576 | >> (fn elems => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 577 | Toplevel.begin_local_theory true (Experiment.experiment_cmd elems #> snd))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 578 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 579 | val interpretation_args = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 580 | Parse.!!! Parse_Spec.locale_expression -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 581 | Scan.optional | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 582 |       (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) [];
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 583 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 584 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 585 |   Outer_Syntax.command @{command_keyword interpret}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 586 | "prove interpretation of locale expression in proof context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 587 | (interpretation_args >> (fn (expr, equations) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 588 | Toplevel.proof (Interpretation.interpret_cmd expr equations))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 589 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 590 | val interpretation_args_with_defs = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 591 | Parse.!!! Parse_Spec.locale_expression -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 592 |     (Scan.optional (@{keyword "defines"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 593 |       -- ((Parse.binding -- Parse.opt_mixfix') --| @{keyword "="} -- Parse.term))) [] --
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 594 | Scan.optional | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 595 |       (@{keyword "rewrites"} |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) []);
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 596 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 597 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 598 |   Outer_Syntax.local_theory_to_proof @{command_keyword global_interpretation}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 599 | "prove interpretation of locale expression into global theory" | 
| 62856 | 600 | (interpretation_args_with_defs >> (fn (expr, (defs, equations)) => | 
| 601 | Interpretation.global_interpretation_cmd expr defs equations)); | |
| 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 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 604 |   Outer_Syntax.command @{command_keyword sublocale}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 605 | "prove sublocale relation between a locale and a locale expression" | 
| 62969 | 606 |     ((Parse.position Parse.name --| (@{keyword "\<subseteq>"} || @{keyword "<"}) --
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 607 | interpretation_args_with_defs >> (fn (loc, (expr, (defs, equations))) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 608 | Toplevel.theory_to_proof (Interpretation.global_sublocale_cmd loc expr defs equations))) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 609 | || interpretation_args_with_defs >> (fn (expr, (defs, equations)) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 610 | Toplevel.local_theory_to_proof NONE NONE (Interpretation.sublocale_cmd expr defs equations))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 611 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 612 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 613 |   Outer_Syntax.command @{command_keyword interpretation}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 614 | "prove interpretation of locale expression in local theory or into global theory" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 615 | (interpretation_args >> (fn (expr, equations) => | 
| 62856 | 616 | Toplevel.local_theory_to_proof NONE NONE | 
| 617 | (Interpretation.isar_interpretation_cmd expr equations))); | |
| 618 | ||
| 619 | in end\<close> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 620 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 621 | |
| 62856 | 622 | subsubsection \<open>Type classes\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 623 | |
| 62856 | 624 | ML \<open> | 
| 625 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 626 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 627 | val class_val = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 628 | Parse_Spec.class_expression -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 629 |     Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 630 | Scan.repeat1 Parse_Spec.context_element >> pair []; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 631 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 632 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 633 |   Outer_Syntax.command @{command_keyword class} "define type class"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 634 |    (Parse.binding -- Scan.optional (@{keyword "="} |-- class_val) ([], []) -- Parse.opt_begin
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 635 | >> (fn ((name, (supclasses, elems)), begin) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 636 | Toplevel.begin_local_theory begin | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 637 | (Class_Declaration.class_cmd name supclasses elems #> snd))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 638 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 639 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 640 |   Outer_Syntax.local_theory_to_proof @{command_keyword subclass} "prove a subclass relation"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 641 | (Parse.class >> Class_Declaration.subclass_cmd); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 642 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 643 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 644 |   Outer_Syntax.command @{command_keyword instantiation} "instantiate and prove type arity"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 645 | (Parse.multi_arity --| Parse.begin | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 646 | >> (fn arities => Toplevel.begin_local_theory true (Class.instantiation_cmd arities))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 647 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 648 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 649 |   Outer_Syntax.command @{command_keyword instance} "prove type arity or subclass relation"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 650 | ((Parse.class -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 651 |     ((@{keyword "\<subseteq>"} || @{keyword "<"}) |-- Parse.!!! Parse.class) >> Class.classrel_cmd ||
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 652 | 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 | 653 | 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 | 654 | |
| 62856 | 655 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 656 | |
| 62856 | 657 | |
| 658 | subsubsection \<open>Arbitrary overloading\<close> | |
| 659 | ||
| 660 | ML \<open> | |
| 661 | local | |
| 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 _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 664 |   Outer_Syntax.command @{command_keyword overloading} "overloaded definitions"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 665 |    (Scan.repeat1 (Parse.name --| (@{keyword "=="} || @{keyword "\<equiv>"}) -- Parse.term --
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 666 |       Scan.optional (@{keyword "("} |-- (@{keyword "unchecked"} >> K false) --| @{keyword ")"}) true
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 667 | >> Scan.triple1) --| Parse.begin | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 668 | >> (fn operations => Toplevel.begin_local_theory true (Overloading.overloading_cmd operations))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 669 | |
| 62856 | 670 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 671 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 672 | |
| 62856 | 673 | subsection \<open>Proof commands\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 674 | |
| 62856 | 675 | ML \<open> | 
| 676 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 677 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 678 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 679 |   Outer_Syntax.local_theory_to_proof @{command_keyword notepad} "begin proof context"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 680 | (Parse.begin >> K Proof.begin_notepad); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 681 | |
| 62856 | 682 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 683 | |
| 62856 | 684 | |
| 685 | subsubsection \<open>Statements\<close> | |
| 686 | ||
| 687 | ML \<open> | |
| 688 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 689 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 690 | val structured_statement = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 691 | 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 | 692 | >> (fn ((shows, (strict, assumes)), fixes) => (strict, fixes, assumes, shows)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 693 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 694 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 695 |   Outer_Syntax.command @{command_keyword have} "state local goal"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 696 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 697 | 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 | 698 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 699 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 700 |   Outer_Syntax.command @{command_keyword show} "state local goal, to refine pending subgoals"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 701 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 702 | 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 | 703 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 704 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 705 |   Outer_Syntax.command @{command_keyword hence} "old-style alias of \"then have\""
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 706 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 707 | 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 | 708 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 709 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 710 |   Outer_Syntax.command @{command_keyword thus} "old-style alias of  \"then show\""
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 711 | (structured_statement >> (fn (a, b, c, d) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 712 | 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 | 713 | |
| 62856 | 714 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 715 | |
| 62856 | 716 | |
| 717 | subsubsection \<open>Local facts\<close> | |
| 718 | ||
| 719 | ML \<open> | |
| 720 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 721 | |
| 62969 | 722 | val facts = Parse.and_list1 Parse.thms1; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 723 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 724 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 725 |   Outer_Syntax.command @{command_keyword then} "forward chaining"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 726 | (Scan.succeed (Toplevel.proof Proof.chain)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 727 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 728 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 729 |   Outer_Syntax.command @{command_keyword from} "forward chaining from given facts"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 730 | (facts >> (Toplevel.proof o Proof.from_thmss_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 731 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 732 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 733 |   Outer_Syntax.command @{command_keyword with} "forward chaining from given and current facts"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 734 | (facts >> (Toplevel.proof o Proof.with_thmss_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 735 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 736 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 737 |   Outer_Syntax.command @{command_keyword note} "define facts"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 738 | (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 | 739 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 740 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 741 |   Outer_Syntax.command @{command_keyword supply} "define facts during goal refinement (unstructured)"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 742 | (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 | 743 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 744 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 745 |   Outer_Syntax.command @{command_keyword using} "augment goal facts"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 746 | (facts >> (Toplevel.proof o Proof.using_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 747 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 748 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 749 |   Outer_Syntax.command @{command_keyword unfolding} "unfold definitions in goal and facts"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 750 | (facts >> (Toplevel.proof o Proof.unfolding_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 751 | |
| 62856 | 752 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 753 | |
| 62856 | 754 | |
| 755 | subsubsection \<open>Proof context\<close> | |
| 756 | ||
| 757 | ML \<open> | |
| 758 | local | |
| 759 | ||
| 760 | val structured_statement = | |
| 761 | Parse_Spec.statement -- Parse_Spec.if_statement' -- Parse.for_fixes | |
| 762 | >> (fn ((shows, assumes), fixes) => (fixes, assumes, shows)); | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 763 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 764 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 765 |   Outer_Syntax.command @{command_keyword fix} "fix local variables (Skolem constants)"
 | 
| 63285 | 766 | (Parse.vars >> (Toplevel.proof o Proof.fix_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 767 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 768 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 769 |   Outer_Syntax.command @{command_keyword assume} "assume propositions"
 | 
| 62856 | 770 | (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 | 771 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 772 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 773 |   Outer_Syntax.command @{command_keyword presume} "assume propositions, to be established later"
 | 
| 62856 | 774 | (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 | 775 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 776 | val _ = | 
| 63039 | 777 |   Outer_Syntax.command @{command_keyword define} "local definition (non-polymorphic)"
 | 
| 63285 | 778 | ((Parse.vars --| Parse.where_) -- Parse_Spec.statement -- Parse.for_fixes | 
| 63039 | 779 | >> (fn ((a, b), c) => Toplevel.proof (Proof.define_cmd a c b))); | 
| 780 | ||
| 781 | val _ = | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 782 |   Outer_Syntax.command @{command_keyword def} "local definition (non-polymorphic)"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 783 | (Parse.and_list1 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 784 | (Parse_Spec.opt_thm_name ":" -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 785 | ((Parse.binding -- Parse.opt_mixfix) -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 786 |           ((@{keyword "\<equiv>"} || @{keyword "=="}) |-- Parse.!!! Parse.termp)))
 | 
| 63043 | 787 | >> (fn args => Toplevel.proof (fn state => | 
| 788 | (legacy_feature "Old 'def' command -- use 'define' instead"; Proof.def_cmd args state)))); | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 789 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 790 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 791 |   Outer_Syntax.command @{command_keyword consider} "state cases rule"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 792 | (Parse_Spec.obtains >> (Toplevel.proof' o Obtain.consider_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 793 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 794 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 795 |   Outer_Syntax.command @{command_keyword obtain} "generalized elimination"
 | 
| 63285 | 796 | (Parse.parbinding -- Scan.optional (Parse.vars --| Parse.where_) [] -- structured_statement | 
| 63059 
3f577308551e
'obtain' supports structured statements (similar to 'define');
 wenzelm parents: 
63043diff
changeset | 797 | >> (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 | 798 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 799 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 800 |   Outer_Syntax.command @{command_keyword guess} "wild guessing (unstructured)"
 | 
| 63285 | 801 | (Scan.optional Parse.vars [] >> (Toplevel.proof' o Obtain.guess_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 802 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 803 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 804 |   Outer_Syntax.command @{command_keyword let} "bind text variables"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 805 |     (Parse.and_list1 (Parse.and_list1 Parse.term -- (@{keyword "="} |-- Parse.term))
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 806 | >> (Toplevel.proof o Proof.let_bind_cmd)); | 
| 
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 _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 809 |   Outer_Syntax.command @{command_keyword write} "add concrete syntax for constants / fixed variables"
 | 
| 62856 | 810 | (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 | 811 | >> (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 | 812 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 813 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 814 |   Outer_Syntax.command @{command_keyword case} "invoke local context"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 815 | (Parse_Spec.opt_thm_name ":" -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 816 |       (@{keyword "("} |--
 | 
| 62969 | 817 | Parse.!!! (Parse.position Parse.name -- Scan.repeat (Parse.maybe Parse.binding) | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 818 |           --| @{keyword ")"}) ||
 | 
| 62969 | 819 | Parse.position Parse.name >> rpair []) >> (Toplevel.proof o Proof.case_cmd)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 820 | |
| 62856 | 821 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 822 | |
| 62856 | 823 | |
| 824 | subsubsection \<open>Proof structure\<close> | |
| 825 | ||
| 826 | ML \<open> | |
| 827 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 828 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 829 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 830 |   Outer_Syntax.command @{command_keyword "{"} "begin explicit proof block"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 831 | (Scan.succeed (Toplevel.proof Proof.begin_block)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 832 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 833 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 834 |   Outer_Syntax.command @{command_keyword "}"} "end explicit proof block"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 835 | (Scan.succeed (Toplevel.proof Proof.end_block)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 836 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 837 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 838 |   Outer_Syntax.command @{command_keyword next} "enter next proof block"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 839 | (Scan.succeed (Toplevel.proof Proof.next_block)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 840 | |
| 62856 | 841 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 842 | |
| 62856 | 843 | |
| 844 | subsubsection \<open>End proof\<close> | |
| 845 | ||
| 846 | ML \<open> | |
| 847 | local | |
| 62849 
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 _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 850 |   Outer_Syntax.command @{command_keyword qed} "conclude proof"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 851 | (Scan.option Method.parse >> (fn m => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 852 | (Option.map Method.report m; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 853 | Isar_Cmd.qed m))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 854 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 855 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 856 |   Outer_Syntax.command @{command_keyword by} "terminal backward proof"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 857 | (Method.parse -- Scan.option Method.parse >> (fn (m1, m2) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 858 | (Method.report m1; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 859 | Option.map Method.report m2; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 860 | Isar_Cmd.terminal_proof (m1, m2)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 861 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 862 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 863 |   Outer_Syntax.command @{command_keyword ".."} "default proof"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 864 | (Scan.succeed Isar_Cmd.default_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 865 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 866 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 867 |   Outer_Syntax.command @{command_keyword "."} "immediate proof"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 868 | (Scan.succeed Isar_Cmd.immediate_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 869 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 870 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 871 |   Outer_Syntax.command @{command_keyword done} "done proof"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 872 | (Scan.succeed Isar_Cmd.done_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 873 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 874 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 875 |   Outer_Syntax.command @{command_keyword sorry} "skip proof (quick-and-dirty mode only!)"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 876 | (Scan.succeed Isar_Cmd.skip_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 877 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 878 | val _ = | 
| 63342 | 879 |   Outer_Syntax.command @{command_keyword \<proof>} "dummy proof (quick-and-dirty mode only!)"
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 880 | (Scan.succeed Isar_Cmd.skip_proof); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 881 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 882 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 883 |   Outer_Syntax.command @{command_keyword oops} "forget proof"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 884 | (Scan.succeed (Toplevel.forget_proof true)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 885 | |
| 62856 | 886 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 887 | |
| 62856 | 888 | |
| 889 | subsubsection \<open>Proof steps\<close> | |
| 890 | ||
| 891 | ML \<open> | |
| 892 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 893 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 894 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 895 |   Outer_Syntax.command @{command_keyword defer} "shuffle internal proof state"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 896 | (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 | 897 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 898 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 899 |   Outer_Syntax.command @{command_keyword prefer} "shuffle internal proof state"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 900 | (Parse.nat >> (Toplevel.proof o Proof.prefer)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 901 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 902 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 903 |   Outer_Syntax.command @{command_keyword apply} "initial goal refinement step (unstructured)"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 904 | (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 | 905 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 906 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 907 |   Outer_Syntax.command @{command_keyword apply_end} "terminal goal refinement step (unstructured)"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 908 | (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 | 909 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 910 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 911 |   Outer_Syntax.command @{command_keyword proof} "backward proof step"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 912 | (Scan.option Method.parse >> (fn m => | 
| 63513 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 913 | (Option.map Method.report m; | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 914 | Toplevel.proof (fn state => | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 915 | let | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 916 | val state' = state |> Proof.proof m |> Seq.the_result ""; | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 917 | val _ = | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 918 | Output.information | 
| 
9f8d06f23c09
information about proof outline with cases (sendback);
 wenzelm parents: 
63510diff
changeset | 919 | (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 | 920 | in state' end)))) | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 921 | |
| 62856 | 922 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 923 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 924 | |
| 62856 | 925 | subsubsection \<open>Subgoal focus\<close> | 
| 926 | ||
| 927 | ML \<open> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 928 | local | 
| 
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 opt_fact_binding = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 931 | Scan.optional (Parse.binding -- Parse.opt_attribs || Parse.attribs >> pair Binding.empty) | 
| 63352 | 932 | Binding.empty_atts; | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 933 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 934 | val for_params = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 935 | Scan.optional | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 936 |     (@{keyword "for"} |--
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 937 | Parse.!!! ((Scan.option Parse.dots >> is_some) -- | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 938 | (Scan.repeat1 (Parse.position (Parse.maybe Parse.name))))) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 939 | (false, []); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 940 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 941 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 942 |   Outer_Syntax.command @{command_keyword subgoal}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 943 | "focus on first subgoal within backward refinement" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 944 |     (opt_fact_binding -- (Scan.option (@{keyword "premises"} |-- Parse.!!! opt_fact_binding)) --
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 945 | for_params >> (fn ((a, b), c) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 946 | 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 | 947 | |
| 62856 | 948 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 949 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 950 | |
| 62856 | 951 | subsubsection \<open>Calculation\<close> | 
| 952 | ||
| 953 | ML \<open> | |
| 954 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 955 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 956 | val calculation_args = | 
| 62969 | 957 |   Scan.option (@{keyword "("} |-- Parse.!!! ((Parse.thms1 --| @{keyword ")"})));
 | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 958 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 959 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 960 |   Outer_Syntax.command @{command_keyword also} "combine calculation and current facts"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 961 | (calculation_args >> (Toplevel.proofs' o Calculation.also_cmd)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 962 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 963 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 964 |   Outer_Syntax.command @{command_keyword finally}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 965 | "combine calculation and current facts, exhibit result" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 966 | (calculation_args >> (Toplevel.proofs' o Calculation.finally_cmd)); | 
| 
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 _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 969 |   Outer_Syntax.command @{command_keyword moreover} "augment calculation by current facts"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 970 | (Scan.succeed (Toplevel.proof' Calculation.moreover)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 971 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 972 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 973 |   Outer_Syntax.command @{command_keyword ultimately}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 974 | "augment calculation by current facts, exhibit result" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 975 | (Scan.succeed (Toplevel.proof' Calculation.ultimately)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 976 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 977 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 978 |   Outer_Syntax.command @{command_keyword print_trans_rules} "print transitivity rules"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 979 | (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 | 980 | |
| 62856 | 981 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 982 | |
| 62856 | 983 | |
| 984 | subsubsection \<open>Proof navigation\<close> | |
| 985 | ||
| 986 | ML \<open> | |
| 987 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 988 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 989 | fun report_back () = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 990 | Output.report [Markup.markup Markup.bad "Explicit backtracking"]; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 991 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 992 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 993 |   Outer_Syntax.command @{command_keyword back} "explicit backtracking of proof command"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 994 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 995 | (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 | 996 | Toplevel.skip_proof report_back)); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 997 | |
| 62856 | 998 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 999 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1000 | |
| 62856 | 1001 | subsection \<open>Diagnostic commands (for interactive mode only)\<close> | 
| 1002 | ||
| 1003 | ML \<open> | |
| 1004 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1005 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1006 | val opt_modes = | 
| 62969 | 1007 |   Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.name --| @{keyword ")"})) [];
 | 
| 62849 
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 _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1010 |   Outer_Syntax.command @{command_keyword help}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1011 | "retrieve outer syntax commands according to name patterns" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1012 | (Scan.repeat Parse.name >> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1013 | (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 | 1014 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1015 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1016 |   Outer_Syntax.command @{command_keyword print_commands} "print outer syntax commands"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1017 | (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 | 1018 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1019 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1020 |   Outer_Syntax.command @{command_keyword print_options} "print configuration options"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1021 | (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 | 1022 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1023 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1024 |   Outer_Syntax.command @{command_keyword print_context}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1025 | "print context of local theory target" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1026 | (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 | 1027 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1028 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1029 |   Outer_Syntax.command @{command_keyword print_theory}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1030 | "print logical theory contents" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1031 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1032 | 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 | 1033 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1034 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1035 |   Outer_Syntax.command @{command_keyword print_definitions}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1036 | "print dependencies of definitional theory content" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1037 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1038 | 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 | 1039 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1040 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1041 |   Outer_Syntax.command @{command_keyword print_syntax}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1042 | "print inner syntax of context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1043 | (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 | 1044 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1045 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1046 |   Outer_Syntax.command @{command_keyword print_defn_rules}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1047 | "print definitional rewrite rules of context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1048 | (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 | 1049 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1050 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1051 |   Outer_Syntax.command @{command_keyword print_abbrevs}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1052 | "print constant abbreviations of context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1053 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1054 | 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 | 1055 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1056 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1057 |   Outer_Syntax.command @{command_keyword print_theorems}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1058 | "print theorems of local theory or proof context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1059 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1060 | 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 | 1061 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1062 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1063 |   Outer_Syntax.command @{command_keyword print_locales}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1064 | "print locales of this theory" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1065 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1066 | Toplevel.keep (Locale.print_locales b o Toplevel.theory_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1067 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1068 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1069 |   Outer_Syntax.command @{command_keyword print_classes}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1070 | "print classes of this theory" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1071 | (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 | 1072 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1073 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1074 |   Outer_Syntax.command @{command_keyword print_locale}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1075 | "print locale of this theory" | 
| 62969 | 1076 | (Parse.opt_bang -- Parse.position Parse.name >> (fn (b, name) => | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1077 | Toplevel.keep (fn state => Locale.print_locale (Toplevel.theory_of state) b name))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1078 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1079 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1080 |   Outer_Syntax.command @{command_keyword print_interps}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1081 | "print interpretations of locale for this theory or proof context" | 
| 62969 | 1082 | (Parse.position Parse.name >> (fn name => | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1083 | Toplevel.keep (fn state => Locale.print_registrations (Toplevel.context_of state) name))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1084 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1085 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1086 |   Outer_Syntax.command @{command_keyword print_dependencies}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1087 | "print dependencies of locale expression" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1088 | (Parse.opt_bang -- Parse_Spec.locale_expression >> (fn (b, expr) => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1089 | Toplevel.keep (fn state => Expression.print_dependencies (Toplevel.context_of state) b expr))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1090 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1091 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1092 |   Outer_Syntax.command @{command_keyword print_attributes}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1093 | "print attributes of this theory" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1094 | (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 | 1095 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1096 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1097 |   Outer_Syntax.command @{command_keyword print_simpset}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1098 | "print context of Simplifier" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1099 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1100 | 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 | 1101 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1102 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1103 |   Outer_Syntax.command @{command_keyword print_rules} "print intro/elim rules"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1104 | (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 | 1105 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1106 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1107 |   Outer_Syntax.command @{command_keyword print_methods} "print methods of this theory"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1108 | (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 | 1109 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1110 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1111 |   Outer_Syntax.command @{command_keyword print_antiquotations}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1112 | "print document antiquotations" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1113 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1114 | Toplevel.keep (Thy_Output.print_antiquotations b o Toplevel.context_of))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1115 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1116 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1117 |   Outer_Syntax.command @{command_keyword print_ML_antiquotations}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1118 | "print ML antiquotations" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1119 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1120 | 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 | 1121 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1122 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1123 |   Outer_Syntax.command @{command_keyword locale_deps} "visualize locale dependencies"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1124 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1125 | (Toplevel.keep (Toplevel.theory_of #> (fn thy => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1126 | Locale.pretty_locale_deps thy | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1127 |         |> map (fn {name, parents, body} =>
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1128 | ((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 | 1129 | |> Graph_Display.display_graph_old)))); | 
| 
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 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1132 |   Outer_Syntax.command @{command_keyword print_term_bindings}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1133 | "print term bindings of proof context" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1134 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1135 | (Toplevel.keep | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1136 | (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 | 1137 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1138 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1139 |   Outer_Syntax.command @{command_keyword print_facts} "print facts of proof context"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1140 | (Parse.opt_bang >> (fn b => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1141 | 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 | 1142 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1143 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1144 |   Outer_Syntax.command @{command_keyword print_cases} "print cases of proof context"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1145 | (Scan.succeed | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1146 | (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 | 1147 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1148 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1149 |   Outer_Syntax.command @{command_keyword print_statement}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1150 | "print theorems as long statements" | 
| 62969 | 1151 | (opt_modes -- Parse.thms1 >> Isar_Cmd.print_stmts); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1152 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1153 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1154 |   Outer_Syntax.command @{command_keyword thm} "print theorems"
 | 
| 62969 | 1155 | (opt_modes -- Parse.thms1 >> Isar_Cmd.print_thms); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1156 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1157 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1158 |   Outer_Syntax.command @{command_keyword prf} "print proof terms of theorems"
 | 
| 62969 | 1159 | (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 | 1160 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1161 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1162 |   Outer_Syntax.command @{command_keyword full_prf} "print full proof terms of theorems"
 | 
| 62969 | 1163 | (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 | 1164 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1165 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1166 |   Outer_Syntax.command @{command_keyword prop} "read and print proposition"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1167 | (opt_modes -- Parse.term >> Isar_Cmd.print_prop); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1168 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1169 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1170 |   Outer_Syntax.command @{command_keyword term} "read and print term"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1171 | (opt_modes -- Parse.term >> Isar_Cmd.print_term); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1172 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1173 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1174 |   Outer_Syntax.command @{command_keyword typ} "read and print type"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1175 |     (opt_modes -- (Parse.typ -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort))
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1176 | >> Isar_Cmd.print_type); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1177 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1178 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1179 |   Outer_Syntax.command @{command_keyword print_codesetup} "print code generator setup"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1180 | (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 | 1181 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1182 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1183 |   Outer_Syntax.command @{command_keyword print_state}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1184 | "print current proof state (if present)" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1185 | (opt_modes >> (fn modes => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1186 | Toplevel.keep (Print_Mode.with_modes modes (Output.state o Toplevel.string_of_state)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1187 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1188 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1189 |   Outer_Syntax.command @{command_keyword welcome} "print welcome message"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1190 | (Scan.succeed (Toplevel.keep (fn _ => writeln (Session.welcome ())))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1191 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1192 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1193 |   Outer_Syntax.command @{command_keyword display_drafts}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1194 | "display raw source files with symbols" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1195 | (Scan.repeat1 Parse.path >> (fn names => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1196 | Toplevel.keep (fn _ => ignore (Present.display_drafts (map Path.explode names))))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1197 | |
| 62856 | 1198 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1199 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1200 | |
| 62856 | 1201 | subsection \<open>Dependencies\<close> | 
| 1202 | ||
| 1203 | ML \<open> | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1204 | local | 
| 62856 | 1205 | |
| 1206 | val theory_bounds = | |
| 62969 | 1207 | Parse.position Parse.theory_name >> single || | 
| 1208 |   (@{keyword "("} |-- Parse.enum "|" (Parse.position Parse.theory_name) --| @{keyword ")"});
 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1209 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1210 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1211 |   Outer_Syntax.command @{command_keyword thy_deps} "visualize theory dependencies"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1212 | (Scan.option theory_bounds -- Scan.option theory_bounds >> | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1213 | (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 | 1214 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1215 | |
| 62856 | 1216 | val class_bounds = | 
| 1217 | Parse.sort >> single || | |
| 1218 |   (@{keyword "("} |-- Parse.enum "|" Parse.sort --| @{keyword ")"});
 | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1219 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1220 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1221 |   Outer_Syntax.command @{command_keyword class_deps} "visualize class dependencies"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1222 | (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 | 1223 | 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 | 1224 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1225 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1226 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1227 |   Outer_Syntax.command @{command_keyword thm_deps} "visualize theorem dependencies"
 | 
| 62969 | 1228 | (Parse.thms1 >> (fn args => | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1229 | Toplevel.keep (fn st => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1230 | Thm_Deps.thm_deps (Toplevel.theory_of st) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1231 | (Attrib.eval_thms (Toplevel.context_of st) args)))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1232 | |
| 62856 | 1233 | |
| 1234 | val thy_names = | |
| 62969 | 1235 | Scan.repeat1 (Scan.unless Parse.minus (Parse.position Parse.theory_name)); | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1236 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1237 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1238 |   Outer_Syntax.command @{command_keyword unused_thms} "find unused theorems"
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1239 | (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 | 1240 | Toplevel.keep (fn st => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1241 | let | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1242 | val thy = Toplevel.theory_of st; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1243 | val ctxt = Toplevel.context_of st; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1244 | fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1245 | val check = Theory.check ctxt; | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1246 | in | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1247 | Thm_Deps.unused_thms | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1248 | (case opt_range of | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1249 | NONE => (Theory.parents_of thy, [thy]) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1250 | | SOME (xs, NONE) => (map check xs, [thy]) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1251 | | SOME (xs, SOME ys) => (map check xs, map check ys)) | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1252 | |> map pretty_thm |> Pretty.writeln_chunks | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1253 | end))); | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1254 | |
| 62856 | 1255 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1256 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1257 | |
| 62856 | 1258 | subsubsection \<open>Find consts and theorems\<close> | 
| 1259 | ||
| 1260 | ML \<open> | |
| 1261 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1262 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1263 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1264 |   Outer_Syntax.command @{command_keyword find_consts}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1265 | "find constants by name / type patterns" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1266 | (Find_Consts.query_parser >> (fn spec => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1267 | Toplevel.keep (fn st => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1268 | 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 | 1269 | |
| 62856 | 1270 | val options = | 
| 1271 | Scan.optional | |
| 1272 |     (Parse.$$$ "(" |--
 | |
| 1273 | Parse.!!! (Scan.option Parse.nat -- | |
| 1274 | Scan.optional (Parse.reserved "with_dups" >> K false) true --| Parse.$$$ ")")) | |
| 1275 | (NONE, true); | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1276 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1277 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1278 |   Outer_Syntax.command @{command_keyword find_theorems}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1279 | "find theorems meeting specified criteria" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1280 | (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 | 1281 | Toplevel.keep (fn st => | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1282 | Pretty.writeln | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1283 | (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 | 1284 | |
| 62856 | 1285 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1286 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1287 | |
| 62856 | 1288 | subsection \<open>Code generation\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1289 | |
| 62856 | 1290 | ML \<open> | 
| 1291 | local | |
| 1292 | ||
| 1293 | val _ = | |
| 1294 |   Outer_Syntax.command @{command_keyword code_datatype}
 | |
| 1295 | "define set of code datatype constructors" | |
| 1296 | (Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd)); | |
| 1297 | ||
| 1298 | in end\<close> | |
| 1299 | ||
| 1300 | ||
| 1301 | subsection \<open>Extraction of programs from proofs\<close> | |
| 1302 | ||
| 1303 | ML \<open> | |
| 1304 | local | |
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1305 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1306 | 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 | 1307 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1308 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1309 |   Outer_Syntax.command @{command_keyword realizers}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1310 | "specify realizers for primitive axioms / theorems, together with correctness proof" | 
| 62969 | 1311 | (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 | 1312 | (fn xs => Toplevel.theory (fn thy => Extraction.add_realizers | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1313 | (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 | 1314 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1315 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1316 |   Outer_Syntax.command @{command_keyword realizability}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1317 | "add equations characterizing realizability" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1318 | (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 | 1319 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1320 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1321 |   Outer_Syntax.command @{command_keyword extract_type}
 | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1322 | "add equations characterizing type of extracted program" | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1323 | (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 | 1324 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1325 | val _ = | 
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1326 |   Outer_Syntax.command @{command_keyword extract} "extract terms from proofs"
 | 
| 62969 | 1327 | (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 | 1328 | 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 | 1329 | |
| 62856 | 1330 | in end\<close> | 
| 62849 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1331 | |
| 
caaa2fc4040d
clarified bootstrap -- avoid 'ML_file' in Pure.thy for uniformity;
 wenzelm parents: 
62848diff
changeset | 1332 | |
| 62944 
3ee643c5ed00
more standard session build process, including browser_info;
 wenzelm parents: 
62913diff
changeset | 1333 | section \<open>Auxiliary lemmas\<close> | 
| 20627 | 1334 | |
| 58611 | 1335 | subsection \<open>Meta-level connectives in assumptions\<close> | 
| 15803 | 1336 | |
| 1337 | lemma meta_mp: | |
| 58612 | 1338 | assumes "PROP P \<Longrightarrow> PROP Q" and "PROP P" | 
| 15803 | 1339 | shows "PROP Q" | 
| 58612 | 1340 | by (rule \<open>PROP P \<Longrightarrow> PROP Q\<close> [OF \<open>PROP P\<close>]) | 
| 15803 | 1341 | |
| 23432 | 1342 | lemmas meta_impE = meta_mp [elim_format] | 
| 1343 | ||
| 15803 | 1344 | lemma meta_spec: | 
| 58612 | 1345 | assumes "\<And>x. PROP P x" | 
| 26958 | 1346 | shows "PROP P x" | 
| 58612 | 1347 | by (rule \<open>\<And>x. PROP P x\<close>) | 
| 15803 | 1348 | |
| 1349 | lemmas meta_allE = meta_spec [elim_format] | |
| 1350 | ||
| 26570 | 1351 | lemma swap_params: | 
| 58612 | 1352 | "(\<And>x y. PROP P x y) \<equiv> (\<And>y x. PROP P x y)" .. | 
| 26570 | 1353 | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1354 | |
| 58611 | 1355 | subsection \<open>Meta-level conjunction\<close> | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1356 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1357 | lemma all_conjunction: | 
| 58612 | 1358 | "(\<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 | 1359 | proof | 
| 58612 | 1360 | assume conj: "\<And>x. PROP A x &&& PROP B x" | 
| 1361 | show "(\<And>x. PROP A x) &&& (\<And>x. PROP B x)" | |
| 19121 | 1362 | proof - | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1363 | fix x | 
| 26958 | 1364 | from conj show "PROP A x" by (rule conjunctionD1) | 
| 1365 | from conj show "PROP B x" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1366 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1367 | next | 
| 58612 | 1368 | 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 | 1369 | fix x | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 1370 | show "PROP A x &&& PROP B x" | 
| 19121 | 1371 | proof - | 
| 26958 | 1372 | show "PROP A x" by (rule conj [THEN conjunctionD1, rule_format]) | 
| 1373 | 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 | 1374 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1375 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1376 | |
| 19121 | 1377 | lemma imp_conjunction: | 
| 58612 | 1378 | "(PROP A \<Longrightarrow> PROP B &&& PROP C) \<equiv> ((PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C))" | 
| 18836 | 1379 | proof | 
| 58612 | 1380 | assume conj: "PROP A \<Longrightarrow> PROP B &&& PROP C" | 
| 1381 | show "(PROP A \<Longrightarrow> PROP B) &&& (PROP A \<Longrightarrow> PROP C)" | |
| 19121 | 1382 | proof - | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1383 | assume "PROP A" | 
| 58611 | 1384 | from conj [OF \<open>PROP A\<close>] show "PROP B" by (rule conjunctionD1) | 
| 1385 | 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 | 1386 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1387 | next | 
| 58612 | 1388 | 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 | 1389 | assume "PROP A" | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 1390 | show "PROP B &&& PROP C" | 
| 19121 | 1391 | proof - | 
| 58611 | 1392 | from \<open>PROP A\<close> show "PROP B" by (rule conj [THEN conjunctionD1]) | 
| 1393 | 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 | 1394 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1395 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1396 | |
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1397 | lemma conjunction_imp: | 
| 58612 | 1398 | "(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 | 1399 | proof | 
| 58612 | 1400 | assume r: "PROP A &&& PROP B \<Longrightarrow> PROP C" | 
| 22933 | 1401 | assume ab: "PROP A" "PROP B" | 
| 1402 | show "PROP C" | |
| 1403 | proof (rule r) | |
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28699diff
changeset | 1404 | from ab show "PROP A &&& PROP B" . | 
| 22933 | 1405 | qed | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1406 | next | 
| 58612 | 1407 | 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 | 1408 | assume conj: "PROP A &&& PROP B" | 
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1409 | show "PROP C" | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1410 | proof (rule r) | 
| 19121 | 1411 | from conj show "PROP A" by (rule conjunctionD1) | 
| 1412 | from conj show "PROP B" by (rule conjunctionD2) | |
| 18466 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1413 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1414 | qed | 
| 
389a6f9c31f4
added locale meta_conjunction_syntax and various conjunction rules;
 wenzelm parents: 
18019diff
changeset | 1415 | |
| 48638 | 1416 | end |