| author | wenzelm | 
| Wed, 26 Mar 2008 20:14:37 +0100 | |
| changeset 26404 | 56fd70fb7571 | 
| parent 26396 | e44c5a1a47bd | 
| child 26435 | bdce320cd426 | 
| permissions | -rw-r--r-- | 
| 5832 | 1  | 
(* Title: Pure/Isar/isar_syn.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
4  | 
||
| 6353 | 5  | 
Isar/Pure outer syntax.  | 
| 5832 | 6  | 
*)  | 
7  | 
||
| 17353 | 8  | 
structure IsarSyn: sig end =  | 
| 5832 | 9  | 
struct  | 
10  | 
||
| 17068 | 11  | 
structure P = OuterParse and K = OuterKeyword;  | 
| 5832 | 12  | 
|
13  | 
||
| 24868 | 14  | 
(** keywords **)  | 
15  | 
||
16  | 
(*keep keywords consistent with the parsers, otherwise be prepared for  | 
|
17  | 
unexpected errors*)  | 
|
18  | 
||
19  | 
val _ = OuterSyntax.keywords  | 
|
20  | 
 ["!!", "!", "%", "(", ")", "+", ",", "--", ":", "::", ";", "<", "<=",
 | 
|
21  | 
"=", "==", "=>", "?", "[", "\\<equiv>", "\\<leftharpoondown>",  | 
|
22  | 
"\\<rightharpoonup>", "\\<rightleftharpoons>", "\\<subseteq>", "]",  | 
|
23  | 
"advanced", "and", "assumes", "attach", "begin", "binder", "concl",  | 
|
24  | 
"constrains", "defines", "fixes", "for", "identifier", "if",  | 
|
25  | 
"imports", "in", "includes", "infix", "infixl", "infixr", "is",  | 
|
| 25003 | 26  | 
"notes", "obtains", "open", "output", "overloaded", "shows",  | 
| 24868 | 27  | 
"structure", "unchecked", "uses", "where", "|"];  | 
28  | 
||
29  | 
||
| 5832 | 30  | 
(** init and exit **)  | 
31  | 
||
| 24868 | 32  | 
val _ =  | 
| 17068 | 33  | 
OuterSyntax.command "theory" "begin theory" (K.tag_theory K.thy_begin)  | 
| 21350 | 34  | 
(ThyHeader.args >> (Toplevel.print oo IsarCmd.theory));  | 
| 5832 | 35  | 
|
| 24868 | 36  | 
val _ =  | 
| 20958 | 37  | 
OuterSyntax.command "end" "end (local) theory" (K.tag_theory K.thy_end)  | 
| 21004 | 38  | 
(Scan.succeed (Toplevel.exit o Toplevel.end_local_theory));  | 
| 6687 | 39  | 
|
| 5832 | 40  | 
|
| 7749 | 41  | 
(** markup commands **)  | 
| 5832 | 42  | 
|
| 24868 | 43  | 
val _ = OuterSyntax.markup_command ThyOutput.Markup "header" "theory header" K.diag  | 
| 12940 | 44  | 
(P.position P.text >> IsarCmd.add_header);  | 
| 6353 | 45  | 
|
| 24868 | 46  | 
val _ = OuterSyntax.markup_command ThyOutput.Markup "chapter" "chapter heading"  | 
| 22117 | 47  | 
K.thy_heading (P.opt_target -- P.position P.text >> IsarCmd.add_chapter);  | 
| 5958 | 48  | 
|
| 24868 | 49  | 
val _ = OuterSyntax.markup_command ThyOutput.Markup "section" "section heading"  | 
| 22117 | 50  | 
K.thy_heading (P.opt_target -- P.position P.text >> IsarCmd.add_section);  | 
| 5958 | 51  | 
|
| 24868 | 52  | 
val _ = OuterSyntax.markup_command ThyOutput.Markup "subsection" "subsection heading"  | 
| 22117 | 53  | 
K.thy_heading (P.opt_target -- P.position P.text >> IsarCmd.add_subsection);  | 
| 5958 | 54  | 
|
| 24868 | 55  | 
val _ =  | 
| 22117 | 56  | 
OuterSyntax.markup_command ThyOutput.Markup "subsubsection" "subsubsection heading"  | 
57  | 
K.thy_heading (P.opt_target -- P.position P.text >> IsarCmd.add_subsubsection);  | 
|
| 5832 | 58  | 
|
| 24868 | 59  | 
val _ = OuterSyntax.markup_command ThyOutput.MarkupEnv "text" "formal comment (theory)"  | 
| 22117 | 60  | 
K.thy_decl (P.opt_target -- P.position P.text >> IsarCmd.add_text);  | 
| 7172 | 61  | 
|
| 24868 | 62  | 
val _ = OuterSyntax.markup_command ThyOutput.Verbatim "text_raw"  | 
| 
17264
 
c5b280a52a67
chapter/section/subsection/subsubsection/text: optional locale specification;
 
wenzelm 
parents: 
17228 
diff
changeset
 | 
63  | 
"raw document preparation text"  | 
| 
 
c5b280a52a67
chapter/section/subsection/subsubsection/text: optional locale specification;
 
wenzelm 
parents: 
17228 
diff
changeset
 | 
64  | 
K.thy_decl (P.position P.text >> IsarCmd.add_text_raw);  | 
| 7172 | 65  | 
|
| 24868 | 66  | 
val _ = OuterSyntax.markup_command ThyOutput.Markup "sect" "formal comment (proof)"  | 
| 17068 | 67  | 
(K.tag_proof K.prf_heading) (P.position P.text >> IsarCmd.add_sect);  | 
| 7172 | 68  | 
|
| 24868 | 69  | 
val _ = OuterSyntax.markup_command ThyOutput.Markup "subsect" "formal comment (proof)"  | 
| 17068 | 70  | 
(K.tag_proof K.prf_heading) (P.position P.text >> IsarCmd.add_subsect);  | 
| 7172 | 71  | 
|
| 24868 | 72  | 
val _ = OuterSyntax.markup_command ThyOutput.Markup "subsubsect"  | 
| 17068 | 73  | 
"formal comment (proof)" (K.tag_proof K.prf_heading)  | 
| 12940 | 74  | 
(P.position P.text >> IsarCmd.add_subsubsect);  | 
| 7172 | 75  | 
|
| 24868 | 76  | 
val _ = OuterSyntax.markup_command ThyOutput.MarkupEnv "txt" "formal comment (proof)"  | 
| 17068 | 77  | 
(K.tag_proof K.prf_decl) (P.position P.text >> IsarCmd.add_txt);  | 
| 7172 | 78  | 
|
| 24868 | 79  | 
val _ = OuterSyntax.markup_command ThyOutput.Verbatim "txt_raw"  | 
| 17068 | 80  | 
"raw document preparation text (proof)" (K.tag_proof K.prf_decl)  | 
| 12940 | 81  | 
(P.position P.text >> IsarCmd.add_txt_raw);  | 
| 7775 | 82  | 
|
| 5832 | 83  | 
|
| 6886 | 84  | 
|
85  | 
(** theory sections **)  | 
|
86  | 
||
| 5832 | 87  | 
(* classes and sorts *)  | 
88  | 
||
| 24868 | 89  | 
val _ =  | 
| 6723 | 90  | 
OuterSyntax.command "classes" "declare type classes" K.thy_decl  | 
| 
11101
 
014e7b5c77ba
support \<subseteq> syntax in classes/classrel/axclass/instance;
 
wenzelm 
parents: 
11017 
diff
changeset
 | 
91  | 
(Scan.repeat1 (P.name -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
92  | 
P.!!! (P.list1 P.xname)) []) >> (Toplevel.theory o fold AxClass.axiomatize_class_cmd));  | 
| 5832 | 93  | 
|
| 24868 | 94  | 
val _ =  | 
| 6723 | 95  | 
OuterSyntax.command "classrel" "state inclusion of type classes (axiomatic!)" K.thy_decl  | 
| 14779 | 96  | 
(P.and_list1 (P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname))  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
97  | 
>> (Toplevel.theory o AxClass.axiomatize_classrel_cmd));  | 
| 5832 | 98  | 
|
| 24868 | 99  | 
val _ =  | 
| 6723 | 100  | 
OuterSyntax.command "defaultsort" "declare default sort" K.thy_decl  | 
| 22796 | 101  | 
(P.sort >> (Toplevel.theory o Sign.add_defsort));  | 
| 5832 | 102  | 
|
| 24868 | 103  | 
val _ =  | 
| 19245 | 104  | 
OuterSyntax.command "axclass" "define axiomatic type class" K.thy_decl  | 
| 
21462
 
74ddf3a522f8
added Isar syntax for adding parameters to axclasses
 
haftmann 
parents: 
21437 
diff
changeset
 | 
