| author | Andreas Lochbihler | 
| Tue, 26 Jul 2011 14:05:28 +0200 | |
| changeset 43979 | 9f27d2bf4087 | 
| parent 43947 | 9b00f09f7721 | 
| child 44186 | 806f0ec1a43d | 
| permissions | -rw-r--r-- | 
| 5832 | 1  | 
(* Title: Pure/Isar/isar_syn.ML  | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
|
3  | 
||
| 6353 | 4  | 
Isar/Pure outer syntax.  | 
| 5832 | 5  | 
*)  | 
6  | 
||
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
7  | 
structure Isar_Syn: sig end =  | 
| 5832 | 8  | 
struct  | 
9  | 
||
| 24868 | 10  | 
(** keywords **)  | 
11  | 
||
12  | 
(*keep keywords consistent with the parsers, otherwise be prepared for  | 
|
13  | 
unexpected errors*)  | 
|
14  | 
||
| 36950 | 15  | 
val _ = List.app Keyword.keyword  | 
| 24868 | 16  | 
 ["!!", "!", "%", "(", ")", "+", ",", "--", ":", "::", ";", "<", "<=",
 | 
17  | 
"=", "==", "=>", "?", "[", "\\<equiv>", "\\<leftharpoondown>",  | 
|
18  | 
"\\<rightharpoonup>", "\\<rightleftharpoons>", "\\<subseteq>", "]",  | 
|
| 26888 | 19  | 
"advanced", "and", "assumes", "attach", "begin", "binder",  | 
| 24868 | 20  | 
"constrains", "defines", "fixes", "for", "identifier", "if",  | 
| 28721 | 21  | 
"imports", "in", "infix", "infixl", "infixr", "is",  | 
| 
33456
 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 
wenzelm 
parents: 
33390 
diff
changeset
 | 
22  | 
"notes", "obtains", "open", "output", "overloaded", "pervasive",  | 
| 
 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 
wenzelm 
parents: 
33390 
diff
changeset
 | 
23  | 
"shows", "structure", "unchecked", "uses", "where", "|"];  | 
| 
 
fbd47f9b9b12
allow "pervasive" local theory declarations, which are applied the background theory;
 
wenzelm 
parents: 
33390 
diff
changeset
 | 
24  | 
|
| 24868 | 25  | 
|
26  | 
||
| 5832 | 27  | 
(** init and exit **)  | 
28  | 
||
| 24868 | 29  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
30  | 
Outer_Syntax.command "theory" "begin theory" (Keyword.tag_theory Keyword.thy_begin)  | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
31  | 
(Thy_Header.args >> (fn (name, imports, uses) =>  | 
| 
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
32  | 
Toplevel.print o  | 
| 
41536
 
47fef6afe756
Toplevel.init_theory: maintain optional master directory, to allow bypassing global Thy_Load.master_path;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
33  | 
Toplevel.init_theory NONE name  | 
| 
 
47fef6afe756
Toplevel.init_theory: maintain optional master directory, to allow bypassing global Thy_Load.master_path;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
34  | 
(fn master =>  | 
| 
 
47fef6afe756
Toplevel.init_theory: maintain optional master directory, to allow bypassing global Thy_Load.master_path;
 
wenzelm 
parents: 
41435 
diff
changeset
 | 
35  | 
Thy_Info.toplevel_begin_theory master name imports (map (apfst Path.explode) uses))));  | 
| 5832 | 36  | 
|
| 24868 | 37  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
38  | 
Outer_Syntax.command "end" "end (local) theory" (Keyword.tag_theory Keyword.thy_end)  | 
| 40960 | 39  | 
(Scan.succeed  | 
40  | 
(Toplevel.exit o Toplevel.end_local_theory o Toplevel.end_proof (K Proof.end_notepad)));  | 
|
| 6687 | 41  | 
|
| 5832 | 42  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
43  | 
|
| 7749 | 44  | 
(** markup commands **)  | 
| 5832 | 45  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
46  | 
val _ = Outer_Syntax.markup_command Thy_Output.Markup "header" "theory header" Keyword.diag  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
47  | 
(Parse.doc_source >> Isar_Cmd.header_markup);  | 
| 6353 | 48  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
49  | 
val _ = Outer_Syntax.markup_command Thy_Output.Markup "chapter" "chapter heading"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
50  | 
Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup);  | 
| 5958 | 51  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
52  | 
val _ = Outer_Syntax.markup_command Thy_Output.Markup "section" "section heading"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
53  | 
Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup);  | 
| 5958 | 54  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
55  | 
val _ = Outer_Syntax.markup_command Thy_Output.Markup "subsection" "subsection heading"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
56  | 
Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup);  | 
| 5958 | 57  | 
|
| 24868 | 58  | 
val _ =  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
59  | 
Outer_Syntax.markup_command Thy_Output.Markup "subsubsection" "subsubsection heading"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
60  | 
Keyword.thy_heading (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup);  | 
| 5832 | 61  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
62  | 
val _ = Outer_Syntax.markup_command Thy_Output.MarkupEnv "text" "formal comment (theory)"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
63  | 
Keyword.thy_decl (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup);  | 
| 7172 | 64  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
65  | 
val _ = Outer_Syntax.markup_command Thy_Output.Verbatim "text_raw" "raw document preparation text"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
66  | 
Keyword.thy_decl (Parse.opt_target -- Parse.doc_source >> Isar_Cmd.local_theory_markup);  | 
| 7172 | 67  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
68  | 
val _ = Outer_Syntax.markup_command Thy_Output.Markup "sect" "formal comment (proof)"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
69  | 
(Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 70  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
71  | 
val _ = Outer_Syntax.markup_command Thy_Output.Markup "subsect" "formal comment (proof)"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
72  | 
(Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 73  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
74  | 
val _ = Outer_Syntax.markup_command Thy_Output.Markup "subsubsect" "formal comment (proof)"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
75  | 
(Keyword.tag_proof Keyword.prf_heading) (Parse.doc_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 76  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
77  | 
val _ = Outer_Syntax.markup_command Thy_Output.MarkupEnv "txt" "formal comment (proof)"  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
78  | 
(Keyword.tag_proof Keyword.prf_decl) (Parse.doc_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 79  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
80  | 
val _ = Outer_Syntax.markup_command Thy_Output.Verbatim "txt_raw"  | 
| 36950 | 81  | 
"raw document preparation text (proof)" (Keyword.tag_proof Keyword.prf_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
82  | 
(Parse.doc_source >> Isar_Cmd.proof_markup);  | 
| 7775 | 83  | 
|
| 5832 | 84  | 
|
| 6886 | 85  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
86  | 
(** theory commands **)  | 
| 6886 | 87  | 
|
| 5832 | 88  | 
(* classes and sorts *)  | 
89  | 
||
| 24868 | 90  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
91  | 
Outer_Syntax.command "classes" "declare type classes" Keyword.thy_decl  | 
| 36950 | 92  | 
(Scan.repeat1 (Parse.binding -- Scan.optional ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") |--  | 
93  | 
Parse.!!! (Parse.list1 Parse.xname)) [])  | 
|
94  | 
>> (Toplevel.theory o fold AxClass.axiomatize_class_cmd));  | 
|
| 5832 | 95  | 
|
| 24868 | 96  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
97  | 
Outer_Syntax.command "classrel" "state inclusion of type classes (axiomatic!)" Keyword.thy_decl  | 
| 36950 | 98  | 
(Parse.and_list1 (Parse.xname -- ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<")  | 
99  | 
|-- Parse.!!! Parse.xname))  | 
|
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
100  | 
>> (Toplevel.theory o AxClass.axiomatize_classrel_cmd));  | 
| 5832 | 101  | 
|
| 24868 | 102  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
103  | 
Outer_Syntax.local_theory "default_sort" "declare default sort for explicit type variables"  | 
| 36950 | 104  | 
Keyword.thy_decl  | 
105  | 
(Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy));  | 
|
| 5832 | 106  | 
|
107  | 
||
108  | 
(* types *)  | 
|
109  | 
||
| 24868 | 110  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
111  | 
Outer_Syntax.local_theory "typedecl" "type declaration" Keyword.thy_decl  | 
| 36950 | 112  | 
(Parse.type_args -- Parse.binding -- Parse.opt_mixfix  | 
| 35835 | 113  | 
>> (fn ((args, a), mx) => Typedecl.typedecl (a, map (rpair dummyS) args, mx) #> snd));  | 
| 5832 | 114  | 
|
| 
41249
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
115  | 
val type_abbrev =  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
116  | 
Parse.type_args -- Parse.binding --  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
117  | 
(Parse.$$$ "=" |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix'));  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
118  | 
|
| 24868 | 119  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
120  | 
Outer_Syntax.local_theory "types" "declare type abbreviations" Keyword.thy_decl  | 
| 
41249
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
121  | 
(Scan.repeat1 type_abbrev >> (fn specs => fn lthy =>  | 
| 43227 | 122  | 
(legacy_feature "Old 'types' command -- use 'type_synonym' instead";  | 
| 
41249
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
123  | 
fold (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs) specs lthy)));  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
124  | 
|
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
125  | 
val _ =  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
126  | 
Outer_Syntax.local_theory "type_synonym" "declare type abbreviation" Keyword.thy_decl  | 
| 
 
26f12f98f50a
Command 'type_synonym' (with single argument) supersedes 'types' (legacy feature);
 
wenzelm 
parents: 
41229 
diff
changeset
 | 
127  | 
(type_abbrev >> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs));  | 
| 5832 | 128  | 
|
| 24868 | 129  | 
val _ =  | 
| 
41229
 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 
wenzelm 
parents: 
40965 
diff
changeset
 | 
130  | 
Outer_Syntax.command "nonterminal"  | 
| 
 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 
wenzelm 
parents: 
40965 
diff
changeset
 | 
131  | 
"declare syntactic type constructors (grammar nonterminal symbols)" Keyword.thy_decl  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42359 
diff
changeset
 | 
132  | 
(Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global));  | 
| 5832 | 133  | 
|
| 24868 | 134  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
135  | 
Outer_Syntax.command "arities" "state type arities (axiomatic!)" Keyword.thy_decl  | 
| 36950 | 136  | 
(Scan.repeat1 Parse.arity >> (Toplevel.theory o fold AxClass.axiomatize_arity_cmd));  | 
| 5832 | 137  | 
|
138  | 
||
139  | 
(* consts and syntax *)  | 
|
140  | 
||
| 24868 | 141  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
142  | 
Outer_Syntax.command "judgment" "declare object-logic judgment" Keyword.thy_decl  | 
| 36950 | 143  | 
(Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd));  | 
| 8227 | 144  | 
|
| 24868 | 145  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
146  | 
Outer_Syntax.command "consts" "declare constants" Keyword.thy_decl  | 
| 36950 | 147  | 
(Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts));  | 
| 5832 | 148  | 
|
| 36950 | 149  | 
val opt_overloaded = Parse.opt_keyword "overloaded";  | 
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
13802 
diff
changeset
 | 
150  | 
|
| 24868 | 151  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
152  | 
Outer_Syntax.command "finalconsts" "declare constants as final" Keyword.thy_decl  | 
| 36950 | 153  | 
(opt_overloaded -- Scan.repeat1 Parse.term >> (uncurry (Toplevel.theory oo Theory.add_finals)));  | 
| 9731 | 154  | 
|
| 12142 | 155  | 
val mode_spec =  | 
| 36950 | 156  | 
  (Parse.$$$ "output" >> K ("", false)) ||
 | 
157  | 
Parse.name -- Scan.optional (Parse.$$$ "output" >> K false) true;  | 
|
| 9731 | 158  | 
|
| 14900 | 159  | 
val opt_mode =  | 
| 36950 | 160  | 
  Scan.optional (Parse.$$$ "(" |-- Parse.!!! (mode_spec --| Parse.$$$ ")")) Syntax.mode_default;
 | 
| 5832 | 161  | 
|
| 24868 | 162  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
163  | 
Outer_Syntax.command "syntax" "declare syntactic constants" Keyword.thy_decl  | 
| 42299 | 164  | 
(opt_mode -- Scan.repeat1 Parse.const_decl >> (Toplevel.theory o uncurry Sign.add_modesyntax));  | 
| 5832 | 165  | 
|
| 24868 | 166  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
167  | 
Outer_Syntax.command "no_syntax" "delete syntax declarations" Keyword.thy_decl  | 
| 42299 | 168  | 
(opt_mode -- Scan.repeat1 Parse.const_decl >> (Toplevel.theory o uncurry Sign.del_modesyntax));  | 
| 15748 | 169  | 
|
| 5832 | 170  | 
|
171  | 
(* translations *)  | 
|
172  | 
||
173  | 
val trans_pat =  | 
|
| 42326 | 174  | 
Scan.optional  | 
175  | 
    (Parse.$$$ "(" |-- Parse.!!! (Parse.xname --| Parse.$$$ ")")) "logic"
 | 
|
176  | 
-- Parse.inner_syntax Parse.string;  | 
|
| 5832 | 177  | 
|
178  | 
fun trans_arrow toks =  | 
|
| 42204 | 179  | 
((Parse.$$$ "\\<rightharpoonup>" || Parse.$$$ "=>") >> K Syntax.Parse_Rule ||  | 
180  | 
(Parse.$$$ "\\<leftharpoondown>" || Parse.$$$ "<=") >> K Syntax.Print_Rule ||  | 
|
181  | 
(Parse.$$$ "\\<rightleftharpoons>" || Parse.$$$ "==") >> K Syntax.Parse_Print_Rule) toks;  | 
|
| 5832 | 182  | 
|
183  | 
val trans_line =  | 
|
| 36950 | 184  | 
trans_pat -- Parse.!!! (trans_arrow -- trans_pat)  | 
| 5832 | 185  | 
>> (fn (left, (arr, right)) => arr (left, right));  | 
186  | 
||
| 24868 | 187  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
188  | 
Outer_Syntax.command "translations" "declare syntax translation rules" Keyword.thy_decl  | 
| 42204 | 189  | 
(Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.translations));  | 
| 5832 | 190  | 
|
| 24868 | 191  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
192  | 
Outer_Syntax.command "no_translations" "remove syntax translation rules" Keyword.thy_decl  | 
| 42204 | 193  | 
(Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.no_translations));  | 
| 19260 | 194  | 
|
| 5832 | 195  | 
|
196  | 
(* axioms and definitions *)  | 
|
197  | 
||
| 24868 | 198  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
199  | 
Outer_Syntax.command "axioms" "state arbitrary propositions (axiomatic!)" Keyword.thy_decl  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
200  | 
(Scan.repeat1 Parse_Spec.spec >>  | 
| 37861 | 201  | 
(fn spec => Toplevel.theory (fn thy =>  | 
202  | 
(legacy_feature "Old 'axioms' command -- use 'axiomatization' instead";  | 
|
203  | 
Isar_Cmd.add_axioms spec thy))));  | 
|
| 5832 | 204  | 
|
| 19631 | 205  | 
val opt_unchecked_overloaded =  | 
| 36950 | 206  | 
  Scan.optional (Parse.$$$ "(" |-- Parse.!!!
 | 
207  | 
(((Parse.$$$ "unchecked" >> K true) -- Scan.optional (Parse.$$$ "overloaded" >> K true) false ||  | 
|
208  | 
Parse.$$$ "overloaded" >> K (false, true)) --| Parse.$$$ ")")) (false, false);  | 
|
| 19631 | 209  | 
|
| 24868 | 210  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
211  | 
Outer_Syntax.command "defs" "define constants" Keyword.thy_decl  | 
| 
35852
 
