| author | wenzelm | 
| Mon, 12 May 2014 17:17:32 +0200 | |
| changeset 56941 | 952833323c99 | 
| parent 56895 | f058120aaad4 | 
| child 57181 | 2d13bf9ea77b | 
| permissions | -rw-r--r-- | 
| 5832 | 1  | 
(* Title: Pure/Isar/isar_syn.ML  | 
| 52546 | 2  | 
Author: Makarius  | 
| 5832 | 3  | 
|
| 52546 | 4  | 
Outer syntax for Isabelle/Pure.  | 
| 5832 | 5  | 
*)  | 
6  | 
||
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
7  | 
structure Isar_Syn: sig end =  | 
| 5832 | 8  | 
struct  | 
9  | 
||
| 7749 | 10  | 
(** markup commands **)  | 
| 5832 | 11  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
12  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
13  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 14  | 
    @{command_spec "header"} "theory header"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
15  | 
(Parse.document_source >> Isar_Cmd.header_markup);  | 
| 6353 | 16  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
17  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
18  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 19  | 
    @{command_spec "chapter"} "chapter heading"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
20  | 
(Parse.opt_target -- Parse.document_source >> Isar_Cmd.local_theory_markup);  | 
| 5958 | 21  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
22  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
23  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 24  | 
    @{command_spec "section"} "section heading"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
25  | 
(Parse.opt_target -- Parse.document_source >> Isar_Cmd.local_theory_markup);  | 
| 5958 | 26  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
27  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
28  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 29  | 
    @{command_spec "subsection"} "subsection heading"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
30  | 
(Parse.opt_target -- Parse.document_source >> Isar_Cmd.local_theory_markup);  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
31  | 
|
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
32  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
33  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 34  | 
    @{command_spec "subsubsection"} "subsubsection heading"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
35  | 
(Parse.opt_target -- Parse.document_source >> Isar_Cmd.local_theory_markup);  | 
| 5958 | 36  | 
|
| 24868 | 37  | 
val _ =  | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
38  | 
Outer_Syntax.markup_command Thy_Output.MarkupEnv  | 
| 48642 | 39  | 
    @{command_spec "text"} "formal comment (theory)"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
40  | 
(Parse.opt_target -- Parse.document_source >> Isar_Cmd.local_theory_markup);  | 
| 7172 | 41  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
42  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
43  | 
Outer_Syntax.markup_command Thy_Output.Verbatim  | 
| 48642 | 44  | 
    @{command_spec "text_raw"} "raw document preparation text"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
45  | 
(Parse.opt_target -- Parse.document_source >> Isar_Cmd.local_theory_markup);  | 
| 7172 | 46  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
47  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
48  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 49  | 
    @{command_spec "sect"} "formal comment (proof)"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
