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