4e3fe0b8687b
minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
 
wenzelm 
parents: 
35835 
diff
changeset
 | 
212  | 
(opt_unchecked_overloaded --  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
213  | 
Scan.repeat1 (Parse_Spec.thm_name ":" -- Parse.prop >> (fn ((x, y), z) => ((x, z), y)))  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
214  | 
>> (Toplevel.theory o Isar_Cmd.add_defs));  | 
| 6370 | 215  | 
|
| 14642 | 216  | 
|
| 
21601
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
217  | 
(* constant definitions and abbreviations *)  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
218  | 
|
| 24868 | 219  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
220  | 
Outer_Syntax.local_theory "definition" "constant definition" Keyword.thy_decl  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
221  | 
(Parse_Spec.constdef >> (fn args => #2 o Specification.definition_cmd args));  | 
| 18780 | 222  | 
|
| 24868 | 223  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
224  | 
Outer_Syntax.local_theory "abbreviation" "constant abbreviation" Keyword.thy_decl  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
225  | 
(opt_mode -- (Scan.option Parse_Spec.constdecl -- Parse.prop)  | 
| 36950 | 226  | 
>> (fn (mode, args) => Specification.abbreviation_cmd mode args));  | 
| 19659 | 227  | 
|
| 24868 | 228  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
229  | 
Outer_Syntax.local_theory "type_notation" "add concrete syntax for type constructors"  | 
| 36950 | 230  | 
Keyword.thy_decl  | 
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
231  | 
(opt_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix)  | 
| 36950 | 232  | 
>> (fn (mode, args) => Specification.type_notation_cmd true mode args));  | 
| 35413 | 233  | 
|
234  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
235  | 
Outer_Syntax.local_theory "no_type_notation" "delete concrete syntax for type constructors"  | 
| 36950 | 236  | 
Keyword.thy_decl  | 
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
237  | 
(opt_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix)  | 
| 36950 | 238  | 
>> (fn (mode, args) => Specification.type_notation_cmd false mode args));  | 
| 35413 | 239  | 
|
240  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
241  | 
Outer_Syntax.local_theory "notation" "add concrete syntax for constants / fixed variables"  | 
| 36950 | 242  | 
Keyword.thy_decl  | 
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
243  | 
(opt_mode -- Parse.and_list1 (Parse.const -- Parse_Spec.locale_mixfix)  | 
| 36950 | 244  | 
>> (fn (mode, args) => Specification.notation_cmd true mode args));  | 
| 24950 | 245  | 
|
246  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
247  | 
Outer_Syntax.local_theory "no_notation" "delete concrete syntax for constants / fixed variables"  | 
| 36950 | 248  | 
Keyword.thy_decl  | 
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
249  | 
(opt_mode -- Parse.and_list1 (Parse.const -- Parse_Spec.locale_mixfix)  | 
| 36950 | 250  | 
>> (fn (mode, args) => Specification.notation_cmd false mode args));  | 
| 19076 | 251  | 
|
| 5832 | 252  | 
|
| 18616 | 253  | 
(* constant specifications *)  | 
254  | 
||
| 24868 | 255  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
256  | 
Outer_Syntax.command "axiomatization" "axiomatic constant specification" Keyword.thy_decl  | 
| 36950 | 257  | 
(Scan.optional Parse.fixes [] --  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
258  | 
Scan.optional (Parse.where_ |-- Parse.!!! (Parse.and_list1 Parse_Spec.specs)) []  | 
| 36950 | 259  | 
>> (fn (x, y) => Toplevel.theory (#2 o Specification.axiomatization_cmd x y)));  | 
| 18616 | 260  | 
|
261  | 
||
| 5914 | 262  | 
(* theorems *)  | 
263  | 
||
| 
26988
 
742e26213212
more uniform treatment of OuterSyntax.local_theory commands;
 
wenzelm 
parents: 
26888 
diff
changeset
 | 
264  | 
fun theorems kind =  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
265  | 
Parse_Spec.name_facts >> (fn args => #2 o Specification.theorems_cmd kind args);  | 
| 12712 | 266  | 
|
| 24868 | 267  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
268  | 
Outer_Syntax.local_theory "theorems" "define theorems" Keyword.thy_decl (theorems Thm.theoremK);  | 
| 5914 | 269  | 
|
| 24868 | 270  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
271  | 
Outer_Syntax.local_theory "lemmas" "define lemmas" Keyword.thy_decl (theorems Thm.lemmaK);  | 
| 5914 | 272  | 
|
| 24868 | 273  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
274  | 
Outer_Syntax.local_theory "declare" "declare theorems" Keyword.thy_decl  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
275  | 
(Parse.and_list1 Parse_Spec.xthms1  | 
| 28965 | 276  | 
>> (fn args => #2 o Specification.theorems_cmd "" [(Attrib.empty_binding, flat args)]));  | 
| 9589 | 277  | 
|
| 5914 | 278  | 
|
| 5832 | 279  | 
(* name space entry path *)  | 
280  | 
||
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36174 
diff
changeset
 | 
281  | 
fun hide_names name hide what =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
282  | 
  Outer_Syntax.command name ("hide " ^ what ^ " from name space") Keyword.thy_decl
 | 
| 36950 | 283  | 
((Parse.opt_keyword "open" >> not) -- Scan.repeat1 Parse.xname >>  | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36174 
diff
changeset
 | 
284  | 
(Toplevel.theory o uncurry hide));  | 
| 
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36174 
diff
changeset
 | 
285  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
286  | 
val _ = hide_names "hide_class" Isar_Cmd.hide_class "classes";  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
287  | 
val _ = hide_names "hide_type" Isar_Cmd.hide_type "types";  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
288  | 
val _ = hide_names "hide_const" Isar_Cmd.hide_const "constants";  | 
| 
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
289  | 
val _ = hide_names "hide_fact" Isar_Cmd.hide_fact "facts";  | 
| 5832 | 290  | 
|
291  | 
||
292  | 
(* use ML text *)  | 
|
293  | 
||
| 24868 | 294  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
295  | 
Outer_Syntax.command "use" "ML text from file" (Keyword.tag_ml Keyword.thy_decl)  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37873 
diff
changeset
 | 
296  | 
(Parse.path >> (fn path => Toplevel.generic_theory (Thy_Load.exec_ml path)));  | 
| 7891 | 297  | 
|
| 24868 | 298  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
299  | 
Outer_Syntax.command "ML" "ML text within theory or local theory"  | 
| 36950 | 300  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
301  | 
(Parse.ML_source >> (fn (txt, pos) =>  | 
|
| 
30579
 
4169ddbfe1f9
command 'use', 'ML': apply ML environment to theory and target as well;
 
wenzelm 
parents: 
30576 
diff
changeset
 | 
302  | 
Toplevel.generic_theory  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37873 
diff
changeset
 | 
303  | 
(ML_Context.exec (fn () => ML_Context.eval_text true pos txt) #>  | 
| 
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37873 
diff
changeset
 | 
304  | 
Local_Theory.propagate_ml_env)));  | 
| 
30579
 
4169ddbfe1f9
command 'use', 'ML': apply ML environment to theory and target as well;
 
wenzelm 
parents: 
30576 
diff
changeset
 | 
305  | 
|
| 
 
4169ddbfe1f9
command 'use', 'ML': apply ML environment to theory and target as well;
 
wenzelm 
parents: 
30576 
diff
changeset
 | 
306  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
307  | 
Outer_Syntax.command "ML_prf" "ML text within proof" (Keyword.tag_proof Keyword.prf_decl)  | 
| 36950 | 308  | 
(Parse.ML_source >> (fn (txt, pos) =>  | 
| 28269 | 309  | 
Toplevel.proof (Proof.map_context (Context.proof_map  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37873 
diff
changeset
 | 
310  | 
(ML_Context.exec (fn () => ML_Context.eval_text true pos txt))) #> Proof.propagate_ml_env)));  | 
| 28269 | 311  | 
|
312  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
313  | 
Outer_Syntax.command "ML_val" "diagnostic ML text" (Keyword.tag_ml Keyword.diag)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
314  | 
(Parse.ML_source >> Isar_Cmd.ml_diag true);  | 
| 26396 | 315  | 
|
316  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
317  | 
Outer_Syntax.command "ML_command" "diagnostic ML text (silent)" (Keyword.tag_ml Keyword.diag)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
318  | 
(Parse.ML_source >> (Toplevel.no_timing oo Isar_Cmd.ml_diag false));  | 
| 5832 | 319  | 
|
| 24868 | 320  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
321  | 
Outer_Syntax.command "setup" "ML theory setup" (Keyword.tag_ml Keyword.thy_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
322  | 
(Parse.ML_source >> (Toplevel.theory o Isar_Cmd.global_setup));  | 
| 30461 | 323  | 
|
324  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
325  | 
Outer_Syntax.local_theory "local_setup" "ML local theory setup" (Keyword.tag_ml Keyword.thy_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
326  | 
(Parse.ML_source >> Isar_Cmd.local_setup);  | 
| 5832 | 327  | 
|
| 24868 | 328  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
329  | 
Outer_Syntax.command "attribute_setup" "define attribute in ML" (Keyword.tag_ml Keyword.thy_decl)  | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
330  | 
(Parse.position Parse.name --  | 
| 
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
331  | 
Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Scan.optional Parse.text "")  | 
| 36950 | 332  | 
>> (fn (name, (txt, cmt)) => Toplevel.theory (Attrib.attribute_setup name txt cmt)));  | 
| 30526 | 333  | 
|
334  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
335  | 
Outer_Syntax.command "method_setup" "define proof method in ML" (Keyword.tag_ml Keyword.thy_decl)  | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
336  | 
(Parse.position Parse.name --  | 
| 
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
337  | 
Parse.!!! (Parse.$$$ "=" |-- Parse.ML_source -- Scan.optional Parse.text "")  | 
| 36950 | 338  | 
>> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt)));  | 
| 9197 | 339  | 
|
| 24868 | 340  | 
val _ =  | 
| 40784 | 341  | 
Outer_Syntax.local_theory "declaration" "generic ML declaration"  | 
342  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
|
343  | 
(Parse.opt_keyword "pervasive" -- Parse.ML_source  | 
|
344  | 
      >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = false, pervasive = pervasive} txt));
 | 
|
345  | 
||
346  | 
val _ =  | 
|
347  | 
Outer_Syntax.local_theory "syntax_declaration" "generic ML declaration"  | 
|
348  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
|
349  | 
(Parse.opt_keyword "pervasive" -- Parse.ML_source  | 
|
350  | 
      >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = true, pervasive = pervasive} txt));
 | 
|
| 22088 | 351  | 
|
| 24868 | 352  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
353  | 
Outer_Syntax.local_theory "simproc_setup" "define simproc in ML" (Keyword.tag_ml Keyword.thy_decl)  | 
| 42464 | 354  | 
(Parse.position Parse.name --  | 
| 36950 | 355  | 
      (Parse.$$$ "(" |-- Parse.enum1 "|" Parse.term --| Parse.$$$ ")" --| Parse.$$$ "=") --
 | 
356  | 
Parse.ML_source -- Scan.optional (Parse.$$$ "identifier" |-- Scan.repeat1 Parse.xname) []  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
357  | 
>> (fn (((a, b), c), d) => Isar_Cmd.simproc_setup a b c d));  | 
| 22202 | 358  | 
|
| 5832 | 359  | 
|
360  | 
(* translation functions *)  | 
|
361  | 
||
| 36950 | 362  | 
val trfun = Parse.opt_keyword "advanced" -- Parse.ML_source;  | 
| 14642 | 363  | 
|
| 24868 | 364  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
365  | 
Outer_Syntax.command "parse_ast_translation" "install parse ast translation functions"  | 
| 36950 | 366  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
367  | 
(trfun >> (Toplevel.theory o Isar_Cmd.parse_ast_translation));  | 
| 5832 | 368  | 
|
| 24868 | 369  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
370  | 
Outer_Syntax.command "parse_translation" "install parse translation functions"  | 
| 36950 | 371  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
372  | 
(trfun >> (Toplevel.theory o Isar_Cmd.parse_translation));  | 
| 5832 | 373  | 
|
| 24868 | 374  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
375  | 
Outer_Syntax.command "print_translation" "install print translation functions"  | 
| 36950 | 376  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
377  | 
(trfun >> (Toplevel.theory o Isar_Cmd.print_translation));  | 
| 5832 | 378  | 
|
| 24868 | 379  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
380  | 
Outer_Syntax.command "typed_print_translation" "install typed print translation functions"  | 
| 36950 | 381  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
382  | 
(trfun >> (Toplevel.theory o Isar_Cmd.typed_print_translation));  | 
| 5832 | 383  | 
|
| 24868 | 384  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
385  | 
Outer_Syntax.command "print_ast_translation" "install print ast translation functions"  | 
| 36950 | 386  | 
(Keyword.tag_ml Keyword.thy_decl)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
387  | 
(trfun >> (Toplevel.theory o Isar_Cmd.print_ast_translation));  | 
| 5832 | 388  | 
|
389  | 
||
390  | 
(* oracles *)  | 
|
391  | 
||
| 24868 | 392  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
393  | 
Outer_Syntax.command "oracle" "declare oracle" (Keyword.tag_ml Keyword.thy_decl)  | 
| 36950 | 394  | 
(Parse.position Parse.name -- (Parse.$$$ "=" |-- Parse.ML_source) >>  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
395  | 
(fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y)));  | 
| 5832 | 396  | 
|
397  | 
||
| 
21306
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
398  | 
(* local theories *)  | 
| 
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
399  | 
|
| 24868 | 400  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
401  | 
Outer_Syntax.command "context" "enter local theory context" Keyword.thy_decl  | 
| 36950 | 402  | 
(Parse.name --| Parse.begin >> (fn name =>  | 
| 
38350
 
480b2de9927c
renamed Theory_Target to the more appropriate Named_Target
 
haftmann 
parents: 
38348 
diff
changeset
 | 
403  | 
Toplevel.print o Toplevel.begin_local_theory true (Named_Target.context_cmd name)));  | 
| 
21306
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
404  | 
|
| 
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
405  | 
|
| 12061 | 406  | 
(* locales *)  | 
407  | 
||
| 12758 | 408  | 
val locale_val =  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
409  | 
Parse_Spec.locale_expression false --  | 
| 
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
410  | 
Scan.optional (Parse.$$$ "+" |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||  | 
| 
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
411  | 
Scan.repeat1 Parse_Spec.context_element >> pair ([], []);  | 
| 12061 | 412  | 
|
| 24868 | 413  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
414  | 
Outer_Syntax.command "locale" "define named proof context" Keyword.thy_decl  | 
| 36950 | 415  | 
(Parse.binding --  | 
416  | 
Scan.optional (Parse.$$$ "=" |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin  | 
|
| 27681 | 417  | 
>> (fn ((name, (expr, elems)), begin) =>  | 
| 21843 | 418  | 
(begin ? Toplevel.print) o Toplevel.begin_local_theory begin  | 
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41536 
diff
changeset
 | 
419  | 
(Expression.add_locale_cmd I name Binding.empty expr elems #> snd)));  | 
| 
28085
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
28084 
diff
changeset
 | 
420  | 
|
| 41270 | 421  | 
fun parse_interpretation_arguments mandatory =  | 
422  | 
Parse.!!! (Parse_Spec.locale_expression mandatory) --  | 
|
423  | 
Scan.optional  | 
|
424  | 
(Parse.where_ |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) [];  | 
|
425  | 
||
| 24868 | 426  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
427  | 
Outer_Syntax.command "sublocale"  | 
| 36950 | 428  | 
"prove sublocale relation between a locale and a locale expression" Keyword.thy_goal  | 
429  | 
(Parse.xname --| (Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") --  | 
|
| 41270 | 430  | 
parse_interpretation_arguments false  | 
431  | 
>> (fn (loc, (expr, equations)) =>  | 
|
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41536 
diff
changeset
 | 
432  | 
Toplevel.print o Toplevel.theory_to_proof (Expression.sublocale_cmd I loc expr equations)));  | 
| 28895 | 433  | 
|
434  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
435  | 
Outer_Syntax.command "interpretation"  | 
| 36950 | 436  | 
"prove interpretation of locale expression in theory" Keyword.thy_goal  | 
| 41270 | 437  | 
(parse_interpretation_arguments true  | 
| 
30727
 
519f8f0fcbf3
interpretation/interpret: prefixes within locale expression are mandatory by default;
 
wenzelm 
parents: 
30703 
diff
changeset
 | 
438  | 
>> (fn (expr, equations) => Toplevel.print o  | 
| 
 
519f8f0fcbf3
interpretation/interpret: prefixes within locale expression are mandatory by default;
 
wenzelm 
parents: 
30703 
diff
changeset
 | 
439  | 
Toplevel.theory_to_proof (Expression.interpretation_cmd expr equations)));  | 
| 29223 | 440  | 
|
441  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
442  | 
Outer_Syntax.command "interpret"  | 
| 29223 | 443  | 
"prove interpretation of locale expression in proof context"  | 
| 36950 | 444  | 
(Keyword.tag_proof Keyword.prf_goal)  | 
| 41270 | 445  | 
(parse_interpretation_arguments true  | 
| 38108 | 446  | 
>> (fn (expr, equations) => Toplevel.print o  | 
447  | 
Toplevel.proof' (Expression.interpret_cmd expr equations)));  | 
|
| 29223 | 448  | 
|
| 15703 | 449  | 
|
| 22299 | 450  | 
(* classes *)  | 
451  | 
||
| 24868 | 452  | 
val class_val =  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
453  | 
Parse_Spec.class_expr --  | 
| 
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
454  | 
Scan.optional (Parse.$$$ "+" |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||  | 
| 
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
455  | 
Scan.repeat1 Parse_Spec.context_element >> pair [];  | 
| 22299 | 456  | 
|
| 24868 | 457  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
458  | 
Outer_Syntax.command "class" "define type class" Keyword.thy_decl  | 
| 36950 | 459  | 
(Parse.binding -- Scan.optional (Parse.$$$ "=" |-- class_val) ([], []) -- Parse.opt_begin  | 
| 26516 | 460  | 
>> (fn ((name, (supclasses, elems)), begin) =>  | 
| 24938 | 461  | 
(begin ? Toplevel.print) o Toplevel.begin_local_theory begin  | 
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41536 
diff
changeset
 | 
462  | 
(Class_Declaration.class_cmd I name supclasses elems #> snd)));  | 
| 22299 | 463  | 
|
| 24868 | 464  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
465  | 
Outer_Syntax.local_theory_to_proof "subclass" "prove a subclass relation" Keyword.thy_goal  | 
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41536 
diff
changeset
 | 
466  | 
(Parse.xname >> Class_Declaration.subclass_cmd I);  | 
| 
24914
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
467  | 
|
| 
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
468  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
469  | 
Outer_Syntax.command "instantiation" "instantiate and prove type arity" Keyword.thy_decl  | 
| 36950 | 470  | 
(Parse.multi_arity --| Parse.begin  | 
| 25462 | 471  | 
>> (fn arities => Toplevel.print o  | 
| 
38348
 
cf7b2121ad9d
moved instantiation target formally to class_target.ML
 
haftmann 
parents: 
38342 
diff
changeset
 | 
472  | 
Toplevel.begin_local_theory true (Class.instantiation_cmd arities)));  | 
| 24589 | 473  | 
|
| 25485 | 474  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
475  | 
Outer_Syntax.command "instance" "prove type arity or subclass relation" Keyword.thy_goal  | 
| 36950 | 476  | 
((Parse.xname -- ((Parse.$$$ "\\<subseteq>" || Parse.$$$ "<") |-- Parse.!!! Parse.xname)  | 
477  | 
>> Class.classrel_cmd ||  | 
|
478  | 
Parse.multi_arity >> Class.instance_arity_cmd)  | 
|
479  | 
>> (Toplevel.print oo Toplevel.theory_to_proof) ||  | 
|
480  | 
Scan.succeed  | 
|
481  | 
(Toplevel.print o Toplevel.local_theory_to_proof NONE (Class.instantiation_instance I)));  | 
|
| 22299 | 482  | 
|
483  | 
||
| 25519 | 484  | 
(* arbitrary overloading *)  | 
485  | 
||
486  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
487  | 
Outer_Syntax.command "overloading" "overloaded definitions" Keyword.thy_decl  | 
| 36950 | 488  | 
(Scan.repeat1 (Parse.name --| (Parse.$$$ "\\<equiv>" || Parse.$$$ "==") -- Parse.term --  | 
489  | 
      Scan.optional (Parse.$$$ "(" |-- (Parse.$$$ "unchecked" >> K false) --| Parse.$$$ ")") true
 | 
|
490  | 
>> Parse.triple1) --| Parse.begin  | 
|
| 25519 | 491  | 
>> (fn operations => Toplevel.print o  | 
| 
38342
 
09d4a04d5c2e
moved overloading target formally to overloading.ML
 
haftmann 
parents: 
38108 
diff
changeset
 | 
492  | 
Toplevel.begin_local_theory true (Overloading.overloading_cmd operations)));  | 
| 25519 | 493  | 
|
494  | 
||
| 22866 | 495  | 
(* code generation *)  | 
496  | 
||
| 24868 | 497  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
498  | 
Outer_Syntax.command "code_datatype" "define set of code datatype constructors" Keyword.thy_decl  | 
| 36950 | 499  | 
(Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd));  | 
| 22866 | 500  | 
|
501  | 
||
| 5832 | 502  | 
|
503  | 
(** proof commands **)  | 
|
504  | 
||
505  | 
(* statements *)  | 
|
506  | 
||
| 
36317
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
507  | 
fun gen_theorem schematic kind =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
508  | 
Outer_Syntax.local_theory_to_proof'  | 
| 
36317
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
509  | 
(if schematic then "schematic_" ^ kind else kind)  | 
| 
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
510  | 
    ("state " ^ (if schematic then "schematic " ^ kind else kind))
 | 
| 36950 | 511  | 
(if schematic then Keyword.thy_schematic_goal else Keyword.thy_goal)  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
512  | 
(Scan.optional (Parse_Spec.opt_thm_name ":" --|  | 
| 
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
513  | 
Scan.ahead (Parse_Spec.locale_keyword || Parse_Spec.statement_keyword)) Attrib.empty_binding --  | 
| 
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
514  | 
Parse_Spec.general_statement >> (fn (a, (elems, concl)) =>  | 
| 
36317
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
515  | 
((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd)  | 
| 
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
516  | 
kind NONE (K I) a elems concl)));  | 
| 5832 | 517  | 
|
| 
36317
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
518  | 
val _ = gen_theorem false Thm.theoremK;  | 
| 
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
519  | 
val _ = gen_theorem false Thm.lemmaK;  | 
| 
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
520  | 
val _ = gen_theorem false Thm.corollaryK;  | 
| 
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
521  | 
val _ = gen_theorem true Thm.theoremK;  | 
| 
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
522  | 
val _ = gen_theorem true Thm.lemmaK;  | 
| 
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
523  | 
val _ = gen_theorem true Thm.corollaryK;  | 
| 5832 | 524  | 
|
| 24868 | 525  | 
val _ =  | 
| 40960 | 526  | 
Outer_Syntax.local_theory_to_proof "notepad"  | 
527  | 
"Isar proof state as formal notepad, without any result" Keyword.thy_decl  | 
|
528  | 
(Parse.begin >> K Proof.begin_notepad);  | 
|
529  | 
||
530  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
531  | 
Outer_Syntax.command "have" "state local goal"  | 
| 36950 | 532  | 
(Keyword.tag_proof Keyword.prf_goal)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
533  | 
(Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.have));  | 
| 17353 | 534  | 
|
| 24868 | 535  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
536  | 
Outer_Syntax.command "hence" "abbreviates \"then have\""  | 
| 36950 | 537  | 
(Keyword.tag_proof Keyword.prf_goal)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
538  | 
(Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.hence));  | 
| 17353 | 539  | 
|
| 24868 | 540  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
541  | 
Outer_Syntax.command "show" "state local goal, solving current obligation"  | 
| 36950 | 542  | 
(Keyword.tag_proof Keyword.prf_asm_goal)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
543  | 
(Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.show));  | 
| 5832 | 544  | 
|
| 24868 | 545  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
546  | 
Outer_Syntax.command "thus" "abbreviates \"then show\""  | 
| 36950 | 547  | 
(Keyword.tag_proof Keyword.prf_asm_goal)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
548  | 
(Parse_Spec.statement >> ((Toplevel.print oo Toplevel.proof') o Isar_Cmd.thus));  | 
| 6501 | 549  | 
|
| 5832 | 550  | 
|
| 5914 | 551  | 
(* facts *)  | 
| 5832 | 552  | 
|
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
553  | 
val facts = Parse.and_list1 Parse_Spec.xthms1;  | 
| 9197 | 554  | 
|
| 24868 | 555  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
556  | 
Outer_Syntax.command "then" "forward chaining"  | 
| 36950 | 557  | 
(Keyword.tag_proof Keyword.prf_chain)  | 
| 17900 | 558  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.chain));  | 
| 5832 | 559  | 
|
| 24868 | 560  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
561  | 
Outer_Syntax.command "from" "forward chaining from given facts"  | 
| 36950 | 562  | 
(Keyword.tag_proof Keyword.prf_chain)  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
563  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss_cmd)));  | 
| 5914 | 564  | 
|
| 24868 | 565  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
566  | 
Outer_Syntax.command "with" "forward chaining from given and current facts"  | 
| 36950 | 567  | 
(Keyword.tag_proof Keyword.prf_chain)  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
568  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss_cmd)));  | 
| 6878 | 569  | 
|
| 24868 | 570  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
571  | 
Outer_Syntax.command "note" "define facts"  | 
| 36950 | 572  | 
(Keyword.tag_proof Keyword.prf_decl)  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
573  | 
(Parse_Spec.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss_cmd)));  | 
| 5832 | 574  | 
|
| 24868 | 575  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
576  | 
Outer_Syntax.command "using" "augment goal facts"  | 
| 36950 | 577  | 
(Keyword.tag_proof Keyword.prf_decl)  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
578  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.using_cmd)));  | 
| 18544 | 579  | 
|
| 24868 | 580  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
581  | 
Outer_Syntax.command "unfolding" "unfold definitions in goal and facts"  | 
| 36950 | 582  | 
(Keyword.tag_proof Keyword.prf_decl)  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
583  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding_cmd)));  | 
| 12926 | 584  | 
|
| 5832 | 585  | 
|
586  | 
(* proof context *)  | 
|
587  | 
||
| 24868 | 588  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
589  | 
Outer_Syntax.command "fix" "fix local variables (Skolem constants)"  | 
| 36950 | 590  | 
(Keyword.tag_proof Keyword.prf_asm)  | 
591  | 
(Parse.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix_cmd)));  | 
|
| 11890 | 592  | 
|
| 24868 | 593  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
594  | 
Outer_Syntax.command "assume" "assume propositions"  | 
| 36950 | 595  | 
(Keyword.tag_proof Keyword.prf_asm)  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
596  | 
(Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume_cmd)));  | 
| 5832 | 597  | 
|
| 24868 | 598  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
599  | 
Outer_Syntax.command "presume" "assume propositions, to be established later"  | 
| 36950 | 600  | 
(Keyword.tag_proof Keyword.prf_asm)  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
601  | 
(Parse_Spec.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume_cmd)));  | 
| 6850 | 602  | 
|
| 24868 | 603  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
604  | 
Outer_Syntax.command "def" "local definition"  | 
| 36950 | 605  | 
(Keyword.tag_proof Keyword.prf_asm)  | 
606  | 
(Parse.and_list1  | 
|
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
607  | 
(Parse_Spec.opt_thm_name ":" --  | 
| 36950 | 608  | 
((Parse.binding -- Parse.opt_mixfix) --  | 
609  | 
((Parse.$$$ "\\<equiv>" || Parse.$$$ "==") |-- Parse.!!! Parse.termp)))  | 
|
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
610  | 
>> (Toplevel.print oo (Toplevel.proof o Proof.def_cmd)));  | 
| 6953 | 611  | 
|
| 24868 | 612  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
613  | 
Outer_Syntax.command "obtain" "generalized existence"  | 
| 36950 | 614  | 
(Keyword.tag_proof Keyword.prf_asm_goal)  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
615  | 
(Parse.parname -- Scan.optional (Parse.fixes --| Parse.where_) [] -- Parse_Spec.statement  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
616  | 
>> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain_cmd x y z)));  | 
| 5832 | 617  | 
|
| 24868 | 618  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
619  | 
Outer_Syntax.command "guess" "wild guessing (unstructured)"  | 
| 36950 | 620  | 
(Keyword.tag_proof Keyword.prf_asm_goal)  | 
621  | 
(Scan.optional Parse.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess_cmd)));  | 
|
| 17854 | 622  | 
|
| 24868 | 623  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
624  | 
Outer_Syntax.command "let" "bind text variables"  | 
| 36950 | 625  | 
(Keyword.tag_proof Keyword.prf_decl)  | 
626  | 
(Parse.and_list1 (Parse.and_list1 Parse.term -- (Parse.$$$ "=" |-- Parse.term))  | 
|
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
627  | 
>> (Toplevel.print oo (Toplevel.proof o Proof.let_bind_cmd)));  | 
| 5832 | 628  | 
|
| 
36505
 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 