105  | 
(P.name -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--  | 
| 
 
74ddf3a522f8
added Isar syntax for adding parameters to axclasses
 
haftmann 
parents: 
21437 
diff
changeset
 | 
106  | 
P.!!! (P.list1 P.xname)) []  | 
| 22117 | 107  | 
-- Scan.repeat (SpecParse.thm_name ":" -- (P.prop >> single))  | 
| 24219 | 108  | 
>> (fn (x, y) => Toplevel.theory (snd o Class.axclass_cmd x y)));  | 
| 19245 | 109  | 
|
| 5832 | 110  | 
|
111  | 
(* types *)  | 
|
112  | 
||
| 24868 | 113  | 
val _ =  | 
| 12624 | 114  | 
OuterSyntax.command "typedecl" "type declaration" K.thy_decl  | 
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
115  | 
(P.type_args -- P.name -- P.opt_infix >> (fn ((args, a), mx) =>  | 
| 25495 | 116  | 
Toplevel.theory (ObjectLogic.typedecl (a, args, mx) #> snd)));  | 
| 5832 | 117  | 
|
| 24868 | 118  | 
val _ =  | 
| 6723 | 119  | 
OuterSyntax.command "types" "declare type abbreviations" K.thy_decl  | 
| 6727 | 120  | 
(Scan.repeat1  | 
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
121  | 
(P.type_args -- P.name -- (P.$$$ "=" |-- P.!!! (P.typ -- P.opt_infix')))  | 
| 22796 | 122  | 
>> (Toplevel.theory o Sign.add_tyabbrs o  | 
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
123  | 
map (fn ((args, a), (T, mx)) => (a, args, T, mx))));  | 
| 5832 | 124  | 
|
| 24868 | 125  | 
val _ =  | 
| 6370 | 126  | 
OuterSyntax.command "nonterminals" "declare types treated as grammar nonterminal symbols"  | 
| 22796 | 127  | 
K.thy_decl (Scan.repeat1 P.name >> (Toplevel.theory o Sign.add_nonterminals));  | 
| 5832 | 128  | 
|
| 24868 | 129  | 
val _ =  | 
| 6723 | 130  | 
OuterSyntax.command "arities" "state type arities (axiomatic!)" K.thy_decl  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
131  | 
(Scan.repeat1 P.arity >> (Toplevel.theory o fold AxClass.axiomatize_arity_cmd));  | 
| 5832 | 132  | 
|
133  | 
||
134  | 
(* consts and syntax *)  | 
|
135  | 
||
| 24868 | 136  | 
val _ =  | 
| 8227 | 137  | 
OuterSyntax.command "judgment" "declare object-logic judgment" K.thy_decl  | 
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
138  | 
(P.const >> (Toplevel.theory o ObjectLogic.add_judgment));  | 
| 8227 | 139  | 
|
| 24868 | 140  | 
val _ =  | 
| 6723 | 141  | 
OuterSyntax.command "consts" "declare constants" K.thy_decl  | 
| 22796 | 142  | 
(Scan.repeat1 P.const >> (Toplevel.theory o Sign.add_consts));  | 
| 5832 | 143  | 
|
| 14642 | 144  | 
val opt_overloaded = P.opt_keyword "overloaded";  | 
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
13802 
diff
changeset
 | 
145  | 
|
| 24868 | 146  | 
val _ =  | 
| 
14223
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
13802 
diff
changeset
 | 
147  | 
OuterSyntax.command "finalconsts" "declare constants as final" K.thy_decl  | 
| 
 
0ee05eef881b
Added support for making constants final, that is, ensuring that no
 
skalberg 
parents: 
13802 
diff
changeset
 | 
148  | 
(opt_overloaded -- Scan.repeat1 P.term >> (uncurry (Toplevel.theory oo Theory.add_finals)));  | 
| 9731 | 149  | 
|
| 12142 | 150  | 
val mode_spec =  | 
| 9731 | 151  | 
  (P.$$$ "output" >> K ("", false)) || P.name -- Scan.optional (P.$$$ "output" >> K false) true;
 | 
152  | 
||
| 14900 | 153  | 
val opt_mode =  | 
| 24960 | 154  | 
  Scan.optional (P.$$$ "(" |-- P.!!! (mode_spec --| P.$$$ ")")) Syntax.mode_default;
 | 
| 5832 | 155  | 
|
| 24868 | 156  | 
val _ =  | 
| 6723 | 157  | 
OuterSyntax.command "syntax" "declare syntactic constants" K.thy_decl  | 
| 22796 | 158  | 
(opt_mode -- Scan.repeat1 P.const >> (Toplevel.theory o uncurry Sign.add_modesyntax));  | 
| 5832 | 159  | 
|
| 24868 | 160  | 
val _ =  | 
| 15748 | 161  | 
OuterSyntax.command "no_syntax" "delete syntax declarations" K.thy_decl  | 
| 22796 | 162  | 
(opt_mode -- Scan.repeat1 P.const >> (Toplevel.theory o uncurry Sign.del_modesyntax));  | 
| 15748 | 163  | 
|
| 5832 | 164  | 
|
165  | 
(* translations *)  | 
|
166  | 
||
167  | 
val trans_pat =  | 
|
| 6723 | 168  | 
  Scan.optional (P.$$$ "(" |-- P.!!! (P.xname --| P.$$$ ")")) "logic" -- P.string;
 | 
| 5832 | 169  | 
|
170  | 
fun trans_arrow toks =  | 
|
| 10688 | 171  | 
((P.$$$ "\\<rightharpoonup>" || P.$$$ "=>") >> K Syntax.ParseRule ||  | 
172  | 
(P.$$$ "\\<leftharpoondown>" || P.$$$ "<=") >> K Syntax.PrintRule ||  | 
|
173  | 
(P.$$$ "\\<rightleftharpoons>" || P.$$$ "==") >> K Syntax.ParsePrintRule) toks;  | 
|
| 5832 | 174  | 
|
175  | 
val trans_line =  | 
|
| 6723 | 176  | 
trans_pat -- P.!!! (trans_arrow -- trans_pat)  | 
| 5832 | 177  | 
>> (fn (left, (arr, right)) => arr (left, right));  | 
178  | 
||
| 24868 | 179  | 
val _ =  | 
| 6723 | 180  | 
OuterSyntax.command "translations" "declare syntax translation rules" K.thy_decl  | 
| 22796 | 181  | 
(Scan.repeat1 trans_line >> (Toplevel.theory o Sign.add_trrules));  | 
| 5832 | 182  | 
|
| 24868 | 183  | 
val _ =  | 
| 19260 | 184  | 
OuterSyntax.command "no_translations" "remove syntax translation rules" K.thy_decl  | 
| 22796 | 185  | 
(Scan.repeat1 trans_line >> (Toplevel.theory o Sign.del_trrules));  | 
| 19260 | 186  | 
|
| 5832 | 187  | 
|
188  | 
(* axioms and definitions *)  | 
|
189  | 
||
| 24868 | 190  | 
val _ =  | 
| 6723 | 191  | 
OuterSyntax.command "axioms" "state arbitrary propositions (axiomatic!)" K.thy_decl  | 
| 22117 | 192  | 
(Scan.repeat1 SpecParse.spec_name >> (Toplevel.theory o IsarCmd.add_axioms));  | 
| 5832 | 193  | 
|
| 19631 | 194  | 
val opt_unchecked_overloaded =  | 
195  | 
  Scan.optional (P.$$$ "(" |-- P.!!!
 | 
|
196  | 
(((P.$$$ "unchecked" >> K true) -- Scan.optional (P.$$$ "overloaded" >> K true) false ||  | 
|
197  | 
P.$$$ "overloaded" >> K (false, true)) --| P.$$$ ")")) (false, false);  | 
|
198  | 
||
| 24868 | 199  | 
val _ =  | 
| 6723 | 200  | 
OuterSyntax.command "defs" "define constants" K.thy_decl  | 
| 22117 | 201  | 
(opt_unchecked_overloaded -- Scan.repeat1 SpecParse.spec_name  | 
| 21350 | 202  | 
>> (Toplevel.theory o IsarCmd.add_defs));  | 
| 6370 | 203  | 
|
| 14642 | 204  | 
|
| 
21601
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
205  | 
(* old constdefs *)  | 
| 14642 | 206  | 
|
| 
21601
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
207  | 
val old_constdecl =  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
208  | 
P.name --| P.where_ >> (fn x => (x, NONE, NoSyn)) ||  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
209  | 
P.name -- (P.$$$ "::" |-- P.!!! P.typ >> SOME) -- P.opt_mixfix'  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
210  | 
--| Scan.option P.where_ >> P.triple1 ||  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
211  | 
P.name -- (P.mixfix >> pair NONE) --| Scan.option P.where_ >> P.triple2;  | 
| 19076 | 212  | 
|
| 22117 | 213  | 
val old_constdef = Scan.option old_constdecl -- (SpecParse.opt_thm_name ":" -- P.prop);  | 
| 14642 | 214  | 
|
| 19076 | 215  | 
val structs =  | 
216  | 
  Scan.optional ((P.$$$ "(" -- P.$$$ "structure") |-- P.!!! (P.simple_fixes --| P.$$$ ")")) [];
 | 
|
217  | 
||
| 24868 | 218  | 
val _ =  | 
| 19076 | 219  | 
OuterSyntax.command "constdefs" "old-style constant definition" K.thy_decl  | 
| 
21601
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
220  | 
(structs -- Scan.repeat1 old_constdef >> (Toplevel.theory o Constdefs.add_constdefs));  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
221  | 
|
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
222  | 
|
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
223  | 
(* constant definitions and abbreviations *)  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
224  | 
|
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
225  | 
val constdecl =  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
226  | 
P.name --  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
227  | 
(P.where_ >> K (NONE, NoSyn) ||  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
228  | 
P.$$$ "::" |-- P.!!! ((P.typ >> SOME) -- P.opt_mixfix' --| P.where_) ||  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
229  | 
      Scan.ahead (P.$$$ "(") |-- P.!!! (P.mixfix' --| P.where_ >> pair NONE))
 | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
230  | 
>> P.triple2;  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
231  | 
|
| 22117 | 232  | 
val constdef = Scan.option constdecl -- (SpecParse.opt_thm_name ":" -- P.prop);  | 
| 5832 | 233  | 
|
| 24868 | 234  | 
val _ =  | 
| 19076 | 235  | 
OuterSyntax.command "definition" "constant definition" K.thy_decl  | 
| 22117 | 236  | 
(P.opt_target -- constdef  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
237  | 
>> (fn (loc, args) => Toplevel.local_theory loc (#2 o Specification.definition_cmd args)));  | 
| 18780 | 238  | 
|
| 24868 | 239  | 
val _ =  | 
| 19076 | 240  | 
OuterSyntax.command "abbreviation" "constant abbreviation" K.thy_decl  | 
| 22117 | 241  | 
(P.opt_target -- opt_mode -- (Scan.option constdecl -- P.prop)  | 
| 19659 | 242  | 
>> (fn ((loc, mode), args) =>  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
243  | 
Toplevel.local_theory loc (Specification.abbreviation_cmd mode args)));  | 
| 19659 | 244  | 
|
| 24868 | 245  | 
val _ =  | 
| 24950 | 246  | 
OuterSyntax.command "notation" "add notation for constants / fixed variables" K.thy_decl  | 
| 24953 | 247  | 
(P.opt_target -- opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)  | 
| 19659 | 248  | 
>> (fn ((loc, mode), args) =>  | 
| 24950 | 249  | 
Toplevel.local_theory loc (Specification.notation_cmd true mode args)));  | 
250  | 
||
251  | 
val _ =  | 
|
252  | 
OuterSyntax.command "no_notation" "delete notation for constants / fixed variables" K.thy_decl  | 
|
| 24953 | 253  | 
(P.opt_target -- opt_mode -- P.and_list1 (P.xname -- SpecParse.locale_mixfix)  | 
| 24950 | 254  | 
>> (fn ((loc, mode), args) =>  | 
255  | 
Toplevel.local_theory loc (Specification.notation_cmd false mode args)));  | 
|
| 19076 | 256  | 
|
| 5832 | 257  | 
|
| 18616 | 258  | 
(* constant specifications *)  | 
259  | 
||
| 24868 | 260  | 
val _ =  | 
| 18616 | 261  | 
OuterSyntax.command "axiomatization" "axiomatic constant specification" K.thy_decl  | 
| 22117 | 262  | 
(P.opt_target --  | 
| 18949 | 263  | 
(Scan.optional P.fixes [] --  | 
| 22117 | 264  | 
Scan.optional (P.where_ |-- P.!!! (P.and_list1 SpecParse.spec)) [])  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
265  | 
>> (fn (loc, (x, y)) => Toplevel.local_theory loc (#2 o Specification.axiomatization_cmd x y)));  | 
| 18616 | 266  | 
|
267  | 
||
| 5914 | 268  | 
(* theorems *)  | 
269  | 
||
| 22117 | 270  | 
fun theorems kind = P.opt_target -- SpecParse.name_facts  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
271  | 
>> (fn (loc, args) => Toplevel.local_theory loc (#2 o Specification.theorems_cmd kind args));  | 
| 12712 | 272  | 
|
| 24868 | 273  | 
val _ =  | 
| 21437 | 274  | 
OuterSyntax.command "theorems" "define theorems" K.thy_decl (theorems Thm.theoremK);  | 
| 5914 | 275  | 
|
| 24868 | 276  | 
val _ =  | 
| 21437 | 277  | 
OuterSyntax.command "lemmas" "define lemmas" K.thy_decl (theorems Thm.lemmaK);  | 
| 5914 | 278  | 
|
| 24868 | 279  | 
val _ =  | 
| 23795 | 280  | 
OuterSyntax.command "declare" "declare theorems (improper)" K.thy_decl  | 
| 22117 | 281  | 
(P.opt_target -- (P.and_list1 SpecParse.xthms1 >> flat)  | 
| 20907 | 282  | 
>> (fn (loc, args) => Toplevel.local_theory loc  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
283  | 
          (#2 o Specification.theorems_cmd "" [(("", []), args)])));
 | 
| 9589 | 284  | 
|
| 5914 | 285  | 
|
| 5832 | 286  | 
(* name space entry path *)  | 
287  | 
||
| 24868 | 288  | 
val _ =  | 
| 6723 | 289  | 
OuterSyntax.command "global" "disable prefixing of theory name" K.thy_decl  | 
| 16447 | 290  | 
(Scan.succeed (Toplevel.theory Sign.root_path));  | 
| 5832 | 291  | 
|
| 24868 | 292  | 
val _ =  | 
| 6723 | 293  | 
OuterSyntax.command "local" "enable prefixing of theory name" K.thy_decl  | 
| 16447 | 294  | 
(Scan.succeed (Toplevel.theory Sign.local_path));  | 
| 8723 | 295  | 
|
| 24868 | 296  | 
val _ =  | 
| 8723 | 297  | 
OuterSyntax.command "hide" "hide names from given name space" K.thy_decl  | 
| 17397 | 298  | 
((P.opt_keyword "open" >> not) -- (P.name -- Scan.repeat1 P.xname) >>  | 
299  | 
(Toplevel.theory o uncurry Sign.hide_names));  | 
|
| 5832 | 300  | 
|
301  | 
||
302  | 
(* use ML text *)  | 
|
303  | 
||
| 24868 | 304  | 
val _ =  | 
| 17068 | 305  | 
OuterSyntax.command "use" "eval ML text from file" (K.tag_ml K.diag)  | 
| 14950 | 306  | 
(P.path >> (Toplevel.no_timing oo IsarCmd.use));  | 
| 5832 | 307  | 
|
| 24868 | 308  | 
val _ =  | 
| 17068 | 309  | 
OuterSyntax.command "ML" "eval ML text (diagnostic)" (K.tag_ml K.diag)  | 
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
310  | 
(P.position P.text >> IsarCmd.use_mltext true);  | 
| 7891 | 311  | 
|
| 24868 | 312  | 
val _ =  | 
| 26396 | 313  | 
OuterSyntax.command "ML_val" "eval ML text (diagnostic)" (K.tag_ml K.diag)  | 
314  | 
(P.position P.text >> IsarCmd.use_mltext true);  | 
|
315  | 
||
316  | 
val _ =  | 
|
317  | 
OuterSyntax.command "ML_command" "silently eval ML text (diagnostic)" (K.tag_ml K.diag)  | 
|
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
318  | 
(P.position P.text >> (Toplevel.no_timing oo IsarCmd.use_mltext false));  | 
| 7616 | 319  | 
|
| 24868 | 320  | 
val _ =  | 
| 17068 | 321  | 
OuterSyntax.command "ML_setup" "eval ML text (may change theory)" (K.tag_ml K.thy_decl)  | 
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
322  | 
(P.position P.text >> IsarCmd.use_mltext_theory);  | 
| 5832 | 323  | 
|
| 24868 | 324  | 
val _ =  | 
| 17068 | 325  | 
OuterSyntax.command "setup" "apply ML theory setup" (K.tag_ml K.thy_decl)  | 
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
326  | 
(Scan.option (P.position P.text) >> (Toplevel.theory o IsarCmd.generic_setup));  | 
| 5832 | 327  | 
|
| 24868 | 328  | 
val _ =  | 
| 17068 | 329  | 
OuterSyntax.command "method_setup" "define proof method in ML" (K.tag_ml K.thy_decl)  | 
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
330  | 
(P.name -- P.!!! (P.$$$ "=" |-- P.position P.text -- P.text)  | 
| 
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
331  | 
>> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt)));  | 
| 9197 | 332  | 
|
| 24868 | 333  | 
val _ =  | 
| 22088 | 334  | 
OuterSyntax.command "declaration" "generic ML declaration" (K.tag_ml K.thy_decl)  | 
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
335  | 
(P.opt_target -- P.position P.text  | 
| 22088 | 336  | 
>> (fn (loc, txt) => Toplevel.local_theory loc (IsarCmd.declaration txt)));  | 
337  | 
||
| 24868 | 338  | 
val _ =  | 
| 22202 | 339  | 
OuterSyntax.command "simproc_setup" "define simproc in ML" (K.tag_ml K.thy_decl)  | 
340  | 
(P.opt_target --  | 
|
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
341  | 
      P.name -- (P.$$$ "(" |-- P.enum1 "|" P.term --| P.$$$ ")" --| P.$$$ "=") --
 | 
| 
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
342  | 
P.position P.text -- Scan.optional (P.$$$ "identifier" |-- Scan.repeat1 P.xname) []  | 
| 22239 | 343  | 
>> (fn ((((loc, a), b), c), d) => Toplevel.local_theory loc (IsarCmd.simproc_setup a b c d)));  | 
| 22202 | 344  | 
|
| 5832 | 345  | 
|
346  | 
(* translation functions *)  | 
|
347  | 
||
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
348  | 
val trfun = P.opt_keyword "advanced" -- P.position P.text;  | 
| 14642 | 349  | 
|
| 24868 | 350  | 
val _ =  | 
| 17068 | 351  | 
OuterSyntax.command "parse_ast_translation" "install parse ast translation functions"  | 
352  | 
(K.tag_ml K.thy_decl)  | 
|
| 22117 | 353  | 
(trfun >> (Toplevel.theory o IsarCmd.parse_ast_translation));  | 
| 5832 | 354  | 
|
| 24868 | 355  | 
val _ =  | 
| 17068 | 356  | 
OuterSyntax.command "parse_translation" "install parse translation functions"  | 
357  | 
(K.tag_ml K.thy_decl)  | 
|
| 22117 | 358  | 
(trfun >> (Toplevel.theory o IsarCmd.parse_translation));  | 
| 5832 | 359  | 
|
| 24868 | 360  | 
val _ =  | 
| 17068 | 361  | 
OuterSyntax.command "print_translation" "install print translation functions"  | 
362  | 
(K.tag_ml K.thy_decl)  | 
|
| 22117 | 363  | 
(trfun >> (Toplevel.theory o IsarCmd.print_translation));  | 
| 5832 | 364  | 
|
| 24868 | 365  | 
val _ =  | 
| 6370 | 366  | 
OuterSyntax.command "typed_print_translation" "install typed print translation functions"  | 
| 17068 | 367  | 
(K.tag_ml K.thy_decl)  | 
| 22117 | 368  | 
(trfun >> (Toplevel.theory o IsarCmd.typed_print_translation));  | 
| 5832 | 369  | 
|
| 24868 | 370  | 
val _ =  | 
| 17068 | 371  | 
OuterSyntax.command "print_ast_translation" "install print ast translation functions"  | 
372  | 
(K.tag_ml K.thy_decl)  | 
|
| 22117 | 373  | 
(trfun >> (Toplevel.theory o IsarCmd.print_ast_translation));  | 
| 5832 | 374  | 
|
| 24868 | 375  | 
val _ =  | 
| 17068 | 376  | 
OuterSyntax.command "token_translation" "install token translation functions"  | 
377  | 
(K.tag_ml K.thy_decl)  | 
|
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
378  | 
(P.position P.text >> (Toplevel.theory o IsarCmd.token_translation));  | 
| 5832 | 379  | 
|
380  | 
||
381  | 
(* oracles *)  | 
|
382  | 
||
| 24868 | 383  | 
val _ =  | 
| 17068 | 384  | 
OuterSyntax.command "oracle" "declare oracle" (K.tag_ml K.thy_decl)  | 
| 16849 | 385  | 
    (P.name -- (P.$$$ "(" |-- P.text --| P.$$$ ")" --| P.$$$ "=")
 | 
| 
26385
 
ae7564661e76
ML runtime compilation: pass position, tuned signature;
 
wenzelm 
parents: 
26248 
diff
changeset
 | 
386  | 
-- P.position P.text >> (fn ((x, y), z) => Toplevel.theory (IsarCmd.oracle x y z)));  | 
| 5832 | 387  | 
|
388  | 
||
| 
21306
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
389  | 
(* local theories *)  | 
| 
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
390  | 
|
| 24868 | 391  | 
val _ =  | 
| 
21306
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
392  | 
OuterSyntax.command "context" "enter local theory context" K.thy_decl  | 
| 
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
393  | 
(P.name --| P.begin >> (fn name =>  | 
| 25290 | 394  | 
Toplevel.print o Toplevel.begin_local_theory true (TheoryTarget.context name)));  | 
| 
21306
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
395  | 
|
| 
 
7ab6e95e6b0b
turned 'context' into plain thy_decl, discontinued thy_switch;
 
wenzelm 
parents: 
21269 
diff
changeset
 | 
396  | 
|
| 12061 | 397  | 
(* locales *)  | 
398  | 
||
| 12758 | 399  | 
val locale_val =  | 
| 24868 | 400  | 
SpecParse.locale_expr --  | 
| 22117 | 401  | 
Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||  | 
| 24868 | 402  | 
Scan.repeat1 SpecParse.context_element >> pair Locale.empty;  | 
| 12061 | 403  | 
|
| 24868 | 404  | 
val _ =  | 
| 12061 | 405  | 
OuterSyntax.command "locale" "define named proof context" K.thy_decl  | 
| 24868 | 406  | 
((P.opt_keyword "open" >> (fn true => NONE | false => SOME "")) --  | 
407  | 
P.name -- Scan.optional (P.$$$ "=" |-- P.!!! locale_val) (Locale.empty, []) -- P.opt_begin  | 
|
| 20958 | 408  | 
>> (fn (((is_open, name), (expr, elems)), begin) =>  | 
| 21843 | 409  | 
(begin ? Toplevel.print) o Toplevel.begin_local_theory begin  | 
| 
22351
 
587845efb4cf
locale: add_locale accepts explicit predicate name, interpretation supports non-mandatory prefixes
 
haftmann 
parents: 
22340 
diff
changeset
 | 
410  | 
(Locale.add_locale is_open name expr elems #-> TheoryTarget.begin)));  | 
| 12061 | 411  | 
|
| 24868 | 412  | 
val _ =  | 
| 
15598
 
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
 
ballarin 
parents: 
15596 
diff
changeset
 | 
413  | 
OuterSyntax.command "interpretation"  | 
| 
16736
 
1e792b32abef
Preparations for interpretation of locales in locales.
 
ballarin 
parents: 
16604 
diff
changeset
 | 
414  | 
"prove and register interpretation of locale expression in theory or locale" K.thy_goal  | 
| 22117 | 415  | 
(P.xname --| (P.$$$ "\\<subseteq>" || P.$$$ "<") -- P.!!! SpecParse.locale_expr  | 
| 20365 | 416  | 
>> (Toplevel.print oo (Toplevel.theory_to_proof o Locale.interpretation_in_locale I)) ||  | 
| 22866 | 417  | 
SpecParse.opt_thm_name ":" -- SpecParse.locale_expr -- SpecParse.locale_insts  | 
418  | 
>> (fn ((x, y), z) => Toplevel.print o  | 
|
419  | 
Toplevel.theory_to_proof (Locale.interpretation I (apfst (pair true) x) y z)));  | 
|
| 12061 | 420  | 
|
| 24868 | 421  | 
val _ =  | 
| 
15624
 
484178635bd8
Further work on interpretation commands.  New command `interpret' for
 
ballarin 
parents: 
15598 
diff
changeset
 | 
422  | 
OuterSyntax.command "interpret"  | 
| 17068 | 423  | 
"prove and register interpretation of locale expression in proof context"  | 
424  | 
(K.tag_proof K.prf_goal)  | 
|
| 22117 | 425  | 
(SpecParse.opt_thm_name ":" -- SpecParse.locale_expr -- SpecParse.locale_insts  | 
| 22866 | 426  | 
>> (fn ((x, y), z) => Toplevel.print o  | 
427  | 
Toplevel.proof' (Locale.interpret Seq.single (apfst (pair true) x) y z)));  | 
|
| 15703 | 428  | 
|
429  | 
||
| 22299 | 430  | 
(* classes *)  | 
431  | 
||
| 24868 | 432  | 
val class_val =  | 
433  | 
SpecParse.class_expr --  | 
|
| 26248 | 434  | 
Scan.optional (P.$$$ "+" |-- P.!!! (Scan.repeat1 SpecParse.context_element)) [] ||  | 
435  | 
Scan.repeat1 SpecParse.context_element >> pair [];  | 
|
| 22299 | 436  | 
|
| 24868 | 437  | 
val _ =  | 
438  | 
OuterSyntax.command "class" "define type class" K.thy_decl  | 
|
439  | 
(P.name --  | 
|
440  | 
     Scan.optional (P.$$$ "(" |-- P.$$$ "attach" |-- Scan.repeat P.term --| P.$$$ ")") [] --  (* FIXME ?? *)
 | 
|
441  | 
(P.$$$ "=" |-- class_val) -- P.opt_begin  | 
|
| 25003 | 442  | 
>> (fn (((name, add_consts), (supclasses, elems)), begin) =>  | 
| 24938 | 443  | 
(begin ? Toplevel.print) o Toplevel.begin_local_theory begin  | 
| 25003 | 444  | 
(Class.class_cmd name supclasses elems add_consts #-> TheoryTarget.begin)));  | 
| 22299 | 445  | 
|
| 24868 | 446  | 
val _ =  | 
| 
24914
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
447  | 
OuterSyntax.command "subclass" "prove a subclass relation" K.thy_goal  | 
| 
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
448  | 
(P.opt_target -- P.xname >> (fn (loc, class) =>  | 
| 
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
449  | 
Toplevel.print o Toplevel.local_theory_to_proof loc (Subclass.subclass_cmd class)));  | 
| 
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
450  | 
|
| 
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
451  | 
val _ =  | 
| 25519 | 452  | 
OuterSyntax.command "instantiation" "instantiate and prove type arity" K.thy_decl  | 
| 25536 | 453  | 
(P.multi_arity --| P.begin  | 
| 25462 | 454  | 
>> (fn arities => Toplevel.print o  | 
455  | 
Toplevel.begin_local_theory true (Instance.instantiation_cmd arities)));  | 
|
| 24589 | 456  | 
|
| 25485 | 457  | 
val _ =  | 
458  | 
OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal  | 
|
459  | 
((P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname) >> Class.classrel_cmd ||  | 
|
| 25536 | 460  | 
P.arity -- Scan.repeat (SpecParse.opt_thm_name ":" -- P.prop)  | 
461  | 
>> (fn (arity, defs) => Instance.instance_cmd arity defs))  | 
|
| 25485 | 462  | 
>> (Toplevel.print oo Toplevel.theory_to_proof)  | 
463  | 
|| Scan.succeed (Toplevel.print o Toplevel.local_theory_to_proof NONE (Class.instantiation_instance I)));  | 
|
| 22299 | 464  | 
|
465  | 
||
| 25519 | 466  | 
(* arbitrary overloading *)  | 
467  | 
||
468  | 
val _ =  | 
|
469  | 
OuterSyntax.command "overloading" "overloaded definitions" K.thy_decl  | 
|
| 25861 | 470  | 
(Scan.repeat1 (P.name --| (P.$$$ "\\<equiv>" || P.$$$ "==") -- P.term --  | 
| 26048 | 471  | 
      Scan.optional (P.$$$ "(" |-- (P.$$$ "unchecked" >> K false) --| P.$$$ ")") true
 | 
472  | 
>> P.triple1) --| P.begin  | 
|
| 25519 | 473  | 
>> (fn operations => Toplevel.print o  | 
474  | 
Toplevel.begin_local_theory true (TheoryTarget.overloading_cmd operations)));  | 
|
475  | 
||
476  | 
||
| 22866 | 477  | 
(* code generation *)  | 
478  | 
||
| 24868 | 479  | 
val _ =  | 
| 22866 | 480  | 
OuterSyntax.command "code_datatype" "define set of code datatype constructors" K.thy_decl  | 
| 
24423
 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 
haftmann 
parents: 
24314 
diff
changeset
 | 
481  | 
(Scan.repeat1 P.term >> (Toplevel.theory o Code.add_datatype_cmd));  | 
| 22866 | 482  | 
|
483  | 
||
| 5832 | 484  | 
|
485  | 
(** proof commands **)  | 
|
486  | 
||
487  | 
(* statements *)  | 
|
488  | 
||
| 17353 | 489  | 
fun gen_theorem kind =  | 
490  | 
  OuterSyntax.command kind ("state " ^ kind) K.thy_goal
 | 
|
| 22117 | 491  | 
(P.opt_target -- Scan.optional (SpecParse.opt_thm_name ":" --|  | 
492  | 
      Scan.ahead (SpecParse.locale_keyword || SpecParse.statement_keyword)) ("", []) --
 | 
|
493  | 
SpecParse.general_statement >> (fn ((loc, a), (elems, concl)) =>  | 
|
| 21033 | 494  | 
(Toplevel.print o  | 
| 
24451
 
7c205d006719
Specification.theorem now also takes "interactive" flag as argument.
 
berghofe 
parents: 
24423 
diff
changeset
 | 
495  | 
Toplevel.local_theory_to_proof' loc  | 
| 
24932
 
86ef9a828a9e
AxClass.axiomatize and Specification: renamed XXX_i to XXX, and XXX to XXX_cmd;
 
wenzelm 
parents: 
24914 
diff
changeset
 | 
496  | 
(Specification.theorem_cmd kind NONE (K I) a elems concl))));  | 
| 5832 | 497  | 
|
| 24868 | 498  | 
val _ = gen_theorem Thm.theoremK;  | 
499  | 
val _ = gen_theorem Thm.lemmaK;  | 
|
500  | 
val _ = gen_theorem Thm.corollaryK;  | 
|
| 5832 | 501  | 
|
| 24868 | 502  | 
val _ =  | 
| 17353 | 503  | 
OuterSyntax.command "have" "state local goal"  | 
504  | 
(K.tag_proof K.prf_goal)  | 
|
| 22117 | 505  | 
(SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.have));  | 
| 17353 | 506  | 
|
| 24868 | 507  | 
val _ =  | 
| 17353 | 508  | 
OuterSyntax.command "hence" "abbreviates \"then have\""  | 
509  | 
(K.tag_proof K.prf_goal)  | 
|
| 22117 | 510  | 
(SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.hence));  | 
| 17353 | 511  | 
|
| 24868 | 512  | 
val _ =  | 
| 17068 | 513  | 
OuterSyntax.command "show" "state local goal, solving current obligation"  | 
| 
21800
 
6035bfcd72d8
classified show/thus as prf_asm_goal, which is usually hilited in PG;
 
wenzelm 
parents: 
21726 
diff
changeset
 | 
514  | 
(K.tag_proof K.prf_asm_goal)  | 
| 22117 | 515  | 
(SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.show));  | 
| 5832 | 516  | 
|
| 24868 | 517  | 
val _ =  | 
| 17068 | 518  | 
OuterSyntax.command "thus" "abbreviates \"then show\""  | 
| 
21800
 
6035bfcd72d8
classified show/thus as prf_asm_goal, which is usually hilited in PG;
 
wenzelm 
parents: 
21726 
diff
changeset
 | 
519  | 
(K.tag_proof K.prf_asm_goal)  | 
| 22117 | 520  | 
(SpecParse.statement >> ((Toplevel.print oo Toplevel.proof') o IsarCmd.thus));  | 
| 6501 | 521  | 
|
| 5832 | 522  | 
|
| 5914 | 523  | 
(* facts *)  | 
| 5832 | 524  | 
|
| 22117 | 525  | 
val facts = P.and_list1 SpecParse.xthms1;  | 
| 9197 | 526  | 
|
| 24868 | 527  | 
val _ =  | 
| 17068 | 528  | 
OuterSyntax.command "then" "forward chaining"  | 
529  | 
(K.tag_proof K.prf_chain)  | 
|
| 17900 | 530  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.chain));  | 
| 5832 | 531  | 
|
| 24868 | 532  | 
val _ =  | 
| 17068 | 533  | 
OuterSyntax.command "from" "forward chaining from given facts"  | 
534  | 
(K.tag_proof K.prf_chain)  | 
|
| 17900 | 535  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss)));  | 
| 5914 | 536  | 
|
| 24868 | 537  | 
val _ =  | 
| 17068 | 538  | 
OuterSyntax.command "with" "forward chaining from given and current facts"  | 
539  | 
(K.tag_proof K.prf_chain)  | 
|
| 17900 | 540  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss)));  | 
| 6878 | 541  | 
|
| 24868 | 542  | 
val _ =  | 
| 17068 | 543  | 
OuterSyntax.command "note" "define facts"  | 
544  | 
(K.tag_proof K.prf_decl)  | 
|
| 22117 | 545  | 
(SpecParse.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss)));  | 
| 5832 | 546  | 
|
| 24868 | 547  | 
val _ =  | 
| 17068 | 548  | 
OuterSyntax.command "using" "augment goal facts"  | 
549  | 
(K.tag_proof K.prf_decl)  | 
|
| 18544 | 550  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.using)));  | 
551  | 
||
| 24868 | 552  | 
val _ =  | 
| 18544 | 553  | 
OuterSyntax.command "unfolding" "unfold definitions in goal and facts"  | 
554  | 
(K.tag_proof K.prf_decl)  | 
|
555  | 
(facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding)));  | 
|
| 12926 | 556  | 
|
| 5832 | 557  | 
|
558  | 
(* proof context *)  | 
|
559  | 
||
| 24868 | 560  | 
val _ =  | 
| 17068 | 561  | 
OuterSyntax.command "fix" "fix local variables (Skolem constants)"  | 
562  | 
(K.tag_proof K.prf_asm)  | 
|
| 19844 | 563  | 
(P.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix)));  | 
| 11890 | 564  | 
|
| 24868 | 565  | 
val _ =  | 
| 17068 | 566  | 
OuterSyntax.command "assume" "assume propositions"  | 
567  | 
(K.tag_proof K.prf_asm)  | 
|
| 22117 | 568  | 
(SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume)));  | 
| 5832 | 569  | 
|
| 24868 | 570  | 
val _ =  | 
| 17068 | 571  | 
OuterSyntax.command "presume" "assume propositions, to be established later"  | 
572  | 
(K.tag_proof K.prf_asm)  | 
|
| 22117 | 573  | 
(SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume)));  | 
| 6850 | 574  | 
|
| 24868 | 575  | 
val _ =  | 
| 17068 | 576  | 
OuterSyntax.command "def" "local definition"  | 
577  | 
(K.tag_proof K.prf_asm)  | 
|
| 18308 | 578  | 
(P.and_list1  | 
| 22117 | 579  | 
(SpecParse.opt_thm_name ":" --  | 
| 19844 | 580  | 
((P.name -- P.opt_mixfix) -- ((P.$$$ "\\<equiv>" || P.$$$ "==") |-- P.!!! P.termp)))  | 
| 18308 | 581  | 
>> (Toplevel.print oo (Toplevel.proof o Proof.def)));  | 
| 6953 | 582  | 
|
| 24868 | 583  | 
val _ =  | 
| 11890 | 584  | 
OuterSyntax.command "obtain" "generalized existence"  | 
| 17068 | 585  | 
(K.tag_proof K.prf_asm_goal)  | 
| 22117 | 586  | 
(P.parname -- Scan.optional (P.fixes --| P.where_) [] -- SpecParse.statement  | 
| 18895 | 587  | 
>> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain x y z)));  | 
| 5832 | 588  | 
|
| 24868 | 589  | 
val _ =  | 
| 17854 | 590  | 
OuterSyntax.command "guess" "wild guessing (unstructured)"  | 
591  | 
(K.tag_proof K.prf_asm_goal)  | 
|
| 19844 | 592  | 
(Scan.optional P.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess)));  | 
| 17854 | 593  | 
|
| 24868 | 594  | 
val _ =  | 
| 17068 | 595  | 
OuterSyntax.command "let" "bind text variables"  | 
596  | 
(K.tag_proof K.prf_decl)  | 
|
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
597  | 
(P.and_list1 (P.enum1 "and" P.term -- (P.$$$ "=" |-- P.term))  | 
| 17900 | 598  | 
>> (Toplevel.print oo (Toplevel.proof o Proof.let_bind)));  | 
| 5832 | 599  | 
|
| 
11793
 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 
wenzelm 
parents: 
11742 
diff
changeset
 | 
600  | 
val case_spec =  | 
| 15703 | 601  | 
  (P.$$$ "(" |-- P.!!! (P.xname -- Scan.repeat1 (P.maybe P.name) --| P.$$$ ")") ||
 | 
| 22117 | 602  | 
P.xname >> rpair []) -- SpecParse.opt_attribs >> P.triple1;  | 
| 
11793
 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 
wenzelm 
parents: 
11742 
diff
changeset
 | 
603  | 
|
| 24868 | 604  | 
val _ =  | 
| 17068 | 605  | 
OuterSyntax.command "case" "invoke local context"  | 
606  | 
(K.tag_proof K.prf_asm)  | 
|
| 17900 | 607  | 
(case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case)));  | 
| 8370 | 608  | 
|
| 5832 | 609  | 
|
610  | 
(* proof structure *)  | 
|
611  | 
||
| 24868 | 612  | 
val _ =  | 
| 17068 | 613  | 
  OuterSyntax.command "{" "begin explicit proof block"
 | 
614  | 
(K.tag_proof K.prf_open)  | 
|
| 17900 | 615  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.begin_block));  | 
| 5832 | 616  | 
|
| 24868 | 617  | 
val _ =  | 
| 17068 | 618  | 
OuterSyntax.command "}" "end explicit proof block"  | 
619  | 
(K.tag_proof K.prf_close)  | 
|
| 20305 | 620  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.end_block));  | 
| 6687 | 621  | 
|
| 24868 | 622  | 
val _ =  | 
| 17068 | 623  | 
OuterSyntax.command "next" "enter next proof block"  | 
624  | 
(K.tag_proof K.prf_block)  | 
|
| 17900 | 625  | 
(Scan.succeed (Toplevel.print o Toplevel.proof Proof.next_block));  | 
| 5832 | 626  | 
|
627  | 
||
628  | 
(* end proof *)  | 
|
629  | 
||
| 24868 | 630  | 
val _ =  | 
| 17068 | 631  | 
OuterSyntax.command "qed" "conclude (sub-)proof"  | 
632  | 
(K.tag_proof K.qed_block)  | 
|
| 22117 | 633  | 
(Scan.option Method.parse >> (Toplevel.print3 oo IsarCmd.qed));  | 
| 5832 | 634  | 
|
| 24868 | 635  | 
val _ =  | 
| 17068 | 636  | 
OuterSyntax.command "by" "terminal backward proof"  | 
637  | 
(K.tag_proof K.qed)  | 
|
| 22117 | 638  | 
(Method.parse -- Scan.option Method.parse >> (Toplevel.print3 oo IsarCmd.terminal_proof));  | 
| 6404 | 639  | 
|
| 24868 | 640  | 
val _ =  | 
| 17068 | 641  | 
OuterSyntax.command ".." "default proof"  | 
642  | 
(K.tag_proof K.qed)  | 
|
| 21350 | 643  | 
(Scan.succeed (Toplevel.print3 o IsarCmd.default_proof));  | 
| 8966 | 644  | 
|
| 24868 | 645  | 
val _ =  | 
| 17068 | 646  | 
OuterSyntax.command "." "immediate proof"  | 
647  | 
(K.tag_proof K.qed)  | 
|
| 21350 | 648  | 
(Scan.succeed (Toplevel.print3 o IsarCmd.immediate_proof));  | 
| 6404 | 649  | 
|
| 24868 | 650  | 
val _ =  | 
| 17068 | 651  | 
OuterSyntax.command "done" "done proof"  | 
652  | 
(K.tag_proof K.qed)  | 
|
| 21350 | 653  | 
(Scan.succeed (Toplevel.print3 o IsarCmd.done_proof));  | 
| 5832 | 654  | 
|
| 24868 | 655  | 
val _ =  | 
| 25108 | 656  | 
OuterSyntax.command "sorry" "skip proof (quick-and-dirty mode only!)"  | 
| 17068 | 657  | 
(K.tag_proof K.qed)  | 
| 21350 | 658  | 
(Scan.succeed (Toplevel.print3 o IsarCmd.skip_proof));  | 
| 
6888
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents: 
6886 
diff
changeset
 | 
659  | 
|
| 24868 | 660  | 
val _ =  | 
| 17068 | 661  | 
OuterSyntax.command "oops" "forget proof"  | 
662  | 
(K.tag_proof K.qed_global)  | 
|
| 18561 | 663  | 
(Scan.succeed Toplevel.forget_proof);  | 
| 8210 | 664  | 
|
| 5832 | 665  | 
|
666  | 
(* proof steps *)  | 
|
667  | 
||
| 24868 | 668  | 
val _ =  | 
| 17068 | 669  | 
OuterSyntax.command "defer" "shuffle internal proof state"  | 
670  | 
(K.tag_proof K.prf_script)  | 
|
| 17900 | 671  | 
(Scan.option P.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.defer)));  | 
| 8165 | 672  | 
|
| 24868 | 673  | 
val _ =  | 
| 17068 | 674  | 
OuterSyntax.command "prefer" "shuffle internal proof state"  | 
675  | 
(K.tag_proof K.prf_script)  | 
|
| 17900 | 676  | 
(P.nat >> (Toplevel.print oo (Toplevel.proofs o Proof.prefer)));  | 
| 8165 | 677  | 
|
| 24868 | 678  | 
val _ =  | 
| 17068 | 679  | 
OuterSyntax.command "apply" "initial refinement step (unstructured)"  | 
680  | 
(K.tag_proof K.prf_script)  | 
|
| 22117 | 681  | 
(Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply)));  | 
| 5832 | 682  | 
|
| 24868 | 683  | 
val _ =  | 
| 17068 | 684  | 
OuterSyntax.command "apply_end" "terminal refinement (unstructured)"  | 
685  | 
(K.tag_proof K.prf_script)  | 
|
| 22117 | 686  | 
(Method.parse >> (Toplevel.print oo (Toplevel.proofs o Proof.apply_end)));  | 
| 5832 | 687  | 
|
| 24868 | 688  | 
val _ =  | 
| 17068 | 689  | 
OuterSyntax.command "proof" "backward proof"  | 
690  | 
(K.tag_proof K.prf_block)  | 
|
| 22117 | 691  | 
(Scan.option Method.parse >> (fn m => Toplevel.print o  | 
| 
17107
 
2c35e00ee2ab
various Toplevel.theory_context commands: proper presentation in context;
 
wenzelm 
parents: 
17068 
diff
changeset
 | 
692  | 
Toplevel.actual_proof (ProofHistory.applys (Proof.proof m)) o  | 
| 
15237
 
250e9be7a09d
Some changes to allow skipping of proof scripts.
 
berghofe 
parents: 
15141 
diff
changeset
 | 
693  | 
Toplevel.skip_proof (History.apply (fn i => i + 1))));  | 
| 5832 | 694  | 
|
695  | 
||
| 6773 | 696  | 
(* calculational proof commands *)  | 
697  | 
||
| 6878 | 698  | 
val calc_args =  | 
| 22117 | 699  | 
  Scan.option (P.$$$ "(" |-- P.!!! ((SpecParse.xthms1 --| P.$$$ ")")));
 | 
| 6878 | 700  | 
|
| 24868 | 701  | 
val _ =  | 
| 17068 | 702  | 
OuterSyntax.command "also" "combine calculation and current facts"  | 
703  | 
(K.tag_proof K.prf_decl)  | 
|
| 17900 | 704  | 
(calc_args >> (Toplevel.proofs' o Calculation.also));  | 
| 6773 | 705  | 
|
| 24868 | 706  | 
val _ =  | 
| 17068 | 707  | 
OuterSyntax.command "finally" "combine calculation and current facts, exhibit result"  | 
708  | 
(K.tag_proof K.prf_chain)  | 
|
| 22573 | 709  | 
(calc_args >> (Toplevel.proofs' o Calculation.finally_));  | 
| 6773 | 710  | 
|
| 24868 | 711  | 
val _ =  | 
| 17068 | 712  | 
OuterSyntax.command "moreover" "augment calculation by current facts"  | 
713  | 
(K.tag_proof K.prf_decl)  | 
|
| 17900 | 714  | 
(Scan.succeed (Toplevel.proof' Calculation.moreover));  | 
| 8562 | 715  | 
|
| 24868 | 716  | 
val _ =  | 
| 8588 | 717  | 
OuterSyntax.command "ultimately" "augment calculation by current facts, exhibit result"  | 
| 17068 | 718  | 
(K.tag_proof K.prf_chain)  | 
| 17900 | 719  | 
(Scan.succeed (Toplevel.proof' Calculation.ultimately));  | 
| 8588 | 720  | 
|
| 6773 | 721  | 
|
| 6742 | 722  | 
(* proof navigation *)  | 
| 5944 | 723  | 
|
| 24868 | 724  | 
val _ =  | 
| 17068 | 725  | 
OuterSyntax.command "back" "backtracking of proof command"  | 
726  | 
(K.tag_proof K.prf_script)  | 
|
727  | 
(Scan.succeed (Toplevel.print o IsarCmd.back));  | 
|
| 5832 | 728  | 
|
729  | 
||
| 6742 | 730  | 
(* history *)  | 
731  | 
||
| 24868 | 732  | 
val _ =  | 
| 20979 | 733  | 
OuterSyntax.improper_command "cannot_undo" "report 'cannot undo' error message" K.control  | 
734  | 
(P.name >> (Toplevel.no_timing oo IsarCmd.cannot_undo));  | 
|
| 6742 | 735  | 
|
| 24868 | 736  | 
val _ =  | 
| 6742 | 737  | 
OuterSyntax.improper_command "redo" "redo last command" K.control  | 
| 9010 | 738  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.print o IsarCmd.redo));  | 
| 6742 | 739  | 
|
| 24868 | 740  | 
val _ =  | 
| 6742 | 741  | 
OuterSyntax.improper_command "undos_proof" "undo last proof commands" K.control  | 
| 9010 | 742  | 
(P.nat >> ((Toplevel.no_timing o Toplevel.print) oo IsarCmd.undos_proof));  | 
| 6742 | 743  | 
|
| 24868 | 744  | 
val _ =  | 
| 6742 | 745  | 
OuterSyntax.improper_command "undo" "undo last command" K.control  | 
| 
25192
 
b568f8c5d5ca
made command 'undo' silent ('ProofGeneral.undo' becomes a historical relic);
 
wenzelm 
parents: 
25108 
diff
changeset
 | 
746  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.undo));  | 
| 6742 | 747  | 
|
| 24868 | 748  | 
val _ =  | 
| 8500 | 749  | 
OuterSyntax.improper_command "kill" "kill current history node" K.control  | 
| 22866 | 750  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.print o IsarCmd.kill));  | 
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
751  | 
|
| 
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
752  | 
|
| 25578 | 753  | 
(* nested command *)  | 
754  | 
||
755  | 
val _ =  | 
|
756  | 
OuterSyntax.improper_command "Isabelle.command" "nested Isabelle command" K.control  | 
|
| 
25794
 
11bec58fc289
Isabelle.command: IsarCmd.nested_command (with properties);
 
wenzelm 
parents: 
25625 
diff
changeset
 | 
757  | 
((Scan.optional P.properties [] -- P.position P.string) :|-- (fn (props, arg) =>  | 
| 
 
11bec58fc289
Isabelle.command: IsarCmd.nested_command (with properties);
 
wenzelm 
parents: 
25625 
diff
changeset
 | 
758  | 
Scan.succeed (K (IsarCmd.nested_command props arg))  | 
| 
 
11bec58fc289
Isabelle.command: IsarCmd.nested_command (with properties);
 
wenzelm 
parents: 
25625 
diff
changeset
 | 
759  | 
handle ERROR msg => Scan.fail_with (K msg)));  | 
| 25578 | 760  | 
|
761  | 
||
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
762  | 
|
| 5832 | 763  | 
(** diagnostic commands (for interactive mode only) **)  | 
764  | 
||
| 8464 | 765  | 
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
 | 
| 20621 | 766  | 
val opt_bang = Scan.optional (P.$$$ "!" >> K true) false;  | 
| 8464 | 767  | 
|
| 24868 | 768  | 
val _ =  | 
| 7124 | 769  | 
OuterSyntax.improper_command "pretty_setmargin" "change default margin for pretty printing"  | 
| 9010 | 770  | 
K.diag (P.nat >> (Toplevel.no_timing oo IsarCmd.pretty_setmargin));  | 
| 7124 | 771  | 
|
| 24868 | 772  | 
val _ =  | 
| 21714 | 773  | 
OuterSyntax.improper_command "help" "print outer syntax commands" K.diag  | 
| 24871 | 774  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative OuterSyntax.print_outer_syntax));  | 
| 21714 | 775  | 
|
| 24868 | 776  | 
val _ =  | 
| 21714 | 777  | 
OuterSyntax.improper_command "print_commands" "print outer syntax commands" K.diag  | 
| 24871 | 778  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative OuterSyntax.print_outer_syntax));  | 
| 5832 | 779  | 
|
| 24868 | 780  | 
val _ =  | 
| 24115 | 781  | 
OuterSyntax.improper_command "print_configs" "print configuration options" K.diag  | 
782  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_configs));  | 
|
| 23989 | 783  | 
|
| 24868 | 784  | 
val _ =  | 
| 7308 | 785  | 
OuterSyntax.improper_command "print_context" "print theory context name" K.diag  | 
| 9010 | 786  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_context));  | 
| 7308 | 787  | 
|
| 24868 | 788  | 
val _ =  | 
| 6723 | 789  | 
OuterSyntax.improper_command "print_theory" "print logical theory contents (verbose!)" K.diag  | 
| 20621 | 790  | 
(opt_bang >> (Toplevel.no_timing oo IsarCmd.print_theory));  | 
| 5832 | 791  | 
|
| 24868 | 792  | 
val _ =  | 
| 21726 | 793  | 
OuterSyntax.improper_command "print_syntax" "print inner syntax of context (verbose!)" K.diag  | 
| 9010 | 794  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_syntax));  | 
| 5832 | 795  | 
|
| 24868 | 796  | 
val _ =  | 
| 21726 | 797  | 
OuterSyntax.improper_command "print_abbrevs" "print constant abbreviation of context" K.diag  | 
798  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_abbrevs));  | 
|
799  | 
||
| 24868 | 800  | 
val _ =  | 
| 17068 | 801  | 
OuterSyntax.improper_command "print_theorems"  | 
802  | 
"print theorems of local theory or proof context" K.diag  | 
|
| 9010 | 803  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_theorems));  | 
| 5881 | 804  | 
|
| 24868 | 805  | 
val _ =  | 
| 12061 | 806  | 
OuterSyntax.improper_command "print_locales" "print locales of this theory" K.diag  | 
807  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_locales));  | 
|
808  | 
||
| 24868 | 809  | 
val _ =  | 
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
810  | 
OuterSyntax.improper_command "print_classes" "print classes of this theory" K.diag  | 
| 
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
811  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory  | 
| 24219 | 812  | 
o Toplevel.keep (Class.print_classes o Toplevel.theory_of)));  | 
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
813  | 
|
| 24868 | 814  | 
val _ =  | 
| 15596 | 815  | 
OuterSyntax.improper_command "print_locale" "print locale expression in this theory" K.diag  | 
| 20621 | 816  | 
(opt_bang -- locale_val >> (Toplevel.no_timing oo IsarCmd.print_locale));  | 
| 12061 | 817  | 
|
| 24868 | 818  | 
val _ =  | 
| 
15598
 
4ab52355bb53
Registrations of global locale interpretations: improved, better naming.
 
ballarin 
parents: 
15596 
diff
changeset
 | 
819  | 
OuterSyntax.improper_command "print_interps"  | 
| 
16102
 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 
ballarin 
parents: 
16074 
diff
changeset
 | 
820  | 
"print interpretations of named locale" K.diag  | 
| 18670 | 821  | 
(Scan.optional (P.$$$ "!" >> K true) false -- P.xname  | 
822  | 
>> (Toplevel.no_timing oo uncurry IsarCmd.print_registrations));  | 
|
| 15596 | 823  | 
|
| 24868 | 824  | 
val _ =  | 
| 12061 | 825  | 
OuterSyntax.improper_command "print_attributes" "print attributes of this theory" K.diag  | 
| 9010 | 826  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_attributes));  | 
| 5832 | 827  | 
|
| 24868 | 828  | 
val _ =  | 
| 16027 | 829  | 
OuterSyntax.improper_command "print_simpset" "print context of Simplifier" K.diag  | 
830  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_simpset));  | 
|
831  | 
||
| 24868 | 832  | 
val _ =  | 
| 12383 | 833  | 
OuterSyntax.improper_command "print_rules" "print intro/elim rules" K.diag  | 
834  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_rules));  | 
|
835  | 
||
| 24868 | 836  | 
val _ =  | 
| 11666 | 837  | 
OuterSyntax.improper_command "print_trans_rules" "print transitivity rules" K.diag  | 
| 9221 | 838  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_trans_rules));  | 
839  | 
||
| 24868 | 840  | 
val _ =  | 
| 12061 | 841  | 
OuterSyntax.improper_command "print_methods" "print methods of this theory" K.diag  | 
| 9010 | 842  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_methods));  | 
| 5832 | 843  | 
|
| 24868 | 844  | 
val _ =  | 
| 9221 | 845  | 
OuterSyntax.improper_command "print_antiquotations" "print antiquotations (global)" K.diag  | 
846  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_antiquotations));  | 
|
847  | 
||
| 24868 | 848  | 
val _ =  | 
| 22485 | 849  | 
OuterSyntax.improper_command "thy_deps" "visualize theory dependencies"  | 
850  | 
K.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.thy_deps));  | 
|
851  | 
||
| 24868 | 852  | 
val _ =  | 
| 20574 | 853  | 
OuterSyntax.improper_command "class_deps" "visualize class dependencies"  | 
854  | 
K.diag (Scan.succeed (Toplevel.no_timing o IsarCmd.class_deps));  | 
|
855  | 
||
| 24868 | 856  | 
val _ =  | 
| 9454 | 857  | 
OuterSyntax.improper_command "thm_deps" "visualize theorem dependencies"  | 
| 22117 | 858  | 
K.diag (SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.thm_deps));  | 
| 9454 | 859  | 
|
| 22866 | 860  | 
local  | 
861  | 
||
| 16027 | 862  | 
val criterion =  | 
863  | 
P.reserved "name" |-- P.!!! (P.$$$ ":" |-- P.xname) >> FindTheorems.Name ||  | 
|
864  | 
P.reserved "intro" >> K FindTheorems.Intro ||  | 
|
865  | 
P.reserved "elim" >> K FindTheorems.Elim ||  | 
|
866  | 
P.reserved "dest" >> K FindTheorems.Dest ||  | 
|
| 16074 | 867  | 
P.reserved "simp" |-- P.!!! (P.$$$ ":" |-- P.term) >> FindTheorems.Simp ||  | 
| 16027 | 868  | 
P.term >> FindTheorems.Pattern;  | 
869  | 
||
| 22866 | 870  | 
val options =  | 
871  | 
Scan.optional  | 
|
872  | 
    (P.$$$ "(" |--
 | 
|
873  | 
P.!!! (Scan.option P.nat -- Scan.optional (P.reserved "with_dups" >> K false) true  | 
|
874  | 
--| P.$$$ ")")) (NONE, true);  | 
|
875  | 
in  | 
|
876  | 
||
| 24868 | 877  | 
val _ =  | 
| 22866 | 878  | 
OuterSyntax.improper_command "find_theorems" "print theorems meeting specified criteria" K.diag  | 
879  | 
(options -- Scan.repeat (((Scan.option P.minus >> is_none) -- criterion))  | 
|
| 
17219
 
515badbfc4d6
renamed 'thms_containing' to 'find_theorems' -- keep old version for the time being;
 
wenzelm 
parents: 
17142 
diff
changeset
 | 
880  | 
>> (Toplevel.no_timing oo IsarCmd.find_theorems));  | 
| 
 
515badbfc4d6
renamed 'thms_containing' to 'find_theorems' -- keep old version for the time being;
 
wenzelm 
parents: 
17142 
diff
changeset
 | 
881  | 
|
| 22866 | 882  | 
end;  | 
883  | 
||
| 24868 | 884  | 
val _ =  | 
| 6723 | 885  | 
OuterSyntax.improper_command "print_binds" "print term bindings of proof context" K.diag  | 
| 9010 | 886  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_binds));  | 
| 5832 | 887  | 
|
| 24868 | 888  | 
val _ =  | 
| 8370 | 889  | 
OuterSyntax.improper_command "print_facts" "print facts of proof context" K.diag  | 
| 21004 | 890  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_facts));  | 
| 5832 | 891  | 
|
| 24868 | 892  | 
val _ =  | 
| 8370 | 893  | 
OuterSyntax.improper_command "print_cases" "print cases of proof context" K.diag  | 
| 9010 | 894  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.print_cases));  | 
| 8370 | 895  | 
|
| 24868 | 896  | 
val _ =  | 
| 19269 | 897  | 
OuterSyntax.improper_command "print_statement" "print theorems as long statements" K.diag  | 
| 22117 | 898  | 
(opt_modes -- SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_stmts));  | 
| 19269 | 899  | 
|
| 24868 | 900  | 
val _ =  | 
| 8464 | 901  | 
OuterSyntax.improper_command "thm" "print theorems" K.diag  | 
| 22117 | 902  | 
(opt_modes -- SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_thms));  | 
| 5832 | 903  | 
|
| 24868 | 904  | 
val _ =  | 
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
905  | 
OuterSyntax.improper_command "prf" "print proof terms of theorems" K.diag  | 
| 22866 | 906  | 
(opt_modes -- Scan.option SpecParse.xthms1  | 
907  | 
>> (Toplevel.no_timing oo IsarCmd.print_prfs false));  | 
|
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
908  | 
|
| 24868 | 909  | 
val _ =  | 
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
910  | 
OuterSyntax.improper_command "full_prf" "print full proof terms of theorems" K.diag  | 
| 22117 | 911  | 
(opt_modes -- Scan.option SpecParse.xthms1 >> (Toplevel.no_timing oo IsarCmd.print_prfs true));  | 
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
912  | 
|
| 24868 | 913  | 
val _ =  | 
| 6723 | 914  | 
OuterSyntax.improper_command "prop" "read and print proposition" K.diag  | 
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
915  | 
(opt_modes -- P.term >> (Toplevel.no_timing oo IsarCmd.print_prop));  | 
| 5832 | 916  | 
|
| 24868 | 917  | 
val _ =  | 
| 6723 | 918  | 
OuterSyntax.improper_command "term" "read and print term" K.diag  | 
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
919  | 
(opt_modes -- P.term >> (Toplevel.no_timing oo IsarCmd.print_term));  | 
| 5832 | 920  | 
|
| 24868 | 921  | 
val _ =  | 
| 6723 | 922  | 
OuterSyntax.improper_command "typ" "read and print type" K.diag  | 
| 
12876
 
a70df1e5bf10
got rid of explicit marginal comments (now stripped earlier from input);
 
wenzelm 
parents: 
12768 
diff
changeset
 | 
923  | 
(opt_modes -- P.typ >> (Toplevel.no_timing oo IsarCmd.print_type));  | 
| 5832 | 924  | 
|
| 24868 | 925  | 
val _ =  | 
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
926  | 
OuterSyntax.improper_command "print_codesetup" "print code generator setup of this theory" K.diag  | 
| 
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
927  | 
(Scan.succeed  | 
| 
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
928  | 
(Toplevel.no_timing o Toplevel.unknown_theory o Toplevel.keep  | 
| 24219 | 929  | 
(Code.print_codesetup o Toplevel.theory_of)));  | 
| 5832 | 930  | 
|
| 26184 | 931  | 
val _ =  | 
932  | 
OuterSyntax.improper_command "unused_thms" "find unused theorems" K.diag  | 
|
933  | 
(Scan.option ((Scan.repeat1 (Scan.unless P.minus P.name) --| P.minus) --  | 
|
934  | 
Scan.option (Scan.repeat1 (Scan.unless P.minus P.name))) >>  | 
|
935  | 
(Toplevel.no_timing oo IsarCmd.unused_thms));  | 
|
936  | 
||
| 5832 | 937  | 
|
938  | 
(** system commands (for interactive mode only) **)  | 
|
939  | 
||
| 24868 | 940  | 
val _ =  | 
| 8650 | 941  | 
OuterSyntax.improper_command "cd" "change current working directory" K.diag  | 
| 14950 | 942  | 
(P.path >> (Toplevel.no_timing oo IsarCmd.cd));  | 
| 5832 | 943  | 
|
| 24868 | 944  | 
val _ =  | 
| 6723 | 945  | 
OuterSyntax.improper_command "pwd" "print current working directory" K.diag  | 
| 9010 | 946  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.pwd));  | 
| 5832 | 947  | 
|
| 24868 | 948  | 
val _ =  | 
| 6723 | 949  | 
OuterSyntax.improper_command "use_thy" "use theory file" K.diag  | 
| 26404 | 950  | 
(P.name >> (fn name =>  | 
951  | 
Toplevel.no_timing o Toplevel.imperative (fn () => ThyInfo.use_thy name)));  | 
|
| 5832 | 952  | 
|
| 24868 | 953  | 
val _ =  | 
| 7102 | 954  | 
OuterSyntax.improper_command "touch_thy" "outdate theory, including descendants" K.diag  | 
| 9010 | 955  | 
(P.name >> (Toplevel.no_timing oo IsarCmd.touch_thy));  | 
| 7102 | 956  | 
|
| 24868 | 957  | 
val _ =  | 
| 7908 | 958  | 
OuterSyntax.improper_command "touch_child_thys" "outdate child theories" K.diag  | 
| 9010 | 959  | 
(P.name >> (Toplevel.no_timing oo IsarCmd.touch_child_thys));  | 
| 7908 | 960  | 
|
| 24868 | 961  | 
val _ =  | 
| 7102 | 962  | 
OuterSyntax.improper_command "remove_thy" "remove theory from loader database" K.diag  | 
| 9010 | 963  | 
(P.name >> (Toplevel.no_timing oo IsarCmd.remove_thy));  | 
| 7102 | 964  | 
|
| 24868 | 965  | 
val _ =  | 
| 7931 | 966  | 
OuterSyntax.improper_command "kill_thy" "kill theory -- try to remove from loader database"  | 
| 9010 | 967  | 
K.diag (P.name >> (Toplevel.no_timing oo IsarCmd.kill_thy));  | 
| 7931 | 968  | 
|
| 24868 | 969  | 
val _ =  | 
| 14934 | 970  | 
OuterSyntax.improper_command "display_drafts" "display raw source files with symbols"  | 
| 14950 | 971  | 
K.diag (Scan.repeat1 P.path >> (Toplevel.no_timing oo IsarCmd.display_drafts));  | 
| 14934 | 972  | 
|
| 24868 | 973  | 
val _ =  | 
| 14934 | 974  | 
OuterSyntax.improper_command "print_drafts" "print raw source files with symbols"  | 
| 14950 | 975  | 
K.diag (Scan.repeat1 P.path >> (Toplevel.no_timing oo IsarCmd.print_drafts));  | 
| 14934 | 976  | 
|
| 9731 | 977  | 
val opt_limits =  | 
978  | 
Scan.option P.nat -- Scan.option (P.$$$ "," |-- P.!!! P.nat);  | 
|
979  | 
||
| 24868 | 980  | 
val _ =  | 
| 8886 | 981  | 
OuterSyntax.improper_command "pr" "print current proof state (if present)" K.diag  | 
| 9731 | 982  | 
(opt_modes -- opt_limits >> (Toplevel.no_timing oo IsarCmd.pr));  | 
| 7199 | 983  | 
|
| 24868 | 984  | 
val _ =  | 
| 7222 | 985  | 
OuterSyntax.improper_command "disable_pr" "disable printing of toplevel state" K.diag  | 
| 9010 | 986  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.disable_pr));  | 
| 5832 | 987  | 
|
| 24868 | 988  | 
val _ =  | 
| 7222 | 989  | 
OuterSyntax.improper_command "enable_pr" "enable printing of toplevel state" K.diag  | 
| 9010 | 990  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.enable_pr));  | 
| 7222 | 991  | 
|
| 24868 | 992  | 
val _ =  | 
| 6723 | 993  | 
OuterSyntax.improper_command "commit" "commit current session to ML database" K.diag  | 
| 9010 | 994  | 
(P.opt_unit >> (Toplevel.no_timing oo K IsarCmd.use_commit));  | 
| 5832 | 995  | 
|
| 24868 | 996  | 
val _ =  | 
| 6723 | 997  | 
OuterSyntax.improper_command "quit" "quit Isabelle" K.control  | 
| 9010 | 998  | 
(P.opt_unit >> (Toplevel.no_timing oo K IsarCmd.quit));  | 
| 5832 | 999  | 
|
| 24868 | 1000  | 
val _ =  | 
| 6723 | 1001  | 
OuterSyntax.improper_command "exit" "exit Isar loop" K.control  | 
| 9010 | 1002  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.exit));  | 
| 5832 | 1003  | 
|
| 24868 | 1004  | 
val _ =  | 
| 7102 | 1005  | 
OuterSyntax.improper_command "init_toplevel" "restart Isar toplevel loop" K.control  | 
| 9010 | 1006  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.init_toplevel));  | 
| 5991 | 1007  | 
|
| 24868 | 1008  | 
val _ =  | 
| 8678 | 1009  | 
OuterSyntax.improper_command "welcome" "print welcome message" K.diag  | 
| 9010 | 1010  | 
(Scan.succeed (Toplevel.no_timing o IsarCmd.welcome));  | 
| 7462 | 1011  | 
|
| 5832 | 1012  | 
end;  |