50  | 
(Parse.document_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 51  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
52  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
53  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 54  | 
    @{command_spec "subsect"} "formal comment (proof)"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
55  | 
(Parse.document_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 56  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
57  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
58  | 
Outer_Syntax.markup_command Thy_Output.Markup  | 
| 48642 | 59  | 
    @{command_spec "subsubsect"} "formal comment (proof)"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
60  | 
(Parse.document_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 61  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
62  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
63  | 
Outer_Syntax.markup_command Thy_Output.MarkupEnv  | 
| 48642 | 64  | 
    @{command_spec "txt"} "formal comment (proof)"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
65  | 
(Parse.document_source >> Isar_Cmd.proof_markup);  | 
| 7172 | 66  | 
|
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
67  | 
val _ =  | 
| 
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
68  | 
Outer_Syntax.markup_command Thy_Output.Verbatim  | 
| 48642 | 69  | 
    @{command_spec "txt_raw"} "raw document preparation text (proof)"
 | 
| 
51627
 
589daaf48dba
tuned signature -- agree with markup terminology;
 
wenzelm 
parents: 
51585 
diff
changeset
 | 
70  | 
(Parse.document_source >> Isar_Cmd.proof_markup);  | 
| 7775 | 71  | 
|
| 5832 | 72  | 
|
| 6886 | 73  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
74  | 
(** theory commands **)  | 
| 6886 | 75  | 
|
| 
55385
 
169e12bbf9a3
discontinued axiomatic 'classes', 'classrel', 'arities';
 
wenzelm 
parents: 
55141 
diff
changeset
 | 
76  | 
(* sorts *)  | 
| 5832 | 77  | 
|
| 24868 | 78  | 
val _ =  | 
| 48642 | 79  | 
  Outer_Syntax.local_theory @{command_spec "default_sort"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
80  | 
"declare default sort for explicit type variables"  | 
| 36950 | 81  | 
(Parse.sort >> (fn s => fn lthy => Local_Theory.set_defsort (Syntax.read_sort lthy s) lthy));  | 
| 5832 | 82  | 
|
83  | 
||
84  | 
(* types *)  | 
|
85  | 
||
| 24868 | 86  | 
val _ =  | 
| 48642 | 87  | 
  Outer_Syntax.local_theory @{command_spec "typedecl"} "type declaration"
 | 
| 36950 | 88  | 
(Parse.type_args -- Parse.binding -- Parse.opt_mixfix  | 
| 35835 | 89  | 
>> (fn ((args, a), mx) => Typedecl.typedecl (a, map (rpair dummyS) args, mx) #> snd));  | 
| 5832 | 90  | 
|
| 24868 | 91  | 
val _ =  | 
| 48642 | 92  | 
  Outer_Syntax.local_theory @{command_spec "type_synonym"} "declare type abbreviation"
 | 
| 45837 | 93  | 
(Parse.type_args -- Parse.binding --  | 
| 48643 | 94  | 
      (@{keyword "="} |-- Parse.!!! (Parse.typ -- Parse.opt_mixfix'))
 | 
| 45837 | 95  | 
>> (fn ((args, a), (rhs, mx)) => snd o Typedecl.abbrev_cmd (a, args, mx) rhs));  | 
| 5832 | 96  | 
|
| 24868 | 97  | 
val _ =  | 
| 48642 | 98  | 
  Outer_Syntax.command @{command_spec "nonterminal"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
99  | 
"declare syntactic type constructors (grammar nonterminal symbols)"  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
42359 
diff
changeset
 | 
100  | 
(Parse.and_list1 Parse.binding >> (Toplevel.theory o Sign.add_nonterminals_global));  | 
| 5832 | 101  | 
|
102  | 
||
103  | 
(* consts and syntax *)  | 
|
104  | 
||
| 24868 | 105  | 
val _ =  | 
| 48642 | 106  | 
  Outer_Syntax.command @{command_spec "judgment"} "declare object-logic judgment"
 | 
| 36950 | 107  | 
(Parse.const_binding >> (Toplevel.theory o Object_Logic.add_judgment_cmd));  | 
| 8227 | 108  | 
|
| 24868 | 109  | 
val _ =  | 
| 48642 | 110  | 
  Outer_Syntax.command @{command_spec "consts"} "declare constants"
 | 
| 56239 | 111  | 
(Scan.repeat1 Parse.const_binding >> (Toplevel.theory o Sign.add_consts_cmd));  | 
| 5832 | 112  | 
|
| 12142 | 113  | 
val mode_spec =  | 
| 48643 | 114  | 
  (@{keyword "output"} >> K ("", false)) ||
 | 
115  | 
    Parse.name -- Scan.optional (@{keyword "output"} >> K false) true;
 | 
|
| 9731 | 116  | 
|
| 14900 | 117  | 
val opt_mode =  | 
| 48643 | 118  | 
  Scan.optional (@{keyword "("} |-- Parse.!!! (mode_spec --| @{keyword ")"})) Syntax.mode_default;
 | 
| 5832 | 119  | 
|
| 24868 | 120  | 
val _ =  | 
| 50214 | 121  | 
  Outer_Syntax.command @{command_spec "syntax"} "add raw syntax clauses"
 | 
| 56239 | 122  | 
(opt_mode -- Scan.repeat1 Parse.const_decl  | 
| 56240 | 123  | 
>> (Toplevel.theory o uncurry Sign.add_syntax_cmd));  | 
| 5832 | 124  | 
|
| 24868 | 125  | 
val _ =  | 
| 50214 | 126  | 
  Outer_Syntax.command @{command_spec "no_syntax"} "delete raw syntax clauses"
 | 
| 56239 | 127  | 
(opt_mode -- Scan.repeat1 Parse.const_decl  | 
| 56240 | 128  | 
>> (Toplevel.theory o uncurry Sign.del_syntax_cmd));  | 
| 15748 | 129  | 
|
| 5832 | 130  | 
|
131  | 
(* translations *)  | 
|
132  | 
||
133  | 
val trans_pat =  | 
|
| 42326 | 134  | 
Scan.optional  | 
| 56006 | 135  | 
    (@{keyword "("} |-- Parse.!!! (Parse.inner_syntax Parse.xname --| @{keyword ")"})) "logic"
 | 
| 42326 | 136  | 
-- Parse.inner_syntax Parse.string;  | 
| 5832 | 137  | 
|
138  | 
fun trans_arrow toks =  | 
|
| 48643 | 139  | 
  ((@{keyword "\<rightharpoonup>"} || @{keyword "=>"}) >> K Syntax.Parse_Rule ||
 | 
140  | 
    (@{keyword "\<leftharpoondown>"} || @{keyword "<="}) >> K Syntax.Print_Rule ||
 | 
|
141  | 
    (@{keyword "\<rightleftharpoons>"} || @{keyword "=="}) >> K Syntax.Parse_Print_Rule) toks;
 | 
|
| 5832 | 142  | 
|
143  | 
val trans_line =  | 
|
| 36950 | 144  | 
trans_pat -- Parse.!!! (trans_arrow -- trans_pat)  | 
| 5832 | 145  | 
>> (fn (left, (arr, right)) => arr (left, right));  | 
146  | 
||
| 24868 | 147  | 
val _ =  | 
| 50214 | 148  | 
  Outer_Syntax.command @{command_spec "translations"} "add syntax translation rules"
 | 
| 42204 | 149  | 
(Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.translations));  | 
| 5832 | 150  | 
|
| 24868 | 151  | 
val _ =  | 
| 50214 | 152  | 
  Outer_Syntax.command @{command_spec "no_translations"} "delete syntax translation rules"
 | 
| 42204 | 153  | 
(Scan.repeat1 trans_line >> (Toplevel.theory o Isar_Cmd.no_translations));  | 
| 19260 | 154  | 
|
| 5832 | 155  | 
|
156  | 
(* axioms and definitions *)  | 
|
157  | 
||
| 19631 | 158  | 
val opt_unchecked_overloaded =  | 
| 48643 | 159  | 
  Scan.optional (@{keyword "("} |-- Parse.!!!
 | 
160  | 
    (((@{keyword "unchecked"} >> K true) --
 | 
|
161  | 
        Scan.optional (@{keyword "overloaded"} >> K true) false ||
 | 
|
162  | 
      @{keyword "overloaded"} >> K (false, true)) --| @{keyword ")"})) (false, false);
 | 
|
| 19631 | 163  | 
|
| 24868 | 164  | 
val _ =  | 
| 48642 | 165  | 
  Outer_Syntax.command @{command_spec "defs"} "define constants"
 | 
| 
35852
 
4e3fe0b8687b
minor renovation of old-style 'axioms' -- make it an alias of iterated 'axiomatization';
 
wenzelm 
parents: 
35835 
diff
changeset
 | 
166  | 
(opt_unchecked_overloaded --  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
167  | 
Scan.repeat1 (Parse_Spec.thm_name ":" -- Parse.prop >> (fn ((x, y), z) => ((x, z), y)))  | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
168  | 
>> (Toplevel.theory o Isar_Cmd.add_defs));  | 
| 6370 | 169  | 
|
| 14642 | 170  | 
|
| 
21601
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
171  | 
(* constant definitions and abbreviations *)  | 
| 
 
6588b947d631
simplified syntax for 'definition', 'abbreviation';
 
wenzelm 
parents: 
21527 
diff
changeset
 | 
172  | 
|
| 24868 | 173  | 
val _ =  | 
| 48642 | 174  | 
  Outer_Syntax.local_theory' @{command_spec "definition"} "constant definition"
 | 
| 
44192
 
a32ca9165928
less verbosity in batch mode -- spam reduction and notable performance improvement;
 
wenzelm 
parents: 
44187 
diff
changeset
 | 
175  | 
(Parse_Spec.constdef >> (fn args => #2 oo Specification.definition_cmd args));  | 
| 18780 | 176  | 
|
| 24868 | 177  | 
val _ =  | 
| 48642 | 178  | 
  Outer_Syntax.local_theory' @{command_spec "abbreviation"} "constant abbreviation"
 | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
179  | 
(opt_mode -- (Scan.option Parse_Spec.constdecl -- Parse.prop)  | 
| 36950 | 180  | 
>> (fn (mode, args) => Specification.abbreviation_cmd mode args));  | 
| 19659 | 181  | 
|
| 24868 | 182  | 
val _ =  | 
| 48642 | 183  | 
  Outer_Syntax.local_theory @{command_spec "type_notation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
184  | 
"add concrete syntax for type constructors"  | 
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
185  | 
(opt_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix)  | 
| 36950 | 186  | 
>> (fn (mode, args) => Specification.type_notation_cmd true mode args));  | 
| 35413 | 187  | 
|
188  | 
val _ =  | 
|
| 48642 | 189  | 
  Outer_Syntax.local_theory @{command_spec "no_type_notation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
190  | 
"delete concrete syntax for type constructors"  | 
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
191  | 
(opt_mode -- Parse.and_list1 (Parse.type_const -- Parse.mixfix)  | 
| 36950 | 192  | 
>> (fn (mode, args) => Specification.type_notation_cmd false mode args));  | 
| 35413 | 193  | 
|
194  | 
val _ =  | 
|
| 48642 | 195  | 
  Outer_Syntax.local_theory @{command_spec "notation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
196  | 
"add concrete syntax for constants / fixed variables"  | 
| 
51654
 
8450b944e58a
just one syntax category "mixfix" -- check structure annotation semantically;
 
wenzelm 
parents: 
51627 
diff
changeset
 | 
197  | 
(opt_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix)  | 
| 36950 | 198  | 
>> (fn (mode, args) => Specification.notation_cmd true mode args));  | 
| 24950 | 199  | 
|
200  | 
val _ =  | 
|
| 48642 | 201  | 
  Outer_Syntax.local_theory @{command_spec "no_notation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
202  | 
"delete concrete syntax for constants / fixed variables"  | 
| 
51654
 
8450b944e58a
just one syntax category "mixfix" -- check structure annotation semantically;
 
wenzelm 
parents: 
51627 
diff
changeset
 | 
203  | 
(opt_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix)  | 
| 36950 | 204  | 
>> (fn (mode, args) => Specification.notation_cmd false mode args));  | 
| 19076 | 205  | 
|
| 5832 | 206  | 
|
| 18616 | 207  | 
(* constant specifications *)  | 
208  | 
||
| 24868 | 209  | 
val _ =  | 
| 48642 | 210  | 
  Outer_Syntax.command @{command_spec "axiomatization"} "axiomatic constant specification"
 | 
| 36950 | 211  | 
(Scan.optional Parse.fixes [] --  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
212  | 
Scan.optional (Parse.where_ |-- Parse.!!! (Parse.and_list1 Parse_Spec.specs)) []  | 
| 36950 | 213  | 
>> (fn (x, y) => Toplevel.theory (#2 o Specification.axiomatization_cmd x y)));  | 
| 18616 | 214  | 
|
215  | 
||
| 5914 | 216  | 
(* theorems *)  | 
217  | 
||
| 
26988
 
742e26213212
more uniform treatment of OuterSyntax.local_theory commands;
 
wenzelm 
parents: 
26888 
diff
changeset
 | 
218  | 
fun theorems kind =  | 
| 
45600
 
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
 
wenzelm 
parents: 
45488 
diff
changeset
 | 
219  | 
Parse_Spec.name_facts -- Parse.for_fixes  | 
| 
 
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
 
wenzelm 
parents: 
45488 
diff
changeset
 | 
220  | 
>> (fn (facts, fixes) => #2 oo Specification.theorems_cmd kind facts fixes);  | 
| 12712 | 221  | 
|
| 24868 | 222  | 
val _ =  | 
| 48642 | 223  | 
  Outer_Syntax.local_theory' @{command_spec "theorems"} "define theorems"
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
224  | 
(theorems Thm.theoremK);  | 
| 5914 | 225  | 
|
| 24868 | 226  | 
val _ =  | 
| 48642 | 227  | 
  Outer_Syntax.local_theory' @{command_spec "lemmas"} "define lemmas" (theorems Thm.lemmaK);
 | 
| 5914 | 228  | 
|
| 24868 | 229  | 
val _ =  | 
| 48642 | 230  | 
  Outer_Syntax.local_theory' @{command_spec "declare"} "declare theorems"
 | 
| 
45600
 
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
 
wenzelm 
parents: 
45488 
diff
changeset
 | 
231  | 
(Parse.and_list1 Parse_Spec.xthms1 -- Parse.for_fixes  | 
| 
 
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
 
wenzelm 
parents: 
45488 
diff
changeset
 | 
232  | 
>> (fn (facts, fixes) =>  | 
| 
 
1bbbac9a0cb0
'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
 
wenzelm 
parents: 
45488 
diff
changeset
 | 
233  | 
#2 oo Specification.theorems_cmd "" [(Attrib.empty_binding, flat facts)] fixes));  | 
| 9589 | 234  | 
|
| 5914 | 235  | 
|
| 
56005
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
236  | 
(* hide names *)  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
237  | 
|
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
238  | 
local  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
239  | 
|
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
240  | 
fun hide_names command_spec what hide parse prep =  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
241  | 
  Outer_Syntax.command command_spec ("hide " ^ what ^ " from name space")
 | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
242  | 
((Parse.opt_keyword "open" >> not) -- Scan.repeat1 parse >> (fn (fully, args) =>  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
243  | 
(Toplevel.theory (fn thy =>  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
244  | 
let val ctxt = Proof_Context.init_global thy  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
245  | 
in fold (hide fully o prep ctxt) args thy end))));  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
246  | 
|
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
247  | 
in  | 
| 5832 | 248  | 
|
| 
56005
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
249  | 
val _ =  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
250  | 
  hide_names @{command_spec "hide_class"} "classes" Sign.hide_class Parse.class
 | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
251  | 
Proof_Context.read_class;  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
252  | 
|
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
253  | 
val _ =  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
254  | 
  hide_names @{command_spec "hide_type"} "types" Sign.hide_type Parse.type_const
 | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
255  | 
    ((#1 o dest_Type) oo Proof_Context.read_type_name {proper = true, strict = false});
 | 
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36174 
diff
changeset
 | 
256  | 
|
| 
56005
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
257  | 
val _ =  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
258  | 
  hide_names @{command_spec "hide_const"} "constants" Sign.hide_const Parse.const
 | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
259  | 
    ((#1 o dest_Const) oo Proof_Context.read_const {proper = true, strict = false});
 | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
260  | 
|
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
261  | 
val _ =  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
262  | 
  hide_names @{command_spec "hide_fact"} "facts" Global_Theory.hide_fact
 | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
263  | 
(Parse.position Parse.xname) (Global_Theory.check_fact o Proof_Context.theory_of);  | 
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
264  | 
|
| 
 
4f4fc80b0613
simplified / modernized hide commands: proper outer parsers and PIDE markup via check;
 
wenzelm 
parents: 
55997 
diff
changeset
 | 
265  | 
end;  | 
| 5832 | 266  | 
|
267  | 
||
268  | 
(* use ML text *)  | 
|
269  | 
||
| 24868 | 270  | 
val _ =  | 
| 
56275
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
271  | 
  Outer_Syntax.command @{command_spec "SML_file"} "read and evaluate Standard ML file"
 | 
| 
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
272  | 
(Resources.provide_parse_files "SML_file" >> (fn files => Toplevel.theory (fn thy =>  | 
| 
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
273  | 
let  | 
| 
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
274  | 
          val ([{lines, pos, ...}], thy') = files thy;
 | 
| 
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
275  | 
          val source = {delimited = true, text = cat_lines lines, pos = pos};
 | 
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
276  | 
          val flags = {SML = true, exchange = false, redirect = true, verbose = true};
 | 
| 
56275
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
277  | 
in  | 
| 
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
278  | 
thy' |> Context.theory_map  | 
| 
56304
 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 
wenzelm 
parents: 
56275 
diff
changeset
 | 
279  | 
(ML_Context.exec (fn () => ML_Context.eval_source flags source))  | 
| 
56275
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
280  | 
end)));  | 
| 
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
281  | 
|
| 
 
600f432ab556
added command 'SML_file' for Standard ML without Isabelle/ML add-ons;
 
wenzelm 
parents: 
56240 
diff
changeset
 | 
282  | 
val _ =  | 
| 
56618
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
283  | 
  Outer_Syntax.command @{command_spec "SML_export"} "evaluate SML within Isabelle/ML environment"
 | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
284  | 
(Parse.ML_source >> (fn source =>  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
285  | 
      let val flags = {SML = true, exchange = true, redirect = false, verbose = true} in
 | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
286  | 
Toplevel.theory  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
287  | 
(Context.theory_map (ML_Context.exec (fn () => ML_Context.eval_source flags source)))  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
288  | 
end));  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
289  | 
|
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
290  | 
val _ =  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
291  | 
  Outer_Syntax.command @{command_spec "SML_import"} "evaluate Isabelle/ML within SML environment"
 | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
292  | 
(Parse.ML_source >> (fn source =>  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
293  | 
      let val flags = {SML = false, exchange = true, redirect = false, verbose = true} in
 | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
294  | 
Toplevel.generic_theory  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
295  | 
(ML_Context.exec (fn () => ML_Context.eval_source flags source) #>  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
296  | 
Local_Theory.propagate_ml_env)  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
297  | 
end));  | 
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
298  | 
|
| 
 
874bdedb2313
added command 'SML_export' and 'SML_import' for exchange of toplevel bindings;
 
wenzelm 
parents: 
56467 
diff
changeset
 | 
299  | 
val _ =  | 
| 48642 | 300  | 
  Outer_Syntax.command @{command_spec "ML"} "ML text within theory or local theory"
 | 
| 
55828
 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 
wenzelm 
parents: 
55795 
diff
changeset
 | 
301  | 
(Parse.ML_source >> (fn source =>  | 
| 
30579
 
4169ddbfe1f9
command 'use', 'ML': apply ML environment to theory and target as well;
 
wenzelm 
parents: 
30576 
diff
changeset
 | 
302  | 
Toplevel.generic_theory  | 
| 
56304
 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 
wenzelm 
parents: 
56275 
diff
changeset
 | 
303  | 
(ML_Context.exec (fn () =>  | 
| 
 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 
wenzelm 
parents: 
56275 
diff
changeset
 | 
304  | 
ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source) #>  | 
| 
37949
 
48a874444164
moved management of auxiliary theory source files to Thy_Load -- as theory data instead of accidental loader state;
 
wenzelm 
parents: 
37873 
diff
changeset
 | 
305  | 
Local_Theory.propagate_ml_env)));  | 
| 
30579
 
4169ddbfe1f9
command 'use', 'ML': apply ML environment to theory and target as well;
 
wenzelm 
parents: 
30576 
diff
changeset
 | 
306  | 
|
| 
 
4169ddbfe1f9
command 'use', 'ML': apply ML environment to theory and target as well;
 
wenzelm 
parents: 
30576 
diff
changeset
 | 
307  | 
val _ =  | 
| 48642 | 308  | 
  Outer_Syntax.command @{command_spec "ML_prf"} "ML text within proof"
 | 
| 
55828
 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 
wenzelm 
parents: 
55795 
diff
changeset
 | 
309  | 
(Parse.ML_source >> (fn source =>  | 
| 28269 | 310  | 
Toplevel.proof (Proof.map_context (Context.proof_map  | 
| 
56304
 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 
wenzelm 
parents: 
56275 
diff
changeset
 | 
311  | 
(ML_Context.exec (fn () =>  | 
| 
 
40274e4f5ebf
redirect ML_Compiler reports more directly: only the (big) parse tree report is deferred via Execution.print (NB: this does not work for asynchronous "diag" commands);
 
wenzelm 
parents: 
56275 
diff
changeset
 | 
312  | 
ML_Context.eval_source (ML_Compiler.verbose true ML_Compiler.flags) source))) #>  | 
| 
55828
 
42ac3cfb89f6
clarified language markup: added "delimited" property;
 
wenzelm 
parents: 
55795 
diff
changeset
 | 
313  | 
Proof.propagate_ml_env)));  | 
| 28269 | 314  | 
|
315  | 
val _ =  | 
|
| 48642 | 316  | 
  Outer_Syntax.command @{command_spec "ML_val"} "diagnostic ML text"
 | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
317  | 
(Parse.ML_source >> Isar_Cmd.ml_diag true);  | 
| 26396 | 318  | 
|
319  | 
val _ =  | 
|
| 48642 | 320  | 
  Outer_Syntax.command @{command_spec "ML_command"} "diagnostic ML text (silent)"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
321  | 
(Parse.ML_source >> Isar_Cmd.ml_diag false);  | 
| 5832 | 322  | 
|
| 24868 | 323  | 
val _ =  | 
| 48642 | 324  | 
  Outer_Syntax.command @{command_spec "setup"} "ML theory setup"
 | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
325  | 
(Parse.ML_source >> (Toplevel.theory o Isar_Cmd.global_setup));  | 
| 30461 | 326  | 
|
327  | 
val _ =  | 
|
| 48642 | 328  | 
  Outer_Syntax.local_theory @{command_spec "local_setup"} "ML local theory setup"
 | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
329  | 
(Parse.ML_source >> Isar_Cmd.local_setup);  | 
| 5832 | 330  | 
|
| 24868 | 331  | 
val _ =  | 
| 48642 | 332  | 
  Outer_Syntax.command @{command_spec "attribute_setup"} "define attribute in ML"
 | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
333  | 
(Parse.position Parse.name --  | 
| 48643 | 334  | 
        Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
 | 
| 36950 | 335  | 
>> (fn (name, (txt, cmt)) => Toplevel.theory (Attrib.attribute_setup name txt cmt)));  | 
| 30526 | 336  | 
|
337  | 
val _ =  | 
|
| 48642 | 338  | 
  Outer_Syntax.command @{command_spec "method_setup"} "define proof method in ML"
 | 
| 
42813
 
6c841fa92fa2
optional description for 'attribute_setup' and 'method_setup';
 
wenzelm 
parents: 
42496 
diff
changeset
 | 
339  | 
(Parse.position Parse.name --  | 
| 48643 | 340  | 
        Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
 | 
| 36950 | 341  | 
>> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt)));  | 
| 9197 | 342  | 
|
| 24868 | 343  | 
val _ =  | 
| 48642 | 344  | 
  Outer_Syntax.local_theory @{command_spec "declaration"} "generic ML declaration"
 | 
| 40784 | 345  | 
(Parse.opt_keyword "pervasive" -- Parse.ML_source  | 
346  | 
      >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = false, pervasive = pervasive} txt));
 | 
|
347  | 
||
348  | 
val _ =  | 
|
| 48642 | 349  | 
  Outer_Syntax.local_theory @{command_spec "syntax_declaration"} "generic ML syntax declaration"
 | 
| 40784 | 350  | 
(Parse.opt_keyword "pervasive" -- Parse.ML_source  | 
351  | 
      >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = true, pervasive = pervasive} txt));
 | 
|
| 22088 | 352  | 
|
| 24868 | 353  | 
val _ =  | 
| 48642 | 354  | 
  Outer_Syntax.local_theory @{command_spec "simproc_setup"} "define simproc in ML"
 | 
| 42464 | 355  | 
(Parse.position Parse.name --  | 
| 48643 | 356  | 
      (@{keyword "("} |-- Parse.enum1 "|" Parse.term --| @{keyword ")"} --| @{keyword "="}) --
 | 
357  | 
      Parse.ML_source -- Scan.optional (@{keyword "identifier"} |-- Scan.repeat1 Parse.xname) []
 | 
|
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
358  | 
>> (fn (((a, b), c), d) => Isar_Cmd.simproc_setup a b c d));  | 
| 22202 | 359  | 
|
| 5832 | 360  | 
|
361  | 
(* translation functions *)  | 
|
362  | 
||
| 24868 | 363  | 
val _ =  | 
| 48642 | 364  | 
  Outer_Syntax.command @{command_spec "parse_ast_translation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
365  | 
"install parse ast translation functions"  | 
| 52143 | 366  | 
(Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_ast_translation));  | 
| 5832 | 367  | 
|
| 24868 | 368  | 
val _ =  | 
| 48642 | 369  | 
  Outer_Syntax.command @{command_spec "parse_translation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
370  | 
"install parse translation functions"  | 
| 52143 | 371  | 
(Parse.ML_source >> (Toplevel.theory o Isar_Cmd.parse_translation));  | 
| 5832 | 372  | 
|
| 24868 | 373  | 
val _ =  | 
| 48642 | 374  | 
  Outer_Syntax.command @{command_spec "print_translation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
375  | 
"install print translation functions"  | 
| 52143 | 376  | 
(Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_translation));  | 
| 5832 | 377  | 
|
| 24868 | 378  | 
val _ =  | 
| 48642 | 379  | 
  Outer_Syntax.command @{command_spec "typed_print_translation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
380  | 
"install typed print translation functions"  | 
| 52143 | 381  | 
(Parse.ML_source >> (Toplevel.theory o Isar_Cmd.typed_print_translation));  | 
| 5832 | 382  | 
|
| 24868 | 383  | 
val _ =  | 
| 48642 | 384  | 
  Outer_Syntax.command @{command_spec "print_ast_translation"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
385  | 
"install print ast translation functions"  | 
| 52143 | 386  | 
(Parse.ML_source >> (Toplevel.theory o Isar_Cmd.print_ast_translation));  | 
| 5832 | 387  | 
|
388  | 
||
389  | 
(* oracles *)  | 
|
390  | 
||
| 24868 | 391  | 
val _ =  | 
| 48642 | 392  | 
  Outer_Syntax.command @{command_spec "oracle"} "declare oracle"
 | 
| 48643 | 393  | 
    (Parse.position Parse.name -- (@{keyword "="} |-- Parse.ML_source) >>
 | 
| 
37216
 
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
 
wenzelm 
parents: 
37198 
diff
changeset
 | 
394  | 
(fn (x, y) => Toplevel.theory (Isar_Cmd.oracle x y)));  | 
| 5832 | 395  | 
|
396  | 
||
| 47057 | 397  | 
(* bundled declarations *)  | 
398  | 
||
399  | 
val _ =  | 
|
| 48642 | 400  | 
  Outer_Syntax.local_theory @{command_spec "bundle"} "define bundle of declarations"
 | 
| 48643 | 401  | 
    ((Parse.binding --| @{keyword "="}) -- Parse_Spec.xthms1 -- Parse.for_fixes
 | 
| 47057 | 402  | 
>> (uncurry Bundle.bundle_cmd));  | 
403  | 
||
404  | 
val _ =  | 
|
| 48642 | 405  | 
  Outer_Syntax.command @{command_spec "include"}
 | 
| 47057 | 406  | 
"include declarations from bundle in proof body"  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
407  | 
(Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.proof o Bundle.include_cmd));  | 
| 47057 | 408  | 
|
409  | 
val _ =  | 
|
| 48642 | 410  | 
  Outer_Syntax.command @{command_spec "including"}
 | 
| 47057 | 411  | 
"include declarations from bundle in goal refinement"  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
412  | 
(Scan.repeat1 (Parse.position Parse.xname) >> (Toplevel.proof o Bundle.including_cmd));  | 
| 
47066
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47057 
diff
changeset
 | 
413  | 
|
| 
 
8a6124d09ff5
basic support for nested contexts including bundles;
 
wenzelm 
parents: 
47057 
diff
changeset
 | 
414  | 
val _ =  | 
| 48642 | 415  | 
  Outer_Syntax.improper_command @{command_spec "print_bundles"}
 | 
416  | 
"print bundles of declarations"  | 
|
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
417  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 47057 | 418  | 
Toplevel.keep (Bundle.print_bundles o Toplevel.context_of)));  | 
419  | 
||
420  | 
||
| 47069 | 421  | 
(* local theories *)  | 
422  | 
||
423  | 
val _ =  | 
|
| 48642 | 424  | 
  Outer_Syntax.command @{command_spec "context"} "begin local theory context"
 | 
| 
47253
 
a00c5c88d8f3
more general context command with auxiliary fixes/assumes etc.;
 
wenzelm 
parents: 
47069 
diff
changeset
 | 
425  | 
((Parse.position Parse.xname >> (fn name =>  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
426  | 
Toplevel.begin_local_theory true (Named_Target.context_cmd name)) ||  | 
| 
47253
 
a00c5c88d8f3
more general context command with auxiliary fixes/assumes etc.;
 
wenzelm 
parents: 
47069 
diff
changeset
 | 
427  | 
Scan.optional Parse_Spec.includes [] -- Scan.repeat Parse_Spec.context_element  | 
| 
 
a00c5c88d8f3
more general context command with auxiliary fixes/assumes etc.;
 
wenzelm 
parents: 
47069 
diff
changeset
 | 
428  | 
>> (fn (incls, elems) => Toplevel.open_target (Bundle.context_cmd incls elems)))  | 
| 47069 | 429  | 
--| Parse.begin);  | 
430  | 
||
431  | 
||
| 12061 | 432  | 
(* locales *)  | 
433  | 
||
| 12758 | 434  | 
val locale_val =  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
435  | 
Parse_Spec.locale_expression false --  | 
| 48643 | 436  | 
    Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
 | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
437  | 
Scan.repeat1 Parse_Spec.context_element >> pair ([], []);  | 
| 12061 | 438  | 
|
| 24868 | 439  | 
val _ =  | 
| 48642 | 440  | 
  Outer_Syntax.command @{command_spec "locale"} "define named proof context"
 | 
| 36950 | 441  | 
(Parse.binding --  | 
| 48643 | 442  | 
      Scan.optional (@{keyword "="} |-- Parse.!!! locale_val) (([], []), []) -- Parse.opt_begin
 | 
| 27681 | 443  | 
>> (fn ((name, (expr, elems)), begin) =>  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
444  | 
Toplevel.begin_local_theory begin  | 
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41536 
diff
changeset
 | 
445  | 
(Expression.add_locale_cmd I name Binding.empty expr elems #> snd)));  | 
| 
28085
 
914183e229e9
Interpretation commands no longer accept interpretation attributes.
 
ballarin 
parents: 
28084 
diff
changeset
 | 
446  | 
|
| 53988 | 447  | 
fun interpretation_args mandatory =  | 
| 41270 | 448  | 
Parse.!!! (Parse_Spec.locale_expression mandatory) --  | 
449  | 
Scan.optional  | 
|
450  | 
(Parse.where_ |-- Parse.and_list1 (Parse_Spec.opt_thm_name ":" -- Parse.prop)) [];  | 
|
451  | 
||
| 24868 | 452  | 
val _ =  | 
| 48642 | 453  | 
  Outer_Syntax.command @{command_spec "sublocale"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
454  | 
"prove sublocale relation between a locale and a locale expression"  | 
| 
51737
 
718866dda2fa
target-sensitive user-level commands interpretation and sublocale
 
haftmann 
parents: 
51717 
diff
changeset
 | 
455  | 
    ((Parse.position Parse.xname --| (@{keyword "\<subseteq>"} || @{keyword "<"}) --
 | 
| 53988 | 456  | 
interpretation_args false >> (fn (loc, (expr, equations)) =>  | 
457  | 
Toplevel.theory_to_proof (Expression.sublocale_global_cmd I loc expr equations)))  | 
|
458  | 
|| interpretation_args false >> (fn (expr, equations) =>  | 
|
459  | 
Toplevel.local_theory_to_proof NONE (Expression.sublocale_cmd expr equations)));  | 
|
| 28895 | 460  | 
|
461  | 
val _ =  | 
|
| 48642 | 462  | 
  Outer_Syntax.command @{command_spec "interpretation"}
 | 
| 
51737
 
718866dda2fa
target-sensitive user-level commands interpretation and sublocale
 
haftmann 
parents: 
51717 
diff
changeset
 | 
463  | 
"prove interpretation of locale expression in local theory"  | 
| 53988 | 464  | 
(interpretation_args true >> (fn (expr, equations) =>  | 
465  | 
Toplevel.local_theory_to_proof NONE (Expression.interpretation_cmd expr equations)));  | 
|
| 29223 | 466  | 
|
467  | 
val _ =  | 
|
| 48642 | 468  | 
  Outer_Syntax.command @{command_spec "interpret"}
 | 
| 29223 | 469  | 
"prove interpretation of locale expression in proof context"  | 
| 53988 | 470  | 
(interpretation_args true >> (fn (expr, equations) =>  | 
471  | 
Toplevel.proof' (Expression.interpret_cmd expr equations)));  | 
|
| 29223 | 472  | 
|
| 15703 | 473  | 
|
| 22299 | 474  | 
(* classes *)  | 
475  | 
||
| 24868 | 476  | 
val class_val =  | 
| 
46922
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
45837 
diff
changeset
 | 
477  | 
Parse_Spec.class_expression --  | 
| 48643 | 478  | 
    Scan.optional (@{keyword "+"} |-- Parse.!!! (Scan.repeat1 Parse_Spec.context_element)) [] ||
 | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
479  | 
Scan.repeat1 Parse_Spec.context_element >> pair [];  | 
| 22299 | 480  | 
|
| 24868 | 481  | 
val _ =  | 
| 48642 | 482  | 
  Outer_Syntax.command @{command_spec "class"} "define type class"
 | 
| 48643 | 483  | 
   (Parse.binding -- Scan.optional (@{keyword "="} |-- class_val) ([], []) -- Parse.opt_begin
 | 
| 26516 | 484  | 
>> (fn ((name, (supclasses, elems)), begin) =>  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
485  | 
Toplevel.begin_local_theory begin  | 
| 
41585
 
45d7da4e4ccf
added before_exit continuation for named targets (locale, class etc.), e.g. for final check/cleanup as in VC management;
 
wenzelm 
parents: 
41536 
diff
changeset
 | 
486  | 
(Class_Declaration.class_cmd I name supclasses elems #> snd)));  | 
| 22299 | 487  | 
|
| 24868 | 488  | 
val _ =  | 
| 48642 | 489  | 
  Outer_Syntax.local_theory_to_proof @{command_spec "subclass"} "prove a subclass relation"
 | 
| 
46922
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
45837 
diff
changeset
 | 
490  | 
(Parse.class >> Class_Declaration.subclass_cmd I);  | 
| 
24914
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
491  | 
|
| 
 
95cda5dd58d5
added proper subclass concept; improved class target
 
haftmann 
parents: 
24871 
diff
changeset
 | 
492  | 
val _ =  | 
| 48642 | 493  | 
  Outer_Syntax.command @{command_spec "instantiation"} "instantiate and prove type arity"
 | 
| 36950 | 494  | 
(Parse.multi_arity --| Parse.begin  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
495  | 
>> (fn arities => Toplevel.begin_local_theory true (Class.instantiation_cmd arities)));  | 
| 24589 | 496  | 
|
| 25485 | 497  | 
val _ =  | 
| 48642 | 498  | 
  Outer_Syntax.command @{command_spec "instance"} "prove type arity or subclass relation"
 | 
| 
46922
 
3717f3878714
source positions for locale and class expressions;
 
wenzelm 
parents: 
45837 
diff
changeset
 | 
499  | 
((Parse.class --  | 
| 48643 | 500  | 
    ((@{keyword "\<subseteq>"} || @{keyword "<"}) |-- Parse.!!! Parse.class) >> Class.classrel_cmd ||
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
501  | 
Parse.multi_arity >> Class.instance_arity_cmd) >> Toplevel.theory_to_proof ||  | 
| 
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
502  | 
Scan.succeed (Toplevel.local_theory_to_proof NONE (Class.instantiation_instance I)));  | 
| 22299 | 503  | 
|
504  | 
||
| 25519 | 505  | 
(* arbitrary overloading *)  | 
506  | 
||
507  | 
val _ =  | 
|
| 48642 | 508  | 
  Outer_Syntax.command @{command_spec "overloading"} "overloaded definitions"
 | 
| 48643 | 509  | 
   (Scan.repeat1 (Parse.name --| (@{keyword "\<equiv>"} || @{keyword "=="}) -- Parse.term --
 | 
510  | 
      Scan.optional (@{keyword "("} |-- (@{keyword "unchecked"} >> K false) --| @{keyword ")"}) true
 | 
|
| 36950 | 511  | 
>> Parse.triple1) --| Parse.begin  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
512  | 
>> (fn operations => Toplevel.begin_local_theory true (Overloading.overloading_cmd operations)));  | 
| 25519 | 513  | 
|
514  | 
||
| 22866 | 515  | 
(* code generation *)  | 
516  | 
||
| 24868 | 517  | 
val _ =  | 
| 48642 | 518  | 
  Outer_Syntax.command @{command_spec "code_datatype"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
519  | 
"define set of code datatype constructors"  | 
| 36950 | 520  | 
(Scan.repeat1 Parse.term >> (Toplevel.theory o Code.add_datatype_cmd));  | 
| 22866 | 521  | 
|
522  | 
||
| 5832 | 523  | 
|
524  | 
(** proof commands **)  | 
|
525  | 
||
526  | 
(* statements *)  | 
|
527  | 
||
| 53988 | 528  | 
fun theorem spec schematic kind =  | 
| 48645 | 529  | 
Outer_Syntax.local_theory_to_proof' spec  | 
| 
36317
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
530  | 
    ("state " ^ (if schematic then "schematic " ^ kind else kind))
 | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
531  | 
(Scan.optional (Parse_Spec.opt_thm_name ":" --|  | 
| 
47067
 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 
wenzelm 
parents: 
47066 
diff
changeset
 | 
532  | 
Scan.ahead (Parse_Spec.includes >> K "" ||  | 
| 
 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 
wenzelm 
parents: 
47066 
diff
changeset
 | 
533  | 
Parse_Spec.locale_keyword || Parse_Spec.statement_keyword)) Attrib.empty_binding --  | 
| 
 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 
wenzelm 
parents: 
47066 
diff
changeset
 | 
534  | 
Scan.optional Parse_Spec.includes [] --  | 
| 
 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 
wenzelm 
parents: 
47066 
diff
changeset
 | 
535  | 
Parse_Spec.general_statement >> (fn ((a, includes), (elems, concl)) =>  | 
| 
36317
 
506d732cb522
explicit 'schematic_theorem' etc. for schematic theorem statements;
 
wenzelm 
parents: 
36178 
diff
changeset
 | 
536  | 
((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd)  | 
| 
47067
 
4ef29b0c568f
optional 'includes' element for long theorem statements;
 
wenzelm 
parents: 
47066 
diff
changeset
 | 
537  | 
kind NONE (K I) a includes elems concl)));  | 
| 5832 | 538  | 
|
| 53988 | 539  | 
val _ = theorem @{command_spec "theorem"} false Thm.theoremK;
 | 
540  | 
val _ = theorem @{command_spec "lemma"} false Thm.lemmaK;
 | 
|
541  | 
val _ = theorem @{command_spec "corollary"} false Thm.corollaryK;
 | 
|
542  | 
val _ = theorem @{command_spec "schematic_theorem"} true Thm.theoremK;
 | 
|
543  | 
val _ = theorem @{command_spec "schematic_lemma"} true Thm.lemmaK;
 | 
|
544  | 
val _ = theorem @{command_spec "schematic_corollary"} true Thm.corollaryK;
 | 
|
| 5832 | 545  | 
|
| 24868 | 546  | 
val _ =  | 
| 48642 | 547  | 
  Outer_Syntax.local_theory_to_proof @{command_spec "notepad"} "begin proof context"
 | 
| 40960 | 548  | 
(Parse.begin >> K Proof.begin_notepad);  | 
549  | 
||
550  | 
val _ =  | 
|
| 48642 | 551  | 
  Outer_Syntax.command @{command_spec "have"} "state local goal"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
552  | 
(Parse_Spec.statement >> (Toplevel.proof' o Isar_Cmd.have));  | 
| 17353 | 553  | 
|
| 24868 | 554  | 
val _ =  | 
| 50214 | 555  | 
  Outer_Syntax.command @{command_spec "hence"} "old-style alias of \"then have\""
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
556  | 
(Parse_Spec.statement >> (Toplevel.proof' o Isar_Cmd.hence));  | 
| 17353 | 557  | 
|
| 24868 | 558  | 
val _ =  | 
| 48642 | 559  | 
  Outer_Syntax.command @{command_spec "show"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
560  | 
"state local goal, solving current obligation"  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
561  | 
(Parse_Spec.statement >> (Toplevel.proof' o Isar_Cmd.show));  | 
| 5832 | 562  | 
|
| 24868 | 563  | 
val _ =  | 
| 50214 | 564  | 
  Outer_Syntax.command @{command_spec "thus"} "old-style alias of  \"then show\""
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
565  | 
(Parse_Spec.statement >> (Toplevel.proof' o Isar_Cmd.thus));  | 
| 6501 | 566  | 
|
| 5832 | 567  | 
|
| 5914 | 568  | 
(* facts *)  | 
| 5832 | 569  | 
|
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
570  | 
val facts = Parse.and_list1 Parse_Spec.xthms1;  | 
| 9197 | 571  | 
|
| 24868 | 572  | 
val _ =  | 
| 48642 | 573  | 
  Outer_Syntax.command @{command_spec "then"} "forward chaining"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
574  | 
(Scan.succeed (Toplevel.proof Proof.chain));  | 
| 5832 | 575  | 
|
| 24868 | 576  | 
val _ =  | 
| 48642 | 577  | 
  Outer_Syntax.command @{command_spec "from"} "forward chaining from given facts"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
578  | 
(facts >> (Toplevel.proof o Proof.from_thmss_cmd));  | 
| 5914 | 579  | 
|
| 24868 | 580  | 
val _ =  | 
| 48642 | 581  | 
  Outer_Syntax.command @{command_spec "with"} "forward chaining from given and current facts"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
582  | 
(facts >> (Toplevel.proof o Proof.with_thmss_cmd));  | 
| 6878 | 583  | 
|
| 24868 | 584  | 
val _ =  | 
| 48642 | 585  | 
  Outer_Syntax.command @{command_spec "note"} "define facts"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
586  | 
(Parse_Spec.name_facts >> (Toplevel.proof o Proof.note_thmss_cmd));  | 
| 5832 | 587  | 
|
| 24868 | 588  | 
val _ =  | 
| 48642 | 589  | 
  Outer_Syntax.command @{command_spec "using"} "augment goal facts"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
590  | 
(facts >> (Toplevel.proof o Proof.using_cmd));  | 
| 18544 | 591  | 
|
| 24868 | 592  | 
val _ =  | 
| 48642 | 593  | 
  Outer_Syntax.command @{command_spec "unfolding"} "unfold definitions in goal and facts"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
594  | 
(facts >> (Toplevel.proof o Proof.unfolding_cmd));  | 
| 12926 | 595  | 
|
| 5832 | 596  | 
|
597  | 
(* proof context *)  | 
|
598  | 
||
| 24868 | 599  | 
val _ =  | 
| 48642 | 600  | 
  Outer_Syntax.command @{command_spec "fix"} "fix local variables (Skolem constants)"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
601  | 
(Parse.fixes >> (Toplevel.proof o Proof.fix_cmd));  | 
| 11890 | 602  | 
|
| 24868 | 603  | 
val _ =  | 
| 48642 | 604  | 
  Outer_Syntax.command @{command_spec "assume"} "assume propositions"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
605  | 
(Parse_Spec.statement >> (Toplevel.proof o Proof.assume_cmd));  | 
| 5832 | 606  | 
|
| 24868 | 607  | 
val _ =  | 
| 48642 | 608  | 
  Outer_Syntax.command @{command_spec "presume"} "assume propositions, to be established later"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
609  | 
(Parse_Spec.statement >> (Toplevel.proof o Proof.presume_cmd));  | 
| 6850 | 610  | 
|
| 24868 | 611  | 
val _ =  | 
| 50214 | 612  | 
  Outer_Syntax.command @{command_spec "def"} "local definition (non-polymorphic)"
 | 
| 36950 | 613  | 
(Parse.and_list1  | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
614  | 
(Parse_Spec.opt_thm_name ":" --  | 
| 36950 | 615  | 
((Parse.binding -- Parse.opt_mixfix) --  | 
| 48643 | 616  | 
          ((@{keyword "\<equiv>"} || @{keyword "=="}) |-- Parse.!!! Parse.termp)))
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
617  | 
>> (Toplevel.proof o Proof.def_cmd));  | 
| 6953 | 618  | 
|
| 24868 | 619  | 
val _ =  | 
| 48642 | 620  | 
  Outer_Syntax.command @{command_spec "obtain"} "generalized elimination"
 | 
| 
36952
 
338c3f8229e4
renamed structure SpecParse to Parse_Spec, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36951 
diff
changeset
 | 
621  | 
(Parse.parname -- Scan.optional (Parse.fixes --| Parse.where_) [] -- Parse_Spec.statement  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
622  | 
>> (fn ((x, y), z) => Toplevel.proof' (Obtain.obtain_cmd x y z)));  | 
| 5832 | 623  | 
|
| 24868 | 624  | 
val _ =  | 
| 48642 | 625  | 
  Outer_Syntax.command @{command_spec "guess"} "wild guessing (unstructured)"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
626  | 
(Scan.optional Parse.fixes [] >> (Toplevel.proof' o Obtain.guess_cmd));  | 
| 17854 | 627  | 
|
| 24868 | 628  | 
val _ =  | 
| 48642 | 629  | 
  Outer_Syntax.command @{command_spec "let"} "bind text variables"
 | 
| 48643 | 630  | 
    (Parse.and_list1 (Parse.and_list1 Parse.term -- (@{keyword "="} |-- Parse.term))
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
631  | 
>> (Toplevel.proof o Proof.let_bind_cmd));  | 
| 5832 | 632  | 
|
| 
36505
 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 
wenzelm 
parents: 
36452 
diff
changeset
 | 
633  | 
val _ =  | 
| 48642 | 634  | 
  Outer_Syntax.command @{command_spec "write"} "add concrete syntax for constants / fixed variables"
 | 
| 
51654
 
8450b944e58a
just one syntax category "mixfix" -- check structure annotation semantically;
 
wenzelm 
parents: 
51627 
diff
changeset
 | 
635  | 
(opt_mode -- Parse.and_list1 (Parse.const -- Parse.mixfix)  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
636  | 
>> (fn (mode, args) => Toplevel.proof (Proof.write_cmd mode args)));  | 
| 
36505
 
79c1d2bbe5a9
ProofContext.read_const: allow for type constraint (for fixed variable);
 
wenzelm 
parents: 
36452 
diff
changeset
 | 
637  | 
|
| 24868 | 638  | 
val _ =  | 
| 48642 | 639  | 
  Outer_Syntax.command @{command_spec "case"} "invoke local context"
 | 
| 
53377
 
21693b7c8fbf
more liberal 'case' syntax: allow parentheses without arguments;
 
wenzelm 
parents: 
52549 
diff
changeset
 | 
640  | 
    ((@{keyword "("} |--
 | 
| 
53378
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
53377 
diff
changeset
 | 
641  | 
Parse.!!! (Parse.position Parse.xname -- Scan.repeat (Parse.maybe Parse.binding)  | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
53377 
diff
changeset
 | 
642  | 
        --| @{keyword ")"}) ||
 | 
| 
 
07990ba8c0ea
cases: more position information and PIDE markup;
 
wenzelm 
parents: 
53377 
diff
changeset
 | 
643  | 
Parse.position Parse.xname >> rpair []) -- Parse_Spec.opt_attribs >> (fn ((c, xs), atts) =>  | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
644  | 
Toplevel.proof (Proof.invoke_case_cmd (c, xs, atts))));  | 
| 8370 | 645  | 
|
| 5832 | 646  | 
|
647  | 
(* proof structure *)  | 
|
648  | 
||
| 24868 | 649  | 
val _ =  | 
| 48642 | 650  | 
  Outer_Syntax.command @{command_spec "{"} "begin explicit proof block"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
651  | 
(Scan.succeed (Toplevel.proof Proof.begin_block));  | 
| 5832 | 652  | 
|
| 24868 | 653  | 
val _ =  | 
| 48642 | 654  | 
  Outer_Syntax.command @{command_spec "}"} "end explicit proof block"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
655  | 
(Scan.succeed (Toplevel.proof Proof.end_block));  | 
| 6687 | 656  | 
|
| 24868 | 657  | 
val _ =  | 
| 48642 | 658  | 
  Outer_Syntax.command @{command_spec "next"} "enter next proof block"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
659  | 
(Scan.succeed (Toplevel.proof Proof.next_block));  | 
| 5832 | 660  | 
|
661  | 
||
662  | 
(* end proof *)  | 
|
663  | 
||
| 24868 | 664  | 
val _ =  | 
| 50214 | 665  | 
  Outer_Syntax.command @{command_spec "qed"} "conclude proof"
 | 
| 56869 | 666  | 
(Scan.option Method.parse >> (fn m =>  | 
667  | 
(Option.map Method.report m;  | 
|
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
668  | 
Isar_Cmd.qed m)));  | 
| 5832 | 669  | 
|
| 24868 | 670  | 
val _ =  | 
| 48642 | 671  | 
  Outer_Syntax.command @{command_spec "by"} "terminal backward proof"
 | 
| 55795 | 672  | 
(Method.parse -- Scan.option Method.parse >> (fn (m1, m2) =>  | 
| 56869 | 673  | 
(Method.report m1;  | 
674  | 
Option.map Method.report m2;  | 
|
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
675  | 
Isar_Cmd.terminal_proof (m1, m2))));  | 
| 6404 | 676  | 
|
| 24868 | 677  | 
val _ =  | 
| 48642 | 678  | 
  Outer_Syntax.command @{command_spec ".."} "default proof"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
679  | 
(Scan.succeed Isar_Cmd.default_proof);  | 
| 8966 | 680  | 
|
| 24868 | 681  | 
val _ =  | 
| 48642 | 682  | 
  Outer_Syntax.command @{command_spec "."} "immediate proof"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
683  | 
(Scan.succeed Isar_Cmd.immediate_proof);  | 
| 6404 | 684  | 
|
| 24868 | 685  | 
val _ =  | 
| 48642 | 686  | 
  Outer_Syntax.command @{command_spec "done"} "done proof"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
687  | 
(Scan.succeed Isar_Cmd.done_proof);  | 
| 5832 | 688  | 
|
| 24868 | 689  | 
val _ =  | 
| 48642 | 690  | 
  Outer_Syntax.command @{command_spec "sorry"} "skip proof (quick-and-dirty mode only!)"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
691  | 
(Scan.succeed Isar_Cmd.skip_proof);  | 
| 
6888
 
d0c68ebdabc5
skip_proof feature 'sorry' (for quick_and_dirty mode only);
 
wenzelm 
parents: 
6886 
diff
changeset
 | 
692  | 
|
| 24868 | 693  | 
val _ =  | 
| 48642 | 694  | 
  Outer_Syntax.command @{command_spec "oops"} "forget proof"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
695  | 
(Scan.succeed Toplevel.forget_proof);  | 
| 8210 | 696  | 
|
| 5832 | 697  | 
|
698  | 
(* proof steps *)  | 
|
699  | 
||
| 24868 | 700  | 
val _ =  | 
| 48642 | 701  | 
  Outer_Syntax.command @{command_spec "defer"} "shuffle internal proof state"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
702  | 
(Scan.optional Parse.nat 1 >> (Toplevel.proof o Proof.defer));  | 
| 8165 | 703  | 
|
| 24868 | 704  | 
val _ =  | 
| 48642 | 705  | 
  Outer_Syntax.command @{command_spec "prefer"} "shuffle internal proof state"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
706  | 
(Parse.nat >> (Toplevel.proof o Proof.prefer));  | 
| 8165 | 707  | 
|
| 24868 | 708  | 
val _ =  | 
| 48642 | 709  | 
  Outer_Syntax.command @{command_spec "apply"} "initial refinement step (unstructured)"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
710  | 
(Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply_results m))));  | 
| 5832 | 711  | 
|
| 24868 | 712  | 
val _ =  | 
| 50214 | 713  | 
  Outer_Syntax.command @{command_spec "apply_end"} "terminal refinement step (unstructured)"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
714  | 
(Method.parse >> (fn m => (Method.report m; Toplevel.proofs (Proof.apply_end_results m))));  | 
| 5832 | 715  | 
|
| 24868 | 716  | 
val _ =  | 
| 50214 | 717  | 
  Outer_Syntax.command @{command_spec "proof"} "backward proof step"
 | 
| 55795 | 718  | 
(Scan.option Method.parse >> (fn m =>  | 
719  | 
(Option.map Method.report m;  | 
|
720  | 
Toplevel.actual_proof (Proof_Node.applys (Proof.proof_results m)) o  | 
|
721  | 
Toplevel.skip_proof (fn i => i + 1))));  | 
|
| 5832 | 722  | 
|
723  | 
||
| 6742 | 724  | 
(* proof navigation *)  | 
| 5944 | 725  | 
|
| 
56065
 
600781e03bf6
more explicit markup and explanation of the improper status of 'back', following the AFP style-guide;
 
wenzelm 
parents: 
56006 
diff
changeset
 | 
726  | 
fun report_back () =  | 
| 
56333
 
38f1422ef473
support bulk messages consisting of small string segments, which are more healthy to the Poly/ML RTS and might prevent spurious GC crashes such as MTGCProcessMarkPointers::ScanAddressesInObject;
 
wenzelm 
parents: 
56304 
diff
changeset
 | 
727  | 
Output.report [Markup.markup Markup.bad "Explicit backtracking"];  | 
| 
56065
 
600781e03bf6
more explicit markup and explanation of the improper status of 'back', following the AFP style-guide;
 
wenzelm 
parents: 
56006 
diff
changeset
 | 
728  | 
|
| 24868 | 729  | 
val _ =  | 
| 
56065
 
600781e03bf6
more explicit markup and explanation of the improper status of 'back', following the AFP style-guide;
 
wenzelm 
parents: 
56006 
diff
changeset
 | 
730  | 
  Outer_Syntax.command @{command_spec "back"} "explicit backtracking of proof command"
 | 
| 
56895
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
731  | 
(Scan.succeed  | 
| 
 
f058120aaad4
discontinued Toplevel.print flag -- print uniformly according to Keyword.is_printed;
 
wenzelm 
parents: 
56893 
diff
changeset
 | 
732  | 
(Toplevel.actual_proof (fn prf => (report_back (); Proof_Node.back prf)) o  | 
| 
56065
 
600781e03bf6
more explicit markup and explanation of the improper status of 'back', following the AFP style-guide;
 
wenzelm 
parents: 
56006 
diff
changeset
 | 
733  | 
Toplevel.skip_proof (fn h => (report_back (); h))));  | 
| 6742 | 734  | 
|
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
735  | 
|
| 56869 | 736  | 
|
737  | 
(** nested commands **)  | 
|
| 25578 | 738  | 
|
| 29309 | 739  | 
val props_text =  | 
| 
43775
 
b361c7d184e7
added Parse.properties (again) -- allow empty list like Parse_Value.properties but unlike Parse.properties of ef86de9c98aa;
 
wenzelm 
parents: 
43227 
diff
changeset
 | 
740  | 
Scan.optional Parse.properties [] -- Parse.position Parse.string  | 
| 36950 | 741  | 
>> (fn (props, (str, pos)) =>  | 
742  | 
(Position.of_properties (Position.default_properties pos props), str));  | 
|
| 29309 | 743  | 
|
| 25578 | 744  | 
val _ =  | 
| 50214 | 745  | 
  Outer_Syntax.improper_command @{command_spec "Isabelle.command"} "evaluate nested Isabelle command"
 | 
| 29309 | 746  | 
(props_text :|-- (fn (pos, str) =>  | 
| 
36953
 
2af1ad9aa1a3
renamed structure OuterSyntax to Outer_Syntax, keeping the old name as alias for some time;
 
wenzelm 
parents: 
36952 
diff
changeset
 | 
747  | 
(case Outer_Syntax.parse pos str of  | 
| 
27838
 
0340fd7cccc3
Isabelle.command: inline former OuterSyntax.prepare_command;
 
wenzelm 
parents: 
27831 
diff
changeset
 | 
748  | 
[tr] => Scan.succeed (K tr)  | 
| 
43947
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43775 
diff
changeset
 | 
749  | 
| _ => Scan.fail_with (K (fn () => "exactly one command expected")))  | 
| 
 
9b00f09f7721
defer evaluation of Scan.message, for improved performance in the frequent situation where failure is handled later (e.g. via ||);
 
wenzelm 
parents: 
43775 
diff
changeset
 | 
750  | 
handle ERROR msg => Scan.fail_with (K (fn () => msg))));  | 
| 25578 | 751  | 
|
752  | 
||
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
753  | 
|
| 5832 | 754  | 
(** diagnostic commands (for interactive mode only) **)  | 
755  | 
||
| 36950 | 756  | 
val opt_modes =  | 
| 48643 | 757  | 
  Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) [];
 | 
| 36950 | 758  | 
|
| 48643 | 759  | 
val opt_bang = Scan.optional (@{keyword "!"} >> K true) false;
 | 
| 8464 | 760  | 
|
| 53988 | 761  | 
val _ = (*Proof General legacy*)  | 
| 48642 | 762  | 
  Outer_Syntax.improper_command @{command_spec "pretty_setmargin"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
763  | 
"change default margin for pretty printing"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
764  | 
(Parse.nat >> (fn n => Toplevel.imperative (fn () => Pretty.margin_default := n)));  | 
| 7124 | 765  | 
|
| 24868 | 766  | 
val _ =  | 
| 50213 | 767  | 
  Outer_Syntax.improper_command @{command_spec "help"}
 | 
768  | 
"retrieve outer syntax commands according to name patterns"  | 
|
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
769  | 
(Scan.repeat Parse.name >>  | 
| 
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
770  | 
(fn pats => Toplevel.imperative (fn () => Outer_Syntax.help_outer_syntax pats)));  | 
| 21714 | 771  | 
|
| 24868 | 772  | 
val _ =  | 
| 48642 | 773  | 
  Outer_Syntax.improper_command @{command_spec "print_commands"} "print outer syntax commands"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
774  | 
(Scan.succeed (Toplevel.imperative Outer_Syntax.print_outer_syntax));  | 
| 5832 | 775  | 
|
| 24868 | 776  | 
val _ =  | 
| 52060 | 777  | 
  Outer_Syntax.improper_command @{command_spec "print_options"} "print configuration options"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
778  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 52060 | 779  | 
Toplevel.keep (Attrib.print_options o Toplevel.context_of)));  | 
| 23989 | 780  | 
|
| 24868 | 781  | 
val _ =  | 
| 56893 | 782  | 
  Outer_Syntax.improper_command @{command_spec "print_context"}
 | 
783  | 
"print context of local theory target"  | 
|
784  | 
(Scan.succeed (Toplevel.keep (Pretty.writeln_chunks o Toplevel.pretty_context)));  | 
|
| 7308 | 785  | 
|
| 24868 | 786  | 
val _ =  | 
| 48642 | 787  | 
  Outer_Syntax.improper_command @{command_spec "print_theory"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
788  | 
"print logical theory contents (verbose!)"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
789  | 
(opt_bang >> (fn b => Toplevel.unknown_theory o  | 
| 50737 | 790  | 
Toplevel.keep (Pretty.writeln o Proof_Display.pretty_full_theory b o Toplevel.theory_of)));  | 
| 5832 | 791  | 
|
| 24868 | 792  | 
val _ =  | 
| 48642 | 793  | 
  Outer_Syntax.improper_command @{command_spec "print_syntax"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
794  | 
"print inner syntax of context (verbose!)"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
795  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 50737 | 796  | 
Toplevel.keep (Proof_Context.print_syntax o Toplevel.context_of)));  | 
| 5832 | 797  | 
|
| 24868 | 798  | 
val _ =  | 
| 51585 | 799  | 
  Outer_Syntax.improper_command @{command_spec "print_defn_rules"}
 | 
800  | 
"print definitional rewrite rules of context"  | 
|
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
801  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 51585 | 802  | 
Toplevel.keep (Local_Defs.print_rules o Toplevel.context_of)));  | 
803  | 
||
804  | 
val _ =  | 
|
| 48642 | 805  | 
  Outer_Syntax.improper_command @{command_spec "print_abbrevs"}
 | 
| 51585 | 806  | 
"print constant abbreviations of context"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
807  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 50737 | 808  | 
Toplevel.keep (Proof_Context.print_abbrevs o Toplevel.context_of)));  | 
| 21726 | 809  | 
|
| 24868 | 810  | 
val _ =  | 
| 48642 | 811  | 
  Outer_Syntax.improper_command @{command_spec "print_theorems"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
812  | 
"print theorems of local theory or proof context"  | 
| 
56868
 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 
wenzelm 
parents: 
56867 
diff
changeset
 | 
813  | 
(opt_bang >> (fn b =>  | 
| 
 
b5fb264d53ba
clarified print operations for "terms" and "theorems";
 
wenzelm 
parents: 
56867 
diff
changeset
 | 
814  | 
Toplevel.unknown_context o Toplevel.keep (Pretty.writeln o Isar_Cmd.pretty_theorems b)));  | 
| 5881 | 815  | 
|
| 24868 | 816  | 
val _ =  | 
| 48642 | 817  | 
  Outer_Syntax.improper_command @{command_spec "print_locales"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
818  | 
"print locales of this theory"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
819  | 
(Scan.succeed (Toplevel.unknown_theory o  | 
| 50737 | 820  | 
Toplevel.keep (Locale.print_locales o Toplevel.theory_of)));  | 
| 12061 | 821  | 
|
| 24868 | 822  | 
val _ =  | 
| 48642 | 823  | 
  Outer_Syntax.improper_command @{command_spec "print_classes"}
 | 
824  | 
"print classes of this theory"  | 
|
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
825  | 
(Scan.succeed (Toplevel.unknown_theory o  | 
| 42359 | 826  | 
Toplevel.keep (Class.print_classes o Toplevel.context_of)));  | 
| 
22744
 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 
haftmann 
parents: 
22573 
diff
changeset
 | 
827  | 
|
| 24868 | 828  | 
val _ =  | 
| 48642 | 829  | 
  Outer_Syntax.improper_command @{command_spec "print_locale"}
 | 
830  | 
"print locale of this theory"  | 
|
| 50737 | 831  | 
(opt_bang -- Parse.position Parse.xname >> (fn (b, name) =>  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
832  | 
Toplevel.unknown_theory o  | 
| 50737 | 833  | 
Toplevel.keep (fn state => Locale.print_locale (Toplevel.theory_of state) b name)));  | 
| 12061 | 834  | 
|
| 24868 | 835  | 
val _ =  | 
| 48642 | 836  | 
  Outer_Syntax.improper_command @{command_spec "print_interps"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
837  | 
"print interpretations of locale for this theory or proof context"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
838  | 
(Parse.position Parse.xname >> (fn name =>  | 
| 
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
839  | 
Toplevel.unknown_context o  | 
| 50737 | 840  | 
Toplevel.keep (fn state => Locale.print_registrations (Toplevel.context_of state) name)));  | 
| 
32804
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
parents: 
31869 
diff
changeset
 | 
841  | 
|
| 
 
ca430e6aee1c
Propagation of mixins for interpretation; reactivated diagnostic command print_interps.
 
ballarin 
parents: 
31869 
diff
changeset
 | 
842  | 
val _ =  | 
| 48642 | 843  | 
  Outer_Syntax.improper_command @{command_spec "print_dependencies"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
844  | 
"print dependencies of locale expression"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
845  | 
(opt_bang -- Parse_Spec.locale_expression true >> (fn (b, expr) =>  | 
| 
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
846  | 
Toplevel.unknown_context o  | 
| 
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
847  | 
Toplevel.keep (fn state => Expression.print_dependencies (Toplevel.context_of state) b expr)));  | 
| 41435 | 848  | 
|
849  | 
val _ =  | 
|
| 48642 | 850  | 
  Outer_Syntax.improper_command @{command_spec "print_attributes"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
851  | 
"print attributes of this theory"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
852  | 
(Scan.succeed (Toplevel.unknown_theory o  | 
| 
55997
 
9dc5ce83202c
modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
853  | 
Toplevel.keep (Attrib.print_attributes o Toplevel.context_of)));  | 
| 5832 | 854  | 
|
| 24868 | 855  | 
val _ =  | 
| 48642 | 856  | 
  Outer_Syntax.improper_command @{command_spec "print_simpset"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
857  | 
"print context of Simplifier"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
858  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51685 
diff
changeset
 | 
859  | 
Toplevel.keep (Pretty.writeln o Simplifier.pretty_simpset o Toplevel.context_of)));  | 
| 16027 | 860  | 
|
| 24868 | 861  | 
val _ =  | 
| 48642 | 862  | 
  Outer_Syntax.improper_command @{command_spec "print_rules"} "print intro/elim rules"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
863  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 50737 | 864  | 
Toplevel.keep (Context_Rules.print_rules o Toplevel.context_of)));  | 
| 12383 | 865  | 
|
| 24868 | 866  | 
val _ =  | 
| 48642 | 867  | 
  Outer_Syntax.improper_command @{command_spec "print_methods"} "print methods of this theory"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
868  | 
(Scan.succeed (Toplevel.unknown_theory o  | 
| 
55742
 
a989bdaf8121
modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
 
wenzelm 
parents: 
55385 
diff
changeset
 | 
869  | 
Toplevel.keep (Method.print_methods o Toplevel.context_of)));  | 
| 5832 | 870  | 
|
| 24868 | 871  | 
val _ =  | 
| 
56069
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
872  | 
  Outer_Syntax.improper_command @{command_spec "print_antiquotations"}
 | 
| 
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
873  | 
"print document antiquotations"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
874  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 50737 | 875  | 
Toplevel.keep (Thy_Output.print_antiquotations o Toplevel.context_of)));  | 
| 9221 | 876  | 
|
| 24868 | 877  | 
val _ =  | 
| 
56069
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
878  | 
  Outer_Syntax.improper_command @{command_spec "print_ML_antiquotations"}
 | 
| 
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
879  | 
"print ML antiquotations"  | 
| 
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
880  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
881  | 
Toplevel.keep (ML_Context.print_antiquotations o Toplevel.context_of)));  | 
| 
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
882  | 
|
| 
 
451d5b73f8cf
simplified programming interface to define ML antiquotations -- NB: the transformed context ignores updates of the context parser;
 
wenzelm 
parents: 
56065 
diff
changeset
 | 
883  | 
val _ =  | 
| 48642 | 884  | 
  Outer_Syntax.improper_command @{command_spec "thy_deps"} "visualize theory dependencies"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
885  | 
(Scan.succeed Isar_Cmd.thy_deps);  | 
| 22485 | 886  | 
|
| 24868 | 887  | 
val _ =  | 
| 
49569
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
48997 
diff
changeset
 | 
888  | 
  Outer_Syntax.improper_command @{command_spec "locale_deps"} "visualize locale dependencies"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
889  | 
(Scan.succeed Isar_Cmd.locale_deps);  | 
| 
49569
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
48997 
diff
changeset
 | 
890  | 
|
| 
 
7b6aaf446496
tuned pretty_locale/print_locale, with more basic pretty_locale_deps based on that;
 
wenzelm 
parents: 
48997 
diff
changeset
 | 
891  | 
val _ =  | 
| 48642 | 892  | 
  Outer_Syntax.improper_command @{command_spec "class_deps"} "visualize class dependencies"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
893  | 
(Scan.succeed Isar_Cmd.class_deps);  | 
| 20574 | 894  | 
|
| 24868 | 895  | 
val _ =  | 
| 48642 | 896  | 
  Outer_Syntax.improper_command @{command_spec "thm_deps"} "visualize theorem dependencies"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
897  | 
(Parse_Spec.xthms1 >> Isar_Cmd.thm_deps);  | 
| 9454 | 898  | 
|
| 24868 | 899  | 
val _ =  | 
| 48642 | 900  | 
  Outer_Syntax.improper_command @{command_spec "print_binds"} "print term bindings of proof context"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
901  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 56867 | 902  | 
Toplevel.keep (Pretty.writeln_chunks o Proof_Context.pretty_binds o Toplevel.context_of)));  | 
| 5832 | 903  | 
|
| 24868 | 904  | 
val _ =  | 
| 48642 | 905  | 
  Outer_Syntax.improper_command @{command_spec "print_facts"} "print facts of proof context"
 | 
| 
56158
 
c2c6d560e7b2
more explicit treatment of verbose mode, which includes concealed entries;
 
wenzelm 
parents: 
56155 
diff
changeset
 | 
906  | 
(opt_bang >> (fn verbose => Toplevel.unknown_context o  | 
| 
 
c2c6d560e7b2
more explicit treatment of verbose mode, which includes concealed entries;
 
wenzelm 
parents: 
56155 
diff
changeset
 | 
907  | 
Toplevel.keep (fn st => Proof_Context.print_local_facts (Toplevel.context_of st) verbose)));  | 
| 5832 | 908  | 
|
| 24868 | 909  | 
val _ =  | 
| 48642 | 910  | 
  Outer_Syntax.improper_command @{command_spec "print_cases"} "print cases of proof context"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
911  | 
(Scan.succeed (Toplevel.unknown_context o  | 
| 56867 | 912  | 
Toplevel.keep (Pretty.writeln_chunks o Proof_Context.pretty_cases o Toplevel.context_of)));  | 
| 8370 | 913  | 
|
| 24868 | 914  | 
val _ =  | 
| 48642 | 915  | 
  Outer_Syntax.improper_command @{command_spec "print_statement"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
916  | 
"print theorems as long statements"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
917  | 
(opt_modes -- Parse_Spec.xthms1 >> Isar_Cmd.print_stmts);  | 
| 19269 | 918  | 
|
| 24868 | 919  | 
val _ =  | 
| 48642 | 920  | 
  Outer_Syntax.improper_command @{command_spec "thm"} "print theorems"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
921  | 
(opt_modes -- Parse_Spec.xthms1 >> Isar_Cmd.print_thms);  | 
| 5832 | 922  | 
|
| 24868 | 923  | 
val _ =  | 
| 48642 | 924  | 
  Outer_Syntax.improper_command @{command_spec "prf"} "print proof terms of theorems"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
925  | 
(opt_modes -- Scan.option Parse_Spec.xthms1 >> Isar_Cmd.print_prfs false);  | 
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
926  | 
|
| 24868 | 927  | 
val _ =  | 
| 48642 | 928  | 
  Outer_Syntax.improper_command @{command_spec "full_prf"} "print full proof terms of theorems"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
929  | 
(opt_modes -- Scan.option Parse_Spec.xthms1 >> Isar_Cmd.print_prfs true);  | 
| 
11524
 
197f2e14a714
Added functions for printing primitive proof terms.
 
berghofe 
parents: 
11101 
diff
changeset
 | 
930  | 
|
| 24868 | 931  | 
val _ =  | 
| 48642 | 932  | 
  Outer_Syntax.improper_command @{command_spec "prop"} "read and print proposition"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
933  | 
(opt_modes -- Parse.term >> Isar_Cmd.print_prop);  | 
| 5832 | 934  | 
|
| 24868 | 935  | 
val _ =  | 
| 48642 | 936  | 
  Outer_Syntax.improper_command @{command_spec "term"} "read and print term"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
937  | 
(opt_modes -- Parse.term >> Isar_Cmd.print_term);  | 
| 5832 | 938  | 
|
| 24868 | 939  | 
val _ =  | 
| 48642 | 940  | 
  Outer_Syntax.improper_command @{command_spec "typ"} "read and print type"
 | 
| 48792 | 941  | 
    (opt_modes -- (Parse.typ -- Scan.option (@{keyword "::"} |-- Parse.!!! Parse.sort))
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
942  | 
>> Isar_Cmd.print_type);  | 
| 5832 | 943  | 
|
| 24868 | 944  | 
val _ =  | 
| 48642 | 945  | 
  Outer_Syntax.improper_command @{command_spec "print_codesetup"} "print code generator setup"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
946  | 
(Scan.succeed (Toplevel.unknown_theory o  | 
| 
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
947  | 
Toplevel.keep (Code.print_codesetup o Toplevel.theory_of)));  | 
| 5832 | 948  | 
|
| 26184 | 949  | 
val _ =  | 
| 48642 | 950  | 
  Outer_Syntax.improper_command @{command_spec "unused_thms"} "find unused theorems"
 | 
| 36950 | 951  | 
(Scan.option ((Scan.repeat1 (Scan.unless Parse.minus Parse.name) --| Parse.minus) --  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
952  | 
Scan.option (Scan.repeat1 (Scan.unless Parse.minus Parse.name))) >> Isar_Cmd.unused_thms);  | 
| 26184 | 953  | 
|
| 5832 | 954  | 
|
| 
30142
 
8d6145694bb5
moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
 
wenzelm 
parents: 
29882 
diff
changeset
 | 
955  | 
|
| 5832 | 956  | 
(** system commands (for interactive mode only) **)  | 
957  | 
||
| 24868 | 958  | 
val _ =  | 
| 48642 | 959  | 
  Outer_Syntax.improper_command @{command_spec "cd"} "change current working directory"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
960  | 
(Parse.path >> (fn name => Toplevel.imperative (fn () => File.cd (Path.explode name))));  | 
| 5832 | 961  | 
|
| 24868 | 962  | 
val _ =  | 
| 48642 | 963  | 
  Outer_Syntax.improper_command @{command_spec "pwd"} "print current working directory"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
964  | 
(Scan.succeed (Toplevel.imperative (fn () => writeln (Path.print (File.pwd ())))));  | 
| 5832 | 965  | 
|
| 24868 | 966  | 
val _ =  | 
| 48642 | 967  | 
  Outer_Syntax.improper_command @{command_spec "use_thy"} "use theory file"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
968  | 
(Parse.position Parse.name >>  | 
| 
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
969  | 
(fn name => Toplevel.imperative (fn () => Thy_Info.use_thy name)));  | 
| 5832 | 970  | 
|
| 24868 | 971  | 
val _ =  | 
| 48642 | 972  | 
  Outer_Syntax.improper_command @{command_spec "remove_thy"} "remove theory from loader database"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
973  | 
(Parse.name >> (fn name => Toplevel.imperative (fn () => Thy_Info.remove_thy name)));  | 
| 7102 | 974  | 
|
| 24868 | 975  | 
val _ =  | 
| 48642 | 976  | 
  Outer_Syntax.improper_command @{command_spec "kill_thy"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
977  | 
"kill theory -- try to remove from loader database"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
978  | 
(Parse.name >> (fn name => Toplevel.imperative (fn () => Thy_Info.kill_thy name)));  | 
| 7931 | 979  | 
|
| 
53403
 
c09f4005d6bd
some explicit indication of Proof General legacy;
 
wenzelm 
parents: 
53378 
diff
changeset
 | 
980  | 
val _ = (*partial Proof General legacy*)  | 
| 48642 | 981  | 
  Outer_Syntax.improper_command @{command_spec "display_drafts"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
982  | 
"display raw source files with symbols"  | 
| 52549 | 983  | 
(Scan.repeat1 Parse.path >> (fn names =>  | 
984  | 
Toplevel.imperative (fn () => ignore (Present.display_drafts (map Path.explode names)))));  | 
|
| 14934 | 985  | 
|
| 52430 | 986  | 
val _ =  | 
987  | 
  Outer_Syntax.improper_command @{command_spec "print_state"}
 | 
|
988  | 
"print current proof state (if present)"  | 
|
| 
56887
 
1ca814da47ae
clarified print_state, which goes back to TTY loop before Proof General, and before separate print_context;
 
wenzelm 
parents: 
56869 
diff
changeset
 | 
989  | 
(opt_modes >> (fn modes => Toplevel.keep (Print_Mode.with_modes modes Toplevel.print_state)));  | 
| 52430 | 990  | 
|
| 
53403
 
c09f4005d6bd
some explicit indication of Proof General legacy;
 
wenzelm 
parents: 
53378 
diff
changeset
 | 
991  | 
val _ = (*Proof General legacy, e.g. for ProofGeneral-3.7.x*)  | 
| 48642 | 992  | 
  Outer_Syntax.improper_command @{command_spec "pr"} "print current proof state (if present)"
 | 
| 39165 | 993  | 
(opt_modes -- Scan.option Parse.nat >> (fn (modes, limit) =>  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
994  | 
Toplevel.keep (fn state =>  | 
| 
52438
 
7b5a5116f3af
back to keyword 'pr' :: diag as required for ProofGeneral command line -- reject this TTY command in Isabelle/jEdit by other means;
 
wenzelm 
parents: 
52430 
diff
changeset
 | 
995  | 
(if Isabelle_Process.is_active () then error "Illegal TTY command" else ();  | 
| 56467 | 996  | 
        case limit of NONE => () | SOME n => Options.default_put_int @{system_option goals_limit} n;
 | 
| 39165 | 997  | 
Toplevel.quiet := false;  | 
| 
56887
 
1ca814da47ae
clarified print_state, which goes back to TTY loop before Proof General, and before separate print_context;
 
wenzelm 
parents: 
56869 
diff
changeset
 | 
998  | 
Print_Mode.with_modes modes Toplevel.print_state state))));  | 
| 7199 | 999  | 
|
| 
53403
 
c09f4005d6bd
some explicit indication of Proof General legacy;
 
wenzelm 
parents: 
53378 
diff
changeset
 | 
1000  | 
val _ = (*Proof General legacy*)  | 
| 48642 | 1001  | 
  Outer_Syntax.improper_command @{command_spec "disable_pr"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
1002  | 
"disable printing of toplevel state"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1003  | 
(Scan.succeed (Toplevel.imperative (fn () => Toplevel.quiet := true)));  | 
| 5832 | 1004  | 
|
| 
53403
 
c09f4005d6bd
some explicit indication of Proof General legacy;
 
wenzelm 
parents: 
53378 
diff
changeset
 | 
1005  | 
val _ = (*Proof General legacy*)  | 
| 48642 | 1006  | 
  Outer_Syntax.improper_command @{command_spec "enable_pr"}
 | 
| 
46961
 
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
 
wenzelm 
parents: 
46958 
diff
changeset
 | 
1007  | 
"enable printing of toplevel state"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1008  | 
(Scan.succeed (Toplevel.imperative (fn () => Toplevel.quiet := false)));  | 
| 7222 | 1009  | 
|
| 24868 | 1010  | 
val _ =  | 
| 48642 | 1011  | 
  Outer_Syntax.improper_command @{command_spec "commit"}
 | 
| 50214 | 1012  | 
"commit current session to ML session image"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1013  | 
(Parse.opt_unit >> K (Toplevel.imperative Secure.commit));  | 
| 5832 | 1014  | 
|
| 24868 | 1015  | 
val _ =  | 
| 56629 | 1016  | 
  Outer_Syntax.improper_command @{command_spec "quit"} "quit Isabelle process"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1017  | 
(Parse.opt_unit >> (K (Toplevel.imperative quit)));  | 
| 5832 | 1018  | 
|
| 24868 | 1019  | 
val _ =  | 
| 48642 | 1020  | 
  Outer_Syntax.improper_command @{command_spec "exit"} "exit Isar loop"
 | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
1021  | 
(Scan.succeed  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1022  | 
(Toplevel.keep (fn state =>  | 
| 
37977
 
3ceccd415145
simplified/clarified theory loader: more explicit task management, kill old versions at start, commit results only in the very end, non-optional master dependency, do not store text in deps;
 
wenzelm 
parents: 
37949 
diff
changeset
 | 
1023  | 
(Context.set_thread_data (try Toplevel.generic_theory_of state);  | 
| 
38888
 
8248cda328de
moved Toplevel.run_command to Pure/PIDE/document.ML;
 
wenzelm 
parents: 
38870 
diff
changeset
 | 
1024  | 
raise Runtime.TERMINATE))));  | 
| 5832 | 1025  | 
|
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1026  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1027  | 
  Outer_Syntax.improper_command @{command_spec "welcome"} "print welcome message"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1028  | 
(Scan.succeed (Toplevel.imperative (writeln o Session.welcome)));  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1029  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1030  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1031  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1032  | 
(** raw Isar read-eval-print loop **)  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1033  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1034  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1035  | 
  Outer_Syntax.improper_command @{command_spec "init_toplevel"} "init toplevel point-of-interest"
 | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1036  | 
(Scan.succeed (Toplevel.imperative Isar.init));  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1037  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1038  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1039  | 
  Outer_Syntax.improper_command @{command_spec "linear_undo"} "undo commands"
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1040  | 
(Scan.optional Parse.nat 1 >>  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1041  | 
(fn n => Toplevel.imperative (fn () => Isar.linear_undo n)));  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1042  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1043  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1044  | 
  Outer_Syntax.improper_command @{command_spec "undo"} "undo commands (skipping closed proofs)"
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1045  | 
(Scan.optional Parse.nat 1 >>  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1046  | 
(fn n => Toplevel.imperative (fn () => Isar.undo n)));  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1047  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1048  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1049  | 
  Outer_Syntax.improper_command @{command_spec "undos_proof"}
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1050  | 
"undo commands (skipping closed proofs)"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1051  | 
(Scan.optional Parse.nat 1 >> (fn n =>  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1052  | 
Toplevel.keep (fn state =>  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1053  | 
if Toplevel.is_proof state then (Isar.undo n; Isar.print ()) else raise Toplevel.UNDEF)));  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1054  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1055  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1056  | 
  Outer_Syntax.improper_command @{command_spec "cannot_undo"}
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1057  | 
"partial undo -- Proof General legacy"  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1058  | 
(Parse.name >>  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1059  | 
(fn "end" => Toplevel.imperative (fn () => Isar.undo 1)  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1060  | 
        | txt => Toplevel.imperative (fn () => error ("Cannot undo " ^ quote txt))));
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1061  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1062  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1063  | 
  Outer_Syntax.improper_command @{command_spec "kill"}
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1064  | 
"kill partial proof or theory development"  | 
| 
51658
 
21c10672633b
discontinued Toplevel.no_timing complication -- also recovers timing of diagnostic commands, e.g. 'find_theorems';
 
wenzelm 
parents: 
51654 
diff
changeset
 | 
1065  | 
(Scan.succeed (Toplevel.imperative Isar.kill));  | 
| 
48646
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1066  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1067  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1068  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1069  | 
(** extraction of programs from proofs **)  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1070  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1071  | 
val parse_vars = Scan.optional (Parse.$$$ "(" |-- Parse.list1 Parse.name --| Parse.$$$ ")") [];
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1072  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1073  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1074  | 
  Outer_Syntax.command @{command_spec "realizers"}
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1075  | 
"specify realizers for primitive axioms / theorems, together with correctness proof"  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1076  | 
(Scan.repeat1 (Parse.xname -- parse_vars --| Parse.$$$ ":" -- Parse.string -- Parse.string) >>  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1077  | 
(fn xs => Toplevel.theory (fn thy => Extraction.add_realizers  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1078  | 
(map (fn (((a, vs), s1), s2) => (Global_Theory.get_thm thy a, (vs, s1, s2))) xs) thy)));  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1079  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1080  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1081  | 
  Outer_Syntax.command @{command_spec "realizability"}
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1082  | 
"add equations characterizing realizability"  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1083  | 
(Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_realizes_eqns));  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1084  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1085  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1086  | 
  Outer_Syntax.command @{command_spec "extract_type"}
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1087  | 
"add equations characterizing type of extracted program"  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1088  | 
(Scan.repeat1 Parse.string >> (Toplevel.theory o Extraction.add_typeof_eqns));  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1089  | 
|
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1090  | 
val _ =  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1091  | 
  Outer_Syntax.command @{command_spec "extract"} "extract terms from proofs"
 | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1092  | 
(Scan.repeat1 (Parse.xname -- parse_vars) >> (fn xs => Toplevel.theory (fn thy =>  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1093  | 
Extraction.extract (map (apfst (Global_Theory.get_thm thy)) xs) thy)));  | 
| 
 
91281e9472d8
more official command specifications, including source position;
 
wenzelm 
parents: 
48645 
diff
changeset
 | 
1094  | 
|
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1095  | 
|
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1096  | 
|
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1097  | 
(** end **)  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1098  | 
|
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1099  | 
val _ =  | 
| 48642 | 1100  | 
  Outer_Syntax.command @{command_spec "end"} "end context"
 | 
| 
48641
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1101  | 
(Scan.succeed  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1102  | 
(Toplevel.exit o Toplevel.end_local_theory o Toplevel.close_target o  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1103  | 
Toplevel.end_proof (K Proof.end_notepad)));  | 
| 
 
92b48b8abfe4
more standard bootstrapping of Pure outer syntax;
 
wenzelm 
parents: 
47416 
diff
changeset
 | 
1104  | 
|
| 5832 | 1105  | 
end;  | 
| 
27614
 
f38c25d106a7
renamed IsarCmd.nested_command to OuterSyntax.prepare_command;
 
wenzelm 
parents: 
27575 
diff
changeset
 | 
1106  |