wenzelm 
parents: 
36452 
diff
changeset
 | 
629  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
630  | 
Outer_Syntax.command "write" "add concrete syntax for constants / fixed variables"  | 
| 36950 | 631  | 
(Keyword.tag_proof Keyword.prf_decl)  | 
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
632  | 
(opt_mode -- Parse.and_list1 (Parse.const -- Parse_Spec.locale_mixfix)  | 
| 
36505
 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 
wenzelm 
parents: 
36452 
diff
changeset
 | 
633  | 
>> (fn (mode, args) => Toplevel.print o Toplevel.proof (Proof.write_cmd mode args)));  | 
| 
 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 
wenzelm 
parents: 
36452 
diff
changeset
 | 
634  | 
|
| 
11793
 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 
wenzelm 
parents: 
11742 
diff
changeset
 | 
635  | 
val case_spec =  | 
| 36950 | 636  | 
  (Parse.$$$ "(" |--
 | 
| 42496 | 637  | 
Parse.!!! (Parse.xname -- Scan.repeat1 (Parse.maybe Parse.binding) --| Parse.$$$ ")") ||  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
638  | 
Parse.xname >> rpair []) -- Parse_Spec.opt_attribs >> Parse.triple1;  | 
| 
11793
 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 
wenzelm 
parents: 
11742 
diff
changeset
 | 
