author | wenzelm |
Mon, 17 Apr 2000 14:06:05 +0200 | |
changeset 8723 | c7de3c2ed7a9 |
parent 8681 | 957a5fe9b212 |
child 8807 | 0046be1769f9 |
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 |
||
8 |
signature ISAR_SYN = |
|
9 |
sig |
|
10 |
val keywords: string list |
|
11 |
val parsers: OuterSyntax.parser list |
|
12 |
end; |
|
13 |
||
14 |
structure IsarSyn: ISAR_SYN = |
|
15 |
struct |
|
16 |
||
6723 | 17 |
structure P = OuterParse and K = OuterSyntax.Keyword; |
5832 | 18 |
|
19 |
||
20 |
(** init and exit **) |
|
21 |
||
22 |
val theoryP = |
|
6723 | 23 |
OuterSyntax.command "theory" "begin theory" K.thy_begin |
6245 | 24 |
(OuterSyntax.theory_header >> (Toplevel.print oo IsarThy.theory)); |
5832 | 25 |
|
6687 | 26 |
val end_excursionP = |
6723 | 27 |
OuterSyntax.command "end" "end current excursion" K.thy_end |
6687 | 28 |
(Scan.succeed (Toplevel.print o Toplevel.exit)); |
29 |
||
6245 | 30 |
val contextP = |
7102 | 31 |
OuterSyntax.improper_command "context" "switch theory context" K.thy_switch |
6723 | 32 |
(P.name >> (Toplevel.print oo IsarThy.context)); |
6245 | 33 |
|
5832 | 34 |
|
35 |
||
7749 | 36 |
(** markup commands **) |
5832 | 37 |
|
7749 | 38 |
val headerP = OuterSyntax.markup_command "header" "theory header" K.diag |
7733 | 39 |
(P.comment >> IsarThy.add_header); |
6353 | 40 |
|
7749 | 41 |
val chapterP = OuterSyntax.markup_command "chapter" "chapter heading" K.thy_heading |
6723 | 42 |
(P.comment >> (Toplevel.theory o IsarThy.add_chapter)); |
5958 | 43 |
|
7749 | 44 |
val sectionP = OuterSyntax.markup_command "section" "section heading" K.thy_heading |
6723 | 45 |
(P.comment >> (Toplevel.theory o IsarThy.add_section)); |
5958 | 46 |
|
7749 | 47 |
val subsectionP = OuterSyntax.markup_command "subsection" "subsection heading" K.thy_heading |
6723 | 48 |
(P.comment >> (Toplevel.theory o IsarThy.add_subsection)); |
5958 | 49 |
|
7749 | 50 |
val subsubsectionP = |
51 |
OuterSyntax.markup_command "subsubsection" "subsubsection heading" K.thy_heading |
|
52 |
(P.comment >> (Toplevel.theory o IsarThy.add_subsubsection)); |
|
5832 | 53 |
|
8661 | 54 |
val textP = OuterSyntax.markup_env_command "text" "formal comment (theory)" K.thy_decl |
7172 | 55 |
(P.comment >> (Toplevel.theory o IsarThy.add_text)); |
56 |
||
7862 | 57 |
val text_rawP = OuterSyntax.verbatim_command "text_raw" "raw document preparation text" |
58 |
K.thy_decl (P.comment >> (Toplevel.theory o IsarThy.add_text_raw)); |
|
7775 | 59 |
|
7172 | 60 |
|
7749 | 61 |
val sectP = OuterSyntax.markup_command "sect" "formal comment (proof)" K.prf_decl |
7172 | 62 |
(P.comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.add_sect))); |
63 |
||
7749 | 64 |
val subsectP = OuterSyntax.markup_command "subsect" "formal comment (proof)" K.prf_decl |
7172 | 65 |
(P.comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.add_subsect))); |
66 |
||
7749 | 67 |
val subsubsectP = OuterSyntax.markup_command "subsubsect" "formal comment (proof)" K.prf_decl |
7172 | 68 |
(P.comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.add_subsubsect))); |
69 |
||
8661 | 70 |
val txtP = OuterSyntax.markup_env_command "txt" "formal comment (proof)" K.prf_decl |
7172 | 71 |
(P.comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.add_txt))); |
72 |
||
7862 | 73 |
val txt_rawP = OuterSyntax.verbatim_command "txt_raw" "raw document preparation text (proof)" |
74 |
K.prf_decl (P.comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.add_txt_raw))); |
|
7775 | 75 |
|
5832 | 76 |
|
6886 | 77 |
|
78 |
(** theory sections **) |
|
79 |
||
5832 | 80 |
(* classes and sorts *) |
81 |
||
82 |
val classesP = |
|
6723 | 83 |
OuterSyntax.command "classes" "declare type classes" K.thy_decl |
84 |
(Scan.repeat1 (P.name -- Scan.optional (P.$$$ "<" |-- P.!!! (P.list1 P.xname)) []) |
|
6727 | 85 |
-- P.marg_comment >> (Toplevel.theory o IsarThy.add_classes)); |
5832 | 86 |
|
87 |
val classrelP = |
|
6723 | 88 |
OuterSyntax.command "classrel" "state inclusion of type classes (axiomatic!)" K.thy_decl |
6727 | 89 |
(P.xname -- (P.$$$ "<" |-- P.!!! P.xname) -- P.marg_comment |
90 |
>> (Toplevel.theory o IsarThy.add_classrel)); |
|
5832 | 91 |
|
92 |
val defaultsortP = |
|
6723 | 93 |
OuterSyntax.command "defaultsort" "declare default sort" K.thy_decl |
6727 | 94 |
(P.sort -- P.marg_comment >> (Toplevel.theory o IsarThy.add_defsort)); |
5832 | 95 |
|
96 |
||
97 |
(* types *) |
|
98 |
||
99 |
val typedeclP = |
|
6723 | 100 |
OuterSyntax.command "typedecl" "Pure type declaration" K.thy_decl |
6727 | 101 |
(P.type_args -- P.name -- P.opt_infix -- P.marg_comment >> (fn (((args, a), mx), cmt) => |
102 |
Toplevel.theory (IsarThy.add_typedecl ((a, args, mx), cmt)))); |
|
5832 | 103 |
|
104 |
val typeabbrP = |
|
6723 | 105 |
OuterSyntax.command "types" "declare type abbreviations" K.thy_decl |
6727 | 106 |
(Scan.repeat1 |
107 |
(P.type_args -- P.name -- (P.$$$ "=" |-- P.!!! (P.typ -- P.opt_infix)) -- P.marg_comment) |
|
108 |
>> (Toplevel.theory o IsarThy.add_tyabbrs o |
|
109 |
map (fn (((args, a), (T, mx)), cmt) => ((a, args, T, mx), cmt)))); |
|
5832 | 110 |
|
111 |
val nontermP = |
|
6370 | 112 |
OuterSyntax.command "nonterminals" "declare types treated as grammar nonterminal symbols" |
6727 | 113 |
K.thy_decl (Scan.repeat1 (P.name -- P.marg_comment) |
114 |
>> (Toplevel.theory o IsarThy.add_nonterminals)); |
|
5832 | 115 |
|
116 |
val aritiesP = |
|
6723 | 117 |
OuterSyntax.command "arities" "state type arities (axiomatic!)" K.thy_decl |
6727 | 118 |
(Scan.repeat1 ((P.xname -- (P.$$$ "::" |-- P.!!! P.arity) >> P.triple2) -- P.marg_comment) |
119 |
>> (Toplevel.theory o IsarThy.add_arities)); |
|
5832 | 120 |
|
121 |
||
122 |
(* consts and syntax *) |
|
123 |
||
8227 | 124 |
val judgmentP = |
125 |
OuterSyntax.command "judgment" "declare object-logic judgment" K.thy_decl |
|
126 |
(P.const -- P.marg_comment >> (Toplevel.theory o IsarThy.add_judgment)); |
|
127 |
||
5832 | 128 |
val constsP = |
6723 | 129 |
OuterSyntax.command "consts" "declare constants" K.thy_decl |
6727 | 130 |
(Scan.repeat1 (P.const -- P.marg_comment) >> (Toplevel.theory o IsarThy.add_consts)); |
5832 | 131 |
|
132 |
val opt_mode = |
|
133 |
Scan.optional |
|
6723 | 134 |
(P.$$$ "(" |-- P.!!! (P.name -- Scan.optional (P.$$$ "output" >> K false) true --| P.$$$ ")")) |
5832 | 135 |
("", true); |
136 |
||
137 |
val syntaxP = |
|
6723 | 138 |
OuterSyntax.command "syntax" "declare syntactic constants" K.thy_decl |
6727 | 139 |
(opt_mode -- Scan.repeat1 (P.const -- P.marg_comment) |
140 |
>> (Toplevel.theory o uncurry IsarThy.add_modesyntax)); |
|
5832 | 141 |
|
142 |
||
143 |
(* translations *) |
|
144 |
||
145 |
val trans_pat = |
|
6723 | 146 |
Scan.optional (P.$$$ "(" |-- P.!!! (P.xname --| P.$$$ ")")) "logic" -- P.string; |
5832 | 147 |
|
148 |
fun trans_arrow toks = |
|
6723 | 149 |
(P.$$$ "=>" >> K Syntax.ParseRule || |
150 |
P.$$$ "<=" >> K Syntax.PrintRule || |
|
151 |
P.$$$ "==" >> K Syntax.ParsePrintRule) toks; |
|
5832 | 152 |
|
153 |
val trans_line = |
|
6723 | 154 |
trans_pat -- P.!!! (trans_arrow -- trans_pat) |
5832 | 155 |
>> (fn (left, (arr, right)) => arr (left, right)); |
156 |
||
157 |
val translationsP = |
|
6723 | 158 |
OuterSyntax.command "translations" "declare syntax translation rules" K.thy_decl |
6727 | 159 |
(Scan.repeat1 (trans_line -- P.marg_comment) >> (Toplevel.theory o IsarThy.add_trrules)); |
5832 | 160 |
|
161 |
||
162 |
(* axioms and definitions *) |
|
163 |
||
164 |
val axiomsP = |
|
6723 | 165 |
OuterSyntax.command "axioms" "state arbitrary propositions (axiomatic!)" K.thy_decl |
6727 | 166 |
(Scan.repeat1 (P.spec_name -- P.marg_comment) >> (Toplevel.theory o IsarThy.add_axioms)); |
5832 | 167 |
|
168 |
val defsP = |
|
6723 | 169 |
OuterSyntax.command "defs" "define constants" K.thy_decl |
7603 | 170 |
(Scan.repeat1 (P.spec_name -- P.marg_comment) >> (Toplevel.theory o IsarThy.add_defs)); |
6370 | 171 |
|
172 |
val constdefsP = |
|
6723 | 173 |
OuterSyntax.command "constdefs" "declare and define constants" K.thy_decl |
6727 | 174 |
(Scan.repeat1 ((P.const -- P.marg_comment) -- (P.term -- P.marg_comment)) |
175 |
>> (Toplevel.theory o IsarThy.add_constdefs)); |
|
5832 | 176 |
|
177 |
||
5914 | 178 |
(* theorems *) |
179 |
||
6723 | 180 |
val facts = P.opt_thm_name "=" -- P.xthms1; |
5914 | 181 |
|
182 |
val theoremsP = |
|
6723 | 183 |
OuterSyntax.command "theorems" "define theorems" K.thy_decl |
6727 | 184 |
(facts -- P.marg_comment >> (Toplevel.theory o IsarThy.have_theorems)); |
5914 | 185 |
|
186 |
val lemmasP = |
|
6723 | 187 |
OuterSyntax.command "lemmas" "define lemmas" K.thy_decl |
6727 | 188 |
(facts -- P.marg_comment >> (Toplevel.theory o IsarThy.have_lemmas)); |
5914 | 189 |
|
190 |
||
5832 | 191 |
(* name space entry path *) |
192 |
||
193 |
val globalP = |
|
6723 | 194 |
OuterSyntax.command "global" "disable prefixing of theory name" K.thy_decl |
8723 | 195 |
(P.marg_comment >> (Toplevel.theory o IsarThy.global_path)); |
5832 | 196 |
|
197 |
val localP = |
|
6723 | 198 |
OuterSyntax.command "local" "enable prefixing of theory name" K.thy_decl |
8723 | 199 |
(P.marg_comment >> (Toplevel.theory o IsarThy.local_path)); |
200 |
||
201 |
val hideP = |
|
202 |
OuterSyntax.command "hide" "hide names from given name space" K.thy_decl |
|
203 |
(P.name -- Scan.repeat1 P.xname -- P.marg_comment >> (Toplevel.theory o IsarThy.hide_space)); |
|
5832 | 204 |
|
205 |
||
206 |
(* use ML text *) |
|
207 |
||
208 |
val useP = |
|
6723 | 209 |
OuterSyntax.command "use" "eval ML text from file" K.diag |
210 |
(P.name >> IsarCmd.use); |
|
5832 | 211 |
|
212 |
val mlP = |
|
7616 | 213 |
OuterSyntax.command "ML" "eval ML text (diagnostic)" K.diag |
7891 | 214 |
(P.text >> IsarCmd.use_mltext true); |
215 |
||
216 |
val ml_commandP = |
|
217 |
OuterSyntax.command "ML_command" "eval ML text" K.diag |
|
218 |
(P.text >> IsarCmd.use_mltext false); |
|
7616 | 219 |
|
220 |
val ml_setupP = |
|
221 |
OuterSyntax.command "ML_setup" "eval ML text (may change theory)" K.thy_decl |
|
222 |
(P.text >> IsarCmd.use_mltext_theory); |
|
5832 | 223 |
|
224 |
val setupP = |
|
7616 | 225 |
OuterSyntax.command "setup" "apply ML theory setup" K.thy_decl |
8349
611342539639
moved use_mltext, use_mltext_theory, use_let, use_setup to context.ML;
wenzelm
parents:
8235
diff
changeset
|
226 |
(P.text >> (Toplevel.theory o Context.use_setup)); |
5832 | 227 |
|
228 |
||
229 |
(* translation functions *) |
|
230 |
||
231 |
val parse_ast_translationP = |
|
6723 | 232 |
OuterSyntax.command "parse_ast_translation" "install parse ast translation functions" K.thy_decl |
233 |
(P.text >> (Toplevel.theory o IsarThy.parse_ast_translation)); |
|
5832 | 234 |
|
235 |
val parse_translationP = |
|
6723 | 236 |
OuterSyntax.command "parse_translation" "install parse translation functions" K.thy_decl |
237 |
(P.text >> (Toplevel.theory o IsarThy.parse_translation)); |
|
5832 | 238 |
|
239 |
val print_translationP = |
|
6723 | 240 |
OuterSyntax.command "print_translation" "install print translation functions" K.thy_decl |
241 |
(P.text >> (Toplevel.theory o IsarThy.print_translation)); |
|
5832 | 242 |
|
243 |
val typed_print_translationP = |
|
6370 | 244 |
OuterSyntax.command "typed_print_translation" "install typed print translation functions" |
6723 | 245 |
K.thy_decl (P.text >> (Toplevel.theory o IsarThy.typed_print_translation)); |
5832 | 246 |
|
247 |
val print_ast_translationP = |
|
6723 | 248 |
OuterSyntax.command "print_ast_translation" "install print ast translation functions" K.thy_decl |
249 |
(P.text >> (Toplevel.theory o IsarThy.print_ast_translation)); |
|
5832 | 250 |
|
251 |
val token_translationP = |
|
6723 | 252 |
OuterSyntax.command "token_translation" "install token translation functions" K.thy_decl |
253 |
(P.text >> (Toplevel.theory o IsarThy.token_translation)); |
|
5832 | 254 |
|
255 |
||
256 |
(* oracles *) |
|
257 |
||
258 |
val oracleP = |
|
6723 | 259 |
OuterSyntax.command "oracle" "install oracle" K.thy_decl |
7140 | 260 |
((P.name --| P.$$$ "=") -- P.text -- P.marg_comment >> (Toplevel.theory o IsarThy.add_oracle)); |
5832 | 261 |
|
262 |
||
263 |
||
264 |
(** proof commands **) |
|
265 |
||
266 |
(* statements *) |
|
267 |
||
6936 | 268 |
fun statement f = (P.opt_thm_name ":" -- P.propp >> P.triple1) -- P.marg_comment >> f; |
5832 | 269 |
|
270 |
val theoremP = |
|
6723 | 271 |
OuterSyntax.command "theorem" "state theorem" K.thy_goal |
5832 | 272 |
(statement IsarThy.theorem >> (fn f => Toplevel.print o Toplevel.theory_to_proof f)); |
273 |
||
274 |
val lemmaP = |
|
6723 | 275 |
OuterSyntax.command "lemma" "state lemma" K.thy_goal |
5832 | 276 |
(statement IsarThy.lemma >> (fn f => Toplevel.print o Toplevel.theory_to_proof f)); |
277 |
||
278 |
val showP = |
|
6723 | 279 |
OuterSyntax.command "show" "state local goal, solving current obligation" K.prf_goal |
5832 | 280 |
(statement IsarThy.show >> (fn f => Toplevel.print o Toplevel.proof f)); |
281 |
||
282 |
val haveP = |
|
6723 | 283 |
OuterSyntax.command "have" "state local goal" K.prf_goal |
5832 | 284 |
(statement IsarThy.have >> (fn f => Toplevel.print o Toplevel.proof f)); |
285 |
||
6501 | 286 |
val thusP = |
6723 | 287 |
OuterSyntax.command "thus" "abbreviates \"then show\"" K.prf_goal |
6501 | 288 |
(statement IsarThy.thus >> (fn f => Toplevel.print o Toplevel.proof f)); |
289 |
||
290 |
val henceP = |
|
6723 | 291 |
OuterSyntax.command "hence" "abbreviates \"then have\"" K.prf_goal |
6501 | 292 |
(statement IsarThy.hence >> (fn f => Toplevel.print o Toplevel.proof f)); |
293 |
||
5832 | 294 |
|
5914 | 295 |
(* facts *) |
5832 | 296 |
|
297 |
val thenP = |
|
6723 | 298 |
OuterSyntax.command "then" "forward chaining" K.prf_chain |
6727 | 299 |
(P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.chain))); |
5832 | 300 |
|
301 |
val fromP = |
|
6723 | 302 |
OuterSyntax.command "from" "forward chaining from given facts" K.prf_chain |
6727 | 303 |
(P.xthms1 -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.from_facts))); |
5914 | 304 |
|
6878 | 305 |
val withP = |
306 |
OuterSyntax.command "with" "forward chaining from given and current facts" K.prf_chain |
|
307 |
(P.xthms1 -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.with_facts))); |
|
308 |
||
6773 | 309 |
val noteP = |
6723 | 310 |
OuterSyntax.command "note" "define facts" K.prf_decl |
6755 | 311 |
(facts -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.have_facts))); |
5832 | 312 |
|
313 |
||
314 |
(* proof context *) |
|
315 |
||
316 |
val assumeP = |
|
6869 | 317 |
OuterSyntax.command "assume" "assume propositions" K.prf_asm |
7269 | 318 |
(P.and_list1 ((P.opt_thm_name ":" -- Scan.repeat1 P.propp >> P.triple1) -- P.marg_comment) |
6727 | 319 |
>> (Toplevel.print oo (Toplevel.proof o IsarThy.assume))); |
5832 | 320 |
|
6850 | 321 |
val presumeP = |
6869 | 322 |
OuterSyntax.command "presume" "assume propositions, to be established later" K.prf_asm |
7269 | 323 |
(P.and_list1 ((P.opt_thm_name ":" -- Scan.repeat1 P.propp >> P.triple1) -- P.marg_comment) |
6850 | 324 |
>> (Toplevel.print oo (Toplevel.proof o IsarThy.presume))); |
325 |
||
6953 | 326 |
val defP = |
327 |
OuterSyntax.command "def" "local definition" K.prf_asm |
|
328 |
((P.opt_thm_name ":" -- (P.name -- Scan.option (P.$$$ "::" |-- P.typ) -- |
|
6972 | 329 |
(P.$$$ "==" |-- P.termp)) >> P.triple1) -- P.marg_comment |
6953 | 330 |
>> (Toplevel.print oo (Toplevel.proof o IsarThy.local_def))); |
331 |
||
5832 | 332 |
val fixP = |
6869 | 333 |
OuterSyntax.command "fix" "fix variables (Skolem constants)" K.prf_asm |
7415 | 334 |
(P.and_list1 (Scan.repeat1 P.name -- Scan.option (P.$$$ "::" |-- P.typ) -- P.marg_comment) |
6727 | 335 |
>> (Toplevel.print oo (Toplevel.proof o IsarThy.fix))); |
5832 | 336 |
|
337 |
val letP = |
|
6723 | 338 |
OuterSyntax.command "let" "bind text variables" K.prf_decl |
8669 | 339 |
(P.and_list1 (P.enum1 "and" P.term -- (P.$$$ "=" |-- P.term) -- P.marg_comment) |
8615 | 340 |
>> (Toplevel.print oo (Toplevel.proof o IsarThy.let_bind))); |
5832 | 341 |
|
8370 | 342 |
val caseP = |
343 |
OuterSyntax.command "case" "invoke local context" K.prf_asm |
|
8450 | 344 |
(P.xthm -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.invoke_case))); |
8370 | 345 |
|
5832 | 346 |
|
347 |
(* proof structure *) |
|
348 |
||
349 |
val beginP = |
|
6723 | 350 |
OuterSyntax.command "{{" "begin explicit proof block" K.prf_block |
5832 | 351 |
(Scan.succeed (Toplevel.print o Toplevel.proof IsarThy.begin_block)); |
352 |
||
6687 | 353 |
val endP = |
6723 | 354 |
OuterSyntax.command "}}" "end explicit proof block" K.prf_block |
6687 | 355 |
(Scan.succeed (Toplevel.print o Toplevel.proof IsarThy.end_block)); |
356 |
||
5832 | 357 |
val nextP = |
6723 | 358 |
OuterSyntax.command "next" "enter next proof block" K.prf_block |
5832 | 359 |
(Scan.succeed (Toplevel.print o Toplevel.proof IsarThy.next_block)); |
360 |
||
361 |
||
362 |
(* end proof *) |
|
363 |
||
6404 | 364 |
val qedP = |
6735 | 365 |
OuterSyntax.command "qed" "conclude (sub-)proof" K.qed_block |
7002 | 366 |
(Scan.option (P.method -- P.interest) -- P.marg_comment >> IsarThy.qed); |
5832 | 367 |
|
6404 | 368 |
val terminal_proofP = |
6723 | 369 |
OuterSyntax.command "by" "terminal backward proof" K.qed |
7002 | 370 |
(P.method -- P.interest -- Scan.option (P.method -- P.interest) |
371 |
-- P.marg_comment >> IsarThy.terminal_proof); |
|
6404 | 372 |
|
373 |
val immediate_proofP = |
|
6723 | 374 |
OuterSyntax.command "." "immediate proof" K.qed |
7002 | 375 |
(P.marg_comment >> IsarThy.immediate_proof); |
6404 | 376 |
|
377 |
val default_proofP = |
|
6723 | 378 |
OuterSyntax.command ".." "default proof" K.qed |
7002 | 379 |
(P.marg_comment >> IsarThy.default_proof); |
5832 | 380 |
|
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
6886
diff
changeset
|
381 |
val skip_proofP = |
7172 | 382 |
OuterSyntax.improper_command "sorry" "skip proof (quick-and-dirty mode only!)" K.qed |
7002 | 383 |
(P.marg_comment >> IsarThy.skip_proof); |
6888
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
wenzelm
parents:
6886
diff
changeset
|
384 |
|
8210 | 385 |
val forget_proofP = |
8521 | 386 |
OuterSyntax.command "oops" "forget proof" K.qed_global |
8210 | 387 |
(P.marg_comment >> IsarThy.forget_proof); |
388 |
||
389 |
||
5832 | 390 |
|
391 |
(* proof steps *) |
|
392 |
||
8165 | 393 |
val deferP = |
8681 | 394 |
OuterSyntax.command "defer" "shuffle internal proof state" K.prf_script |
395 |
(Scan.option P.nat -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.defer))); |
|
8165 | 396 |
|
397 |
val preferP = |
|
8681 | 398 |
OuterSyntax.command "prefer" "shuffle internal proof state" K.prf_script |
399 |
(P.nat -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.prefer))); |
|
8165 | 400 |
|
6850 | 401 |
val applyP = |
8681 | 402 |
OuterSyntax.command "apply" "initial refinement step (unstructured)" K.prf_script |
403 |
(P.method -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.apply))); |
|
5832 | 404 |
|
8235 | 405 |
val apply_endP = |
8681 | 406 |
OuterSyntax.command "apply_end" "terminal refinement (unstructured)" K.prf_script |
407 |
(P.method -- P.marg_comment >> (Toplevel.print oo (Toplevel.proof o IsarThy.apply_end))); |
|
5832 | 408 |
|
6404 | 409 |
val proofP = |
6723 | 410 |
OuterSyntax.command "proof" "backward proof" K.prf_block |
7002 | 411 |
(P.interest -- Scan.option (P.method -- P.interest) -- P.marg_comment |
6727 | 412 |
>> (Toplevel.print oo (Toplevel.proof o IsarThy.proof))); |
5832 | 413 |
|
414 |
||
6773 | 415 |
(* calculational proof commands *) |
416 |
||
6878 | 417 |
val calc_args = |
418 |
Scan.option (P.$$$ "(" |-- P.!!! ((P.xthms1 --| P.$$$ ")") -- P.interest)); |
|
419 |
||
6773 | 420 |
val alsoP = |
8562 | 421 |
OuterSyntax.command "also" "combine calculation and current facts" K.prf_decl |
6878 | 422 |
(calc_args -- P.marg_comment >> IsarThy.also); |
6773 | 423 |
|
424 |
val finallyP = |
|
8562 | 425 |
OuterSyntax.command "finally" "combine calculation and current facts, exhibit result" K.prf_chain |
6878 | 426 |
(calc_args -- P.marg_comment >> IsarThy.finally); |
6773 | 427 |
|
8562 | 428 |
val moreoverP = |
429 |
OuterSyntax.command "moreover" "augment calculation by current facts" K.prf_decl |
|
430 |
(P.marg_comment >> IsarThy.moreover); |
|
431 |
||
8588 | 432 |
val ultimatelyP = |
433 |
OuterSyntax.command "ultimately" "augment calculation by current facts, exhibit result" |
|
434 |
K.prf_chain (P.marg_comment >> IsarThy.ultimately); |
|
435 |
||
6773 | 436 |
|
6742 | 437 |
(* proof navigation *) |
5944 | 438 |
|
5832 | 439 |
val backP = |
8235 | 440 |
OuterSyntax.command "back" "backtracking of proof command" K.prf_script |
7368 | 441 |
(Scan.optional (P.$$$ "!" >> K true) false >> |
442 |
(Toplevel.print oo (Toplevel.proof o ProofHistory.back))); |
|
5832 | 443 |
|
444 |
||
6742 | 445 |
(* history *) |
446 |
||
447 |
val cannot_undoP = |
|
448 |
OuterSyntax.improper_command "cannot_undo" "report 'cannot undo' error message" K.control |
|
7269 | 449 |
(P.name >> IsarCmd.cannot_undo); |
6742 | 450 |
|
7733 | 451 |
val clear_undosP = |
452 |
OuterSyntax.improper_command "clear_undos" "clear theory-level undo information" K.control |
|
453 |
(P.nat >> IsarCmd.clear_undos_theory); |
|
6742 | 454 |
|
455 |
val redoP = |
|
456 |
OuterSyntax.improper_command "redo" "redo last command" K.control |
|
457 |
(Scan.succeed (Toplevel.print o IsarCmd.redo)); |
|
458 |
||
459 |
val undos_proofP = |
|
460 |
OuterSyntax.improper_command "undos_proof" "undo last proof commands" K.control |
|
461 |
(P.nat >> (Toplevel.print oo IsarCmd.undos_proof)); |
|
462 |
||
463 |
val undoP = |
|
464 |
OuterSyntax.improper_command "undo" "undo last command" K.control |
|
8454 | 465 |
(Scan.succeed (Toplevel.print o IsarCmd.undo)); |
6742 | 466 |
|
8500 | 467 |
val killP = |
468 |
OuterSyntax.improper_command "kill" "kill current history node" K.control |
|
469 |
(Scan.succeed (Toplevel.print o IsarCmd.kill)); |
|
470 |
||
6742 | 471 |
|
5832 | 472 |
|
473 |
(** diagnostic commands (for interactive mode only) **) |
|
474 |
||
8464 | 475 |
val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) []; |
476 |
||
477 |
||
7124 | 478 |
val pretty_setmarginP = |
479 |
OuterSyntax.improper_command "pretty_setmargin" "change default margin for pretty printing" |
|
480 |
K.diag (P.nat >> IsarCmd.pretty_setmargin); |
|
481 |
||
5832 | 482 |
val print_commandsP = |
6723 | 483 |
OuterSyntax.improper_command "help" "print outer syntax (global)" K.diag |
7368 | 484 |
(Scan.succeed OuterSyntax.print_help); |
5832 | 485 |
|
7308 | 486 |
val print_contextP = |
487 |
OuterSyntax.improper_command "print_context" "print theory context name" K.diag |
|
488 |
(Scan.succeed IsarCmd.print_context); |
|
489 |
||
5832 | 490 |
val print_theoryP = |
6723 | 491 |
OuterSyntax.improper_command "print_theory" "print logical theory contents (verbose!)" K.diag |
5832 | 492 |
(Scan.succeed IsarCmd.print_theory); |
493 |
||
494 |
val print_syntaxP = |
|
6723 | 495 |
OuterSyntax.improper_command "print_syntax" "print inner syntax of theory (verbose!)" K.diag |
5832 | 496 |
(Scan.succeed IsarCmd.print_syntax); |
497 |
||
5881 | 498 |
val print_theoremsP = |
6723 | 499 |
OuterSyntax.improper_command "print_theorems" "print theorems known in this theory" K.diag |
5881 | 500 |
(Scan.succeed IsarCmd.print_theorems); |
501 |
||
5832 | 502 |
val print_attributesP = |
6723 | 503 |
OuterSyntax.improper_command "print_attributes" "print attributes known in this theory" K.diag |
5832 | 504 |
(Scan.succeed IsarCmd.print_attributes); |
505 |
||
506 |
val print_methodsP = |
|
6723 | 507 |
OuterSyntax.improper_command "print_methods" "print methods known in this theory" K.diag |
5832 | 508 |
(Scan.succeed IsarCmd.print_methods); |
509 |
||
7616 | 510 |
val thms_containingP = |
511 |
OuterSyntax.improper_command "thms_containing" "print theorems containing all given constants" |
|
512 |
K.diag (Scan.repeat P.xname >> IsarCmd.print_thms_containing); |
|
513 |
||
5832 | 514 |
val print_bindsP = |
6723 | 515 |
OuterSyntax.improper_command "print_binds" "print term bindings of proof context" K.diag |
5832 | 516 |
(Scan.succeed IsarCmd.print_binds); |
517 |
||
518 |
val print_lthmsP = |
|
8370 | 519 |
OuterSyntax.improper_command "print_facts" "print facts of proof context" K.diag |
5832 | 520 |
(Scan.succeed IsarCmd.print_lthms); |
521 |
||
8370 | 522 |
val print_casesP = |
523 |
OuterSyntax.improper_command "print_cases" "print cases of proof context" K.diag |
|
524 |
(Scan.succeed IsarCmd.print_cases); |
|
525 |
||
5896 | 526 |
val print_thmsP = |
8464 | 527 |
OuterSyntax.improper_command "thm" "print theorems" K.diag |
528 |
(opt_modes -- P.xthms1 >> IsarCmd.print_thms); |
|
5832 | 529 |
|
530 |
val print_propP = |
|
6723 | 531 |
OuterSyntax.improper_command "prop" "read and print proposition" K.diag |
8464 | 532 |
(opt_modes -- P.term >> IsarCmd.print_prop); |
5832 | 533 |
|
534 |
val print_termP = |
|
6723 | 535 |
OuterSyntax.improper_command "term" "read and print term" K.diag |
8464 | 536 |
(opt_modes -- P.term >> IsarCmd.print_term); |
5832 | 537 |
|
538 |
val print_typeP = |
|
6723 | 539 |
OuterSyntax.improper_command "typ" "read and print type" K.diag |
8464 | 540 |
(opt_modes -- P.typ >> IsarCmd.print_type); |
5832 | 541 |
|
542 |
||
543 |
||
544 |
(** system commands (for interactive mode only) **) |
|
545 |
||
546 |
val cdP = |
|
8650 | 547 |
OuterSyntax.improper_command "cd" "change current working directory" K.diag |
6723 | 548 |
(P.name >> IsarCmd.cd); |
5832 | 549 |
|
550 |
val pwdP = |
|
6723 | 551 |
OuterSyntax.improper_command "pwd" "print current working directory" K.diag |
5832 | 552 |
(Scan.succeed IsarCmd.pwd); |
553 |
||
554 |
val use_thyP = |
|
6723 | 555 |
OuterSyntax.improper_command "use_thy" "use theory file" K.diag |
556 |
(P.name >> IsarCmd.use_thy); |
|
5832 | 557 |
|
6694 | 558 |
val use_thy_onlyP = |
7102 | 559 |
OuterSyntax.improper_command "use_thy_only" "use theory file only, ignoring associated ML" |
560 |
K.diag (P.name >> IsarCmd.use_thy_only); |
|
6694 | 561 |
|
6196 | 562 |
val update_thyP = |
6723 | 563 |
OuterSyntax.improper_command "update_thy" "update theory file" K.diag |
564 |
(P.name >> IsarCmd.update_thy); |
|
5832 | 565 |
|
7102 | 566 |
val update_thy_onlyP = |
567 |
OuterSyntax.improper_command "update_thy_only" "update theory file, ignoring associated ML" |
|
7885 | 568 |
K.diag (P.name >> IsarCmd.update_thy_only); |
7102 | 569 |
|
570 |
val touch_thyP = |
|
571 |
OuterSyntax.improper_command "touch_thy" "outdate theory, including descendants" K.diag |
|
572 |
(P.name >> IsarCmd.touch_thy); |
|
573 |
||
574 |
val touch_all_thysP = |
|
575 |
OuterSyntax.improper_command "touch_all_thys" "outdate all non-finished theories" K.diag |
|
576 |
(Scan.succeed IsarCmd.touch_all_thys); |
|
577 |
||
7908 | 578 |
val touch_child_thysP = |
579 |
OuterSyntax.improper_command "touch_child_thys" "outdate child theories" K.diag |
|
580 |
(P.name >> IsarCmd.touch_child_thys); |
|
581 |
||
7102 | 582 |
val remove_thyP = |
583 |
OuterSyntax.improper_command "remove_thy" "remove theory from loader database" K.diag |
|
584 |
(P.name >> IsarCmd.remove_thy); |
|
585 |
||
7931 | 586 |
val kill_thyP = |
587 |
OuterSyntax.improper_command "kill_thy" "kill theory -- try to remove from loader database" |
|
588 |
K.diag (P.name >> IsarCmd.kill_thy); |
|
589 |
||
5832 | 590 |
val prP = |
6723 | 591 |
OuterSyntax.improper_command "pr" "print current toplevel state" K.diag |
8464 | 592 |
(opt_modes -- Scan.option P.nat >> IsarCmd.pr); |
7199 | 593 |
|
7222 | 594 |
val disable_prP = |
595 |
OuterSyntax.improper_command "disable_pr" "disable printing of toplevel state" K.diag |
|
8454 | 596 |
(Scan.succeed IsarCmd.disable_pr); |
5832 | 597 |
|
7222 | 598 |
val enable_prP = |
599 |
OuterSyntax.improper_command "enable_pr" "enable printing of toplevel state" K.diag |
|
8454 | 600 |
(Scan.succeed IsarCmd.enable_pr); |
7222 | 601 |
|
5832 | 602 |
val commitP = |
6723 | 603 |
OuterSyntax.improper_command "commit" "commit current session to ML database" K.diag |
7931 | 604 |
(P.opt_unit >> (K IsarCmd.use_commit)); |
5832 | 605 |
|
606 |
val quitP = |
|
6723 | 607 |
OuterSyntax.improper_command "quit" "quit Isabelle" K.control |
7931 | 608 |
(P.opt_unit >> (K IsarCmd.quit)); |
5832 | 609 |
|
610 |
val exitP = |
|
6723 | 611 |
OuterSyntax.improper_command "exit" "exit Isar loop" K.control |
5832 | 612 |
(Scan.succeed IsarCmd.exit); |
613 |
||
7102 | 614 |
val init_toplevelP = |
615 |
OuterSyntax.improper_command "init_toplevel" "restart Isar toplevel loop" K.control |
|
616 |
(Scan.succeed IsarCmd.init_toplevel); |
|
5991 | 617 |
|
7462 | 618 |
val welcomeP = |
8678 | 619 |
OuterSyntax.improper_command "welcome" "print welcome message" K.diag |
7462 | 620 |
(Scan.succeed IsarCmd.welcome); |
621 |
||
5832 | 622 |
|
623 |
||
624 |
(** the Pure outer syntax **) |
|
625 |
||
626 |
(*keep keywords consistent with the parsers, including those in |
|
627 |
outer_parse.ML, otherwise be prepared for unexpected errors*) |
|
628 |
||
629 |
val keywords = |
|
7415 | 630 |
["!", "!!", "%", "%%", "(", ")", "+", ",", "--", ":", "::", ";", "<", |
8669 | 631 |
"<=", "=", "==", "=>", "?", "[", "]", "and", "binder", "concl", |
632 |
"files", "in", "infixl", "infixr", "is", "output", "{", "|", "}"]; |
|
5832 | 633 |
|
634 |
val parsers = [ |
|
635 |
(*theory structure*) |
|
8500 | 636 |
theoryP, end_excursionP, contextP, |
7775 | 637 |
(*markup commands*) |
7733 | 638 |
headerP, chapterP, sectionP, subsectionP, subsubsectionP, textP, |
7862 | 639 |
text_rawP, sectP, subsectP, subsubsectP, txtP, txt_rawP, |
5832 | 640 |
(*theory sections*) |
7172 | 641 |
classesP, classrelP, defaultsortP, typedeclP, typeabbrP, nontermP, |
8227 | 642 |
aritiesP, judgmentP, constsP, syntaxP, translationsP, axiomsP, |
8723 | 643 |
defsP, constdefsP, theoremsP, lemmasP, globalP, localP, hideP, useP, |
644 |
mlP, ml_commandP, ml_setupP, setupP, parse_ast_translationP, |
|
7891 | 645 |
parse_translationP, print_translationP, typed_print_translationP, |
7616 | 646 |
print_ast_translationP, token_translationP, oracleP, |
5832 | 647 |
(*proof commands*) |
6850 | 648 |
theoremP, lemmaP, showP, haveP, thusP, henceP, assumeP, presumeP, |
8464 | 649 |
defP, fixP, letP, caseP, thenP, fromP, withP, noteP, beginP, endP, |
650 |
nextP, qedP, terminal_proofP, immediate_proofP, default_proofP, |
|
651 |
skip_proofP, forget_proofP, deferP, preferP, applyP, apply_endP, |
|
8588 | 652 |
proofP, alsoP, finallyP, moreoverP, ultimatelyP, backP, |
653 |
cannot_undoP, clear_undosP, redoP, undos_proofP, undoP, killP, |
|
5832 | 654 |
(*diagnostic commands*) |
7308 | 655 |
pretty_setmarginP, print_commandsP, print_contextP, print_theoryP, |
7616 | 656 |
print_syntaxP, print_theoremsP, print_attributesP, print_methodsP, |
8370 | 657 |
thms_containingP, print_bindsP, print_lthmsP, print_casesP, |
658 |
print_thmsP, print_propP, print_termP, print_typeP, |
|
5832 | 659 |
(*system commands*) |
7102 | 660 |
cdP, pwdP, use_thyP, use_thy_onlyP, update_thyP, update_thy_onlyP, |
7931 | 661 |
touch_thyP, touch_all_thysP, touch_child_thysP, remove_thyP, |
662 |
kill_thyP, prP, disable_prP, enable_prP, commitP, quitP, exitP, |
|
663 |
init_toplevelP, welcomeP]; |
|
5832 | 664 |
|
665 |
||
666 |
end; |
|
667 |
||
668 |
||
6196 | 669 |
(*install the Pure outer syntax*) |
5832 | 670 |
OuterSyntax.add_keywords IsarSyn.keywords; |
671 |
OuterSyntax.add_parsers IsarSyn.parsers; |