639  | 
|
| 24868 | 640  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
641  | 
Outer_Syntax.command "case" "invoke local context"  | 
| 36950 | 642  | 
(Keyword.tag_proof Keyword.prf_asm)  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
643  | 
(case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case_cmd)));  | 
| 8370 | 644  | 
|
| 5832 | 645  | 
|
646  | 
(* proof structure *)  | 
|
647  | 
||
| 24868 | 648  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
649  | 
  Outer_Syntax.command "{" "begin explicit proof block"
 | 
| 36950 | 650  | 
(Keyword.tag_proof Keyword.prf_open)  | 
| 17900 | 651  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.begin_block));  | 
| 5832 | 652  | 
|
| 24868 | 653  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
654  | 
Outer_Syntax.command "}" "end explicit proof block"  | 
| 36950 | 655  | 
(Keyword.tag_proof Keyword.prf_close)  | 
| 20305 | 656  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.end_block));  | 
| 6687 | 657  | 
|
| 24868 | 658  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
659  | 
Outer_Syntax.command "next" "enter next proof block"  | 
| 36950 | 660  | 
(Keyword.tag_proof Keyword.prf_block)  | 
| 17900 | 661  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.next_block));  | 
| 5832 | 662  | 
|
663  | 
||
664  | 
(* end proof *)  | 
|
665  | 
||
| 24868 | 666  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
667  | 
Outer_Syntax.command "qed" "conclude (sub-)proof"  | 
| 36950 | 668  | 
(Keyword.tag_proof Keyword.qed_block)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
669  | 
(Scan.option Method.parse >> Isar_Cmd.qed);  | 
| 5832 | 670  | 
|
| 24868 | 671  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
672  | 
Outer_Syntax.command "by" "terminal backward proof"  | 
| 36950 | 673  | 
(Keyword.tag_proof Keyword.qed)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
674  | 
(Method.parse -- Scan.option Method.parse >> Isar_Cmd.terminal_proof);  | 
| 6404 | 675  | 
|
| 24868 | 676  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
677  | 
Outer_Syntax.command ".." "default proof"  | 
| 36950 | 678  | 
(Keyword.tag_proof Keyword.qed)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
679  | 
(Scan.succeed Isar_Cmd.default_proof);  | 
| 8966 | 680  | 
|
| 24868 | 681  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
682  | 
Outer_Syntax.command "." "immediate proof"  | 
| 36950 | 683  | 
(Keyword.tag_proof Keyword.qed)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
684  | 
(Scan.succeed Isar_Cmd.immediate_proof);  | 
| 6404 | 685  | 
|
| 24868 | 686  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
687  | 
Outer_Syntax.command "done" "done proof"  | 
| 36950 | 688  | 
(Keyword.tag_proof Keyword.qed)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
689  | 
(Scan.succeed Isar_Cmd.done_proof);  | 
| 5832 | 690  | 
|
| 24868 | 691  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
692  | 
Outer_Syntax.command "sorry" "skip proof (quick-and-dirty mode only!)"  | 
| 36950 | 693  | 
(Keyword.tag_proof Keyword.qed)  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
694  | 
(Scan.succeed Isar_Cmd.skip_proof);  | 
| 
6888
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents: 
6886 
diff
changeset
 | 
695  | 
|
| 24868 | 696  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
697  | 
Outer_Syntax.command "oops" "forget proof"  | 
| 36950 | 698  | 
(Keyword.tag_proof Keyword.qed_global)  | 
| 18561 | 699  | 
(Scan.succeed Toplevel.forget_proof);  | 
| 8210 | 700  | 
|
| 5832 | 701  | 
|
702  | 
(* proof steps *)  | 
|
703  | 
||
| 24868 | 704  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
705  | 
Outer_Syntax.command "defer" "shuffle internal proof state"  | 
| 36950 | 706  | 
(Keyword.tag_proof Keyword.prf_script)  | 
707  | 
(Scan.option Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.defer)));  | 
|
| 8165 | 708  | 
|
| 24868 | 709  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
710  | 
Outer_Syntax.command "prefer" "shuffle internal proof state"  | 
| 36950 | 711  | 
(Keyword.tag_proof Keyword.prf_script)  | 
712  | 
(Parse.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.prefer)));  | 
|
| 8165 | 713  | 
|
| 24868 | 714  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
715  | 
Outer_Syntax.command "apply" "initial refinement step (unstructured)"  | 
| 36950 | 716  | 
(Keyword.tag_proof Keyword.prf_script)  | 
| 22117 | 717  | 
(Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply)));  | 
| 5832 | 718  | 
|
| 24868 | 719  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
720  | 
Outer_Syntax.command "apply_end" "terminal refinement (unstructured)"  | 
| 36950 | 721  | 
(Keyword.tag_proof Keyword.prf_script)  | 
| 22117 | 722  | 
(Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply_end)));  | 
| 5832 | 723  | 
|
| 24868 | 724  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
725  | 
Outer_Syntax.command "proof" "backward proof"  | 
| 36950 | 726  | 
(Keyword.tag_proof Keyword.prf_block)  | 
| 22117 | 727  | 
(Scan.option Method.parse >> (fn m => Toplevel.print o  | 
| 33390 | 728  | 
Toplevel.actual_proof (Proof_Node.applys (Proof.proof m)) o  | 
| 27563 | 729  | 
Toplevel.skip_proof (fn i => i + 1)));  | 
| 5832 | 730  | 
|
731  | 
||
| 6773 | 732  | 
(* calculational proof commands *)  | 
733  | 
||
| 6878 | 734  | 
val calc_args =  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
735  | 
  Scan.option (Parse.$$$ "(" |-- Parse.!!! ((Parse_Spec.xthms1 --| Parse.$$$ ")")));
 | 
| 6878 | 736  | 
|
| 24868 | 737  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
738  | 
Outer_Syntax.command "also" "combine calculation and current facts"  | 
| 36950 | 739  | 
(Keyword.tag_proof Keyword.prf_decl)  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
740  | 
(calc_args >> (Toplevel.proofs' o Calculation.also_cmd));  | 
| 6773 | 741  | 
|
| 24868 | 742  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
743  | 
Outer_Syntax.command "finally" "combine calculation and current facts, exhibit result"  | 
| 36950 | 744  | 
(Keyword.tag_proof Keyword.prf_chain)  | 
| 
36323
 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 
wenzelm 
parents: 
36317 
diff
changeset
 | 
745  | 
(calc_args >> (Toplevel.proofs' o Calculation.finally_cmd));  | 
| 6773 | 746  | 
|
| 24868 | 747  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
748  | 
Outer_Syntax.command "moreover" "augment calculation by current facts"  | 
| 36950 | 749  | 
(Keyword.tag_proof Keyword.prf_decl)  | 
| 17900 | 750  | 
(Scan.succeed (Toplevel.proof' Calculation.moreover));  | 
| 8562 | 751  | 
|
| 24868 | 752  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
753  | 
Outer_Syntax.command "ultimately" "augment calculation by current facts, exhibit result"  | 
| 36950 | 754  | 
(Keyword.tag_proof Keyword.prf_chain)  | 
| 17900 | 755  | 
(Scan.succeed (Toplevel.proof' Calculation.ultimately));  | 
| 8588 | 756  | 
|
| 6773 | 757  | 
|
| 6742 | 758  | 
(* proof navigation *)  | 
| 5944 | 759  | 
|
| 24868 | 760  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
761  | 
Outer_Syntax.command "back" "backtracking of proof command"  | 
| 36950 | 762  | 
(Keyword.tag_proof Keyword.prf_script)  | 
| 33390 | 763  | 
(Scan.succeed (Toplevel.print o Toplevel.actual_proof Proof_Node.back o Toplevel.skip_proof I));  | 
| 6742 | 764  | 
|
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
765  | 
|
| 
27614
 
f38c25d106a7
renamed IsarCmd.nested_command to OuterSyntax.prepare_command;
 
wenzelm 
parents: 
27575 
diff
changeset
 | 
766  | 
(* nested commands *)  | 
| 25578 | 767  | 
|
| 29309 | 768  | 
val props_text =  | 
| 
43775
 
b361c7d184e7
added Parse.properties (again) -- allow empty list like Parse_Value.properties but unlike Parse.properties of ef86de9c98aa;
 
wenzelm 
parents: 
43227 
diff
changeset
 | 
769  | 
Scan.optional Parse.properties [] -- Parse.position Parse.string  | 
| 36950 | 770  | 
>> (fn (props, (str, pos)) =>  | 
771  | 
(Position.of_properties (Position.default_properties pos props), str));  | 
|
| 29309 | 772  | 
|
| 25578 | 773  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
774  | 
Outer_Syntax.improper_command "Isabelle.command" "nested Isabelle command" Keyword.control  | 
| 29309 | 775  | 
(props_text :|-- (fn (pos, str) =>  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
776  | 
(case Outer_Syntax.parse pos str of  | 
| 
27838
 
0340fd7cccc3
Isabelle.command: inline former OuterSyntax.prepare_command;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
777  | 
[tr] => Scan.succeed (K tr)  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43775 
diff
changeset
 | 
778  | 
| _ => Scan.fail_with (K (fn () => "exactly one command expected")))  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43775 
diff
changeset
 | 
779  | 
handle ERROR msg => Scan.fail_with (K (fn () => msg))));  | 
| 25578 | 780  | 
|
781  | 
||
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
782  | 
|
| 5832 | 783  | 
(** diagnostic commands (for interactive mode only) **)  | 
784  | 
||
| 36950 | 785  | 
val opt_modes =  | 
786  | 
  Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.xname --| Parse.$$$ ")")) [];
 | 
|
787  | 
||
788  | 
val opt_bang = Scan.optional (Parse.$$$ "!" >> K true) false;  | 
|
| 8464 | 789  | 
|
| 24868 | 790  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
791  | 
Outer_Syntax.improper_command "pretty_setmargin" "change default margin for pretty printing"  | 
| 38870 | 792  | 
Keyword.diag (Parse.nat >>  | 
793  | 
(fn n => Toplevel.no_timing o Toplevel.imperative (fn () => Pretty.margin_default := n)));  | 
|
| 7124 | 794  | 
|
| 24868 | 795  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
796  | 
Outer_Syntax.improper_command "help" "print outer syntax commands" Keyword.diag  | 
| 
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
797  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax));  | 
| 21714 | 798  | 
|
| 24868 | 799  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
800  | 
Outer_Syntax.improper_command "print_commands" "print outer syntax commands" Keyword.diag  | 
| 
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
801  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative Outer_Syntax.print_outer_syntax));  | 
| 5832 | 802  | 
|
| 24868 | 803  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
804  | 
Outer_Syntax.improper_command "print_configs" "print configuration options" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
805  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_configs));  | 
| 23989 | 806  | 
|
| 24868 | 807  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
808  | 
Outer_Syntax.improper_command "print_context" "print theory context name" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
809  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_context));  | 
| 7308 | 810  | 
|
| 24868 | 811  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
812  | 
Outer_Syntax.improper_command "print_theory" "print logical theory contents (verbose!)"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
813  | 
Keyword.diag (opt_bang >> (Toplevel.no_timing oo Isar_Cmd.print_theory));  | 
| 5832 | 814  | 
|
| 24868 | 815  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
816  | 
Outer_Syntax.improper_command "print_syntax" "print inner syntax of context (verbose!)"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
817  | 
Keyword.diag (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_syntax));  | 
| 5832 | 818  | 
|
| 24868 | 819  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
820  | 
Outer_Syntax.improper_command "print_abbrevs" "print constant abbreviation of context"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
821  | 
Keyword.diag (Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_abbrevs));  | 
| 21726 | 822  | 
|
| 24868 | 823  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
824  | 
Outer_Syntax.improper_command "print_theorems"  | 
| 36950 | 825  | 
"print theorems of local theory or proof context" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
826  | 
(opt_bang >> (Toplevel.no_timing oo Isar_Cmd.print_theorems));  | 
| 5881 | 827  | 
|
| 24868 | 828  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
829  | 
Outer_Syntax.improper_command "print_locales" "print locales of this theory" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
830  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_locales));  | 
| 12061 | 831  | 
|
| 24868 | 832  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
833  | 
Outer_Syntax.improper_command "print_classes" "print classes of this theory" Keyword.diag  | 
| 42359 | 834  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o  | 
835  | 
Toplevel.keep (Class.print_classes o Toplevel.context_of)));  | 
|
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
836  | 
|
| 24868 | 837  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
838  | 
Outer_Syntax.improper_command "print_locale" "print locale of this theory" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
839  | 
(opt_bang -- Parse.xname >> (Toplevel.no_timing oo Isar_Cmd.print_locale));  | 
| 12061 | 840  | 
|
| 24868 | 841  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
842  | 
Outer_Syntax.improper_command "print_interps"  | 
| 41435 | 843  | 
"print interpretations of locale for this theory or proof context" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
844  | 
(Parse.xname >> (Toplevel.no_timing oo Isar_Cmd.print_registrations));  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
parents: 
31869 
diff
changeset
 | 
845  | 
|
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
parents: 
31869 
diff
changeset
 | 
846  | 
val _ =  | 
| 41435 | 847  | 
Outer_Syntax.improper_command "print_dependencies"  | 
848  | 
"print dependencies of locale expression" Keyword.diag  | 
|
849  | 
(opt_bang -- Parse_Spec.locale_expression true >>  | 
|
850  | 
(Toplevel.no_timing oo Isar_Cmd.print_dependencies));  | 
|
851  | 
||
852  | 
val _ =  | 
|
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
853  | 
Outer_Syntax.improper_command "print_attributes" "print attributes of this theory" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
854  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_attributes));  | 
| 5832 | 855  | 
|
| 24868 | 856  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
857  | 
Outer_Syntax.improper_command "print_simpset" "print context of Simplifier" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
858  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_simpset));  | 
| 16027 | 859  | 
|
| 24868 | 860  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
861  | 
Outer_Syntax.improper_command "print_rules" "print intro/elim rules" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
862  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_rules));  | 
| 12383 | 863  | 
|
| 24868 | 864  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
865  | 
Outer_Syntax.improper_command "print_trans_rules" "print transitivity rules" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
866  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_trans_rules));  | 
| 9221 | 867  | 
|
| 24868 | 868  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
869  | 
Outer_Syntax.improper_command "print_methods" "print methods of this theory" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
870  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_methods));  | 
| 5832 | 871  | 
|
| 24868 | 872  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
873  | 
Outer_Syntax.improper_command "print_antiquotations" "print antiquotations (global)" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
874  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_antiquotations));  | 
| 9221 | 875  | 
|
| 24868 | 876  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
877  | 
Outer_Syntax.improper_command "thy_deps" "visualize theory dependencies"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
878  | 
Keyword.diag (Scan.succeed (Toplevel.no_timing o Isar_Cmd.thy_deps));  | 
| 22485 | 879  | 
|
| 24868 | 880  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
881  | 
Outer_Syntax.improper_command "class_deps" "visualize class dependencies"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
882  | 
Keyword.diag (Scan.succeed (Toplevel.no_timing o Isar_Cmd.class_deps));  | 
| 20574 | 883  | 
|
| 24868 | 884  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
885  | 
Outer_Syntax.improper_command "thm_deps" "visualize theorem dependencies"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
886  | 
Keyword.diag (Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.thm_deps));  | 
| 9454 | 887  | 
|
| 24868 | 888  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
889  | 
Outer_Syntax.improper_command "print_binds" "print term bindings of proof context" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
890  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_binds));  | 
| 5832 | 891  | 
|
| 24868 | 892  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
893  | 
Outer_Syntax.improper_command "print_facts" "print facts of proof context" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
894  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_facts));  | 
| 5832 | 895  | 
|
| 24868 | 896  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
897  | 
Outer_Syntax.improper_command "print_cases" "print cases of proof context" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
898  | 
(Scan.succeed (Toplevel.no_timing o Isar_Cmd.print_cases));  | 
| 8370 | 899  | 
|
| 24868 | 900  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
901  | 
Outer_Syntax.improper_command "print_statement" "print theorems as long statements" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
902  | 
(opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.print_stmts));  | 
| 19269 | 903  | 
|
| 24868 | 904  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
905  | 
Outer_Syntax.improper_command "thm" "print theorems" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
906  | 
(opt_modes -- Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.print_thms));  | 
| 5832 | 907  | 
|
| 24868 | 908  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
909  | 
Outer_Syntax.improper_command "prf" "print proof terms of theorems" Keyword.diag  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
910  | 
(opt_modes -- Scan.option Parse_Spec.xthms1  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
911  | 
>> (Toplevel.no_timing oo Isar_Cmd.print_prfs false));  | 
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
912  | 
|
| 24868 | 913  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
914  | 
Outer_Syntax.improper_command "full_prf" "print full proof terms of theorems" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
915  | 
(opt_modes -- Scan.option Parse_Spec.xthms1 >> (Toplevel.no_timing oo Isar_Cmd.print_prfs true));  | 
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
916  | 
|
| 24868 | 917  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
918  | 
Outer_Syntax.improper_command "prop" "read and print proposition" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
919  | 
(opt_modes -- Parse.term >> (Toplevel.no_timing oo Isar_Cmd.print_prop));  | 
| 5832 | 920  | 
|
| 24868 | 921  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
922  | 
Outer_Syntax.improper_command "term" "read and print term" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
923  | 
(opt_modes -- Parse.term >> (Toplevel.no_timing oo Isar_Cmd.print_term));  | 
| 5832 | 924  | 
|
| 24868 | 925  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
926  | 
Outer_Syntax.improper_command "typ" "read and print type" Keyword.diag  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
927  | 
(opt_modes -- Parse.typ >> (Toplevel.no_timing oo Isar_Cmd.print_type));  | 
| 5832 | 928  | 
|
| 24868 | 929  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
930  | 
Outer_Syntax.improper_command "print_codesetup" "print code generator setup" Keyword.diag  | 
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
931  | 
(Scan.succeed  | 
| 
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
932  | 
(Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep  | 
| 24219 | 933  | 
(Code.print_codesetup o Toplevel.theory_of)));  | 
| 5832 | 934  | 
|
| 26184 | 935  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
936  | 
Outer_Syntax.improper_command "unused_thms" "find unused theorems" Keyword.diag  | 
| 36950 | 937  | 
(Scan.option ((Scan.repeat1 (Scan.unless Parse.minus Parse.name) --| Parse.minus) --  | 
938  | 
Scan.option (Scan.repeat1 (Scan.unless Parse.minus Parse.name))) >>  | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
939  | 
(Toplevel.no_timing oo Isar_Cmd.unused_thms));  | 
| 26184 | 940  | 
|
| 5832 | 941  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
942  | 
|
| 5832 | 943  | 
(** system commands (for interactive mode only) **)  | 
944  | 
||
| 24868 | 945  | 
val _ =  | 
| 
40395
 
4985aaade799
mark 'cd' and 'commit' as control command -- not usable in asynchronous document model, likely to cause confusion in Proof General;
 
wenzelm 
parents: 
39214 
diff
changeset
 | 
946  | 
Outer_Syntax.improper_command "cd" "change current working directory" Keyword.control  | 
| 
 
4985aaade799
mark 'cd' and 'commit' as control command -- not usable in asynchronous document model, likely to cause confusion in Proof General;
 
wenzelm 
parents: 
39214 
diff
changeset
 | 
947  | 
(Parse.path >> (fn path => Toplevel.no_timing o Toplevel.imperative (fn () => File.cd path)));  | 
| 5832 | 948  | 
|
| 24868 | 949  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
950  | 
Outer_Syntax.improper_command "pwd" "print current working directory" Keyword.diag  | 
| 
40395
 
4985aaade799
mark 'cd' and 'commit' as control command -- not usable in asynchronous document model, likely to cause confusion in Proof General;
 
wenzelm 
parents: 
39214 
diff
changeset
 | 
951  | 
(Scan.succeed (Toplevel.no_timing o  | 
| 
41944
 
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
 
wenzelm 
parents: 
41585 
diff
changeset
 | 
952  | 
Toplevel.imperative (fn () => writeln (Path.print (File.pwd ())))));  | 
| 5832 | 953  | 
|
| 24868 | 954  | 
val _ =  | 
| 
37987
 
aac4eb1fa1d8
explicit Keyword.control markup for various control commands -- to prevent them from occurring in proof documents;
 
wenzelm 
parents: 
37981 
diff
changeset
 | 
955  | 
Outer_Syntax.improper_command "use_thy" "use theory file" Keyword.control  | 
| 36950 | 956  | 
(Parse.name >> (fn name =>  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
957  | 
Toplevel.no_timing o Toplevel.imperative (fn () => Thy_Info.use_thy name)));  | 
| 5832 | 958  | 
|
| 24868 | 959  | 
val _ =  | 
| 
37987
 
aac4eb1fa1d8
explicit Keyword.control markup for various control commands -- to prevent them from occurring in proof documents;
 
wenzelm 
parents: 
37981 
diff
changeset
 | 
960  | 
Outer_Syntax.improper_command "remove_thy" "remove theory from loader database" Keyword.control  | 
| 36950 | 961  | 
(Parse.name >> (fn name =>  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
962  | 
Toplevel.no_timing o Toplevel.imperative (fn () => Thy_Info.remove_thy name)));  | 
| 7102 | 963  | 
|
| 24868 | 964  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
965  | 
Outer_Syntax.improper_command "kill_thy" "kill theory -- try to remove from loader database"  | 
| 
37987
 
aac4eb1fa1d8
explicit Keyword.control markup for various control commands -- to prevent them from occurring in proof documents;
 
wenzelm 
parents: 
37981 
diff
changeset
 | 
966  | 
Keyword.control (Parse.name >> (fn name =>  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
967  | 
Toplevel.no_timing o Toplevel.imperative (fn () => Thy_Info.kill_thy name)));  | 
| 7931 | 968  | 
|
| 24868 | 969  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
970  | 
Outer_Syntax.improper_command "display_drafts" "display raw source files with symbols"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
971  | 
Keyword.diag (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo Isar_Cmd.display_drafts));  | 
| 14934 | 972  | 
|
| 24868 | 973  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
974  | 
Outer_Syntax.improper_command "print_drafts" "print raw source files with symbols"  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
975  | 
Keyword.diag (Scan.repeat1 Parse.path >> (Toplevel.no_timing oo Isar_Cmd.print_drafts));  | 
| 14934 | 976  | 
|
| 39165 | 977  | 
val _ = (* FIXME tty only *)  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
978  | 
Outer_Syntax.improper_command "pr" "print current proof state (if present)" Keyword.diag  | 
| 39165 | 979  | 
(opt_modes -- Scan.option Parse.nat >> (fn (modes, limit) =>  | 
980  | 
Toplevel.no_timing o Toplevel.keep (fn state =>  | 
|
981  | 
(case limit of NONE => () | SOME n => Goal_Display.goals_limit_default := n;  | 
|
982  | 
Toplevel.quiet := false;  | 
|
983  | 
Print_Mode.with_modes modes (Toplevel.print_state true) state))));  | 
|
| 7199 | 984  | 
|
| 24868 | 985  | 
val _ =  | 
| 
37987
 
aac4eb1fa1d8
explicit Keyword.control markup for various control commands -- to prevent them from occurring in proof documents;
 
wenzelm 
parents: 
37981 
diff
changeset
 | 
986  | 
Outer_Syntax.improper_command "disable_pr" "disable printing of toplevel state" Keyword.control  | 
| 39165 | 987  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Toplevel.quiet := true)));  | 
| 5832 | 988  | 
|
| 24868 | 989  | 
val _ =  | 
| 
37987
 
aac4eb1fa1d8
explicit Keyword.control markup for various control commands -- to prevent them from occurring in proof documents;
 
wenzelm 
parents: 
37981 
diff
changeset
 | 
990  | 
Outer_Syntax.improper_command "enable_pr" "enable printing of toplevel state" Keyword.control  | 
| 39165 | 991  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative (fn () => Toplevel.quiet := false)));  | 
| 7222 | 992  | 
|
| 24868 | 993  | 
val _ =  | 
| 
40395
 
4985aaade799
mark 'cd' and 'commit' as control command -- not usable in asynchronous document model, likely to cause confusion in Proof General;
 
wenzelm 
parents: 
39214 
diff
changeset
 | 
994  | 
Outer_Syntax.improper_command "commit" "commit current session to ML database" Keyword.control  | 
| 36950 | 995  | 
(Parse.opt_unit >> K (Toplevel.no_timing o Toplevel.imperative Secure.commit));  | 
| 5832 | 996  | 
|
| 24868 | 997  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
998  | 
Outer_Syntax.improper_command "quit" "quit Isabelle" Keyword.control  | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
999  | 
(Parse.opt_unit >> (Toplevel.no_timing oo K (Toplevel.imperative quit)));  | 
| 5832 | 1000  | 
|
| 24868 | 1001  | 
val _ =  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
1002  | 
Outer_Syntax.improper_command "exit" "exit Isar loop" Keyword.control  | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
1003  | 
(Scan.succeed  | 
| 
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
1004  | 
(Toplevel.no_timing o Toplevel.keep (fn state =>  | 
| 
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
1005  | 
(Context.set_thread_data (try Toplevel.generic_theory_of state);  | 
| 
38888
 
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
 
wenzelm 
parents: 
38870 
diff
changeset
 | 
1006  | 
raise Runtime.TERMINATE))));  | 
| 5832 | 1007  | 
|
1008  | 
end;  | 
|
| 
27614
 
f38c25d106a7
renamed IsarCmd.nested_command to OuterSyntax.prepare_command;
 
wenzelm 
parents: 
27575 
diff
changeset
 | 
1009  |