| author | wenzelm | 
| Wed, 19 May 2021 15:53:55 +0200 | |
| changeset 73746 | b2d47981c8dc | 
| parent 69349 | 7cef9e386ffe | 
| child 74563 | 042041c0ebeb | 
| permissions | -rw-r--r-- | 
| 
5822
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/Isar/args.ML  | 
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Markus Wenzel, TU Muenchen  | 
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 
58011
 
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
 
wenzelm 
parents: 
57944 
diff
changeset
 | 
4  | 
Quasi-inner syntax based on outer tokens: concrete argument syntax of  | 
| 27811 | 5  | 
attributes, methods etc.  | 
| 
5822
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
signature ARGS =  | 
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
sig  | 
| 30513 | 10  | 
val context: Proof.context context_parser  | 
11  | 
val theory: theory context_parser  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
12  | 
val symbolic: Token.T parser  | 
| 30513 | 13  | 
val $$$ : string -> string parser  | 
14  | 
val add: string parser  | 
|
15  | 
val del: string parser  | 
|
16  | 
val colon: string parser  | 
|
17  | 
val query: string parser  | 
|
18  | 
val bang: string parser  | 
|
19  | 
val query_colon: string parser  | 
|
20  | 
val bang_colon: string parser  | 
|
| 53168 | 21  | 
val parens: 'a parser -> 'a parser  | 
22  | 
val bracks: 'a parser -> 'a parser  | 
|
23  | 
val mode: string -> bool parser  | 
|
| 30513 | 24  | 
val maybe: 'a parser -> 'a option parser  | 
| 60577 | 25  | 
val name_token: Token.T parser  | 
26  | 
val name: string parser  | 
|
| 
69349
 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 
wenzelm 
parents: 
63120 
diff
changeset
 | 
27  | 
val name_position: (string * Position.T) parser  | 
| 55111 | 28  | 
val cartouche_inner_syntax: string parser  | 
| 59809 | 29  | 
val cartouche_input: Input.source parser  | 
| 59857 | 30  | 
val text_token: Token.T parser  | 
| 
63120
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
31  | 
val embedded_token: Token.T parser  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
32  | 
val embedded_inner_syntax: string parser  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
33  | 
val embedded_input: Input.source parser  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
34  | 
val embedded: string parser  | 
| 
69349
 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 
wenzelm 
parents: 
63120 
diff
changeset
 | 
35  | 
val embedded_position: (string * Position.T) parser  | 
| 59809 | 36  | 
val text_input: Input.source parser  | 
| 
56500
 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 
wenzelm 
parents: 
56499 
diff
changeset
 | 
37  | 
val text: string parser  | 
| 30513 | 38  | 
val binding: binding parser  | 
39  | 
val alt_name: string parser  | 
|
40  | 
val liberal_name: string parser  | 
|
41  | 
val var: indexname parser  | 
|
| 58012 | 42  | 
val internal_source: Token.src parser  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
43  | 
val internal_name: Token.name_value parser  | 
| 30513 | 44  | 
val internal_typ: typ parser  | 
45  | 
val internal_term: term parser  | 
|
46  | 
val internal_fact: thm list parser  | 
|
47  | 
val internal_attribute: (morphism -> attribute) parser  | 
|
| 58017 | 48  | 
val internal_declaration: declaration parser  | 
| 58012 | 49  | 
val named_source: (Token.T -> Token.src) -> Token.src parser  | 
| 30513 | 50  | 
val named_typ: (string -> typ) -> typ parser  | 
51  | 
val named_term: (string -> term) -> term parser  | 
|
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
56500 
diff
changeset
 | 
52  | 
val named_fact: (string -> string option * thm list) -> thm list parser  | 
| 58017 | 53  | 
val named_attribute: (string * Position.T -> morphism -> attribute) ->  | 
54  | 
(morphism -> attribute) parser  | 
|
| 59064 | 55  | 
val text_declaration: (Input.source -> declaration) -> declaration parser  | 
| 59649 | 56  | 
val cartouche_declaration: (Input.source -> declaration) -> declaration parser  | 
| 30513 | 57  | 
val typ_abbrev: typ context_parser  | 
58  | 
val typ: typ context_parser  | 
|
59  | 
val term: term context_parser  | 
|
| 55155 | 60  | 
val term_pattern: term context_parser  | 
| 30513 | 61  | 
val term_abbrev: term context_parser  | 
62  | 
val prop: term context_parser  | 
|
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55915 
diff
changeset
 | 
63  | 
  val type_name: {proper: bool, strict: bool} -> string context_parser
 | 
| 55954 | 64  | 
  val const: {proper: bool, strict: bool} -> string context_parser
 | 
| 30514 | 65  | 
val goal_spec: ((int -> tactic) -> tactic) context_parser  | 
| 
5822
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
end;  | 
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
structure Args: ARGS =  | 
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
struct  | 
| 
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
|
| 27811 | 71  | 
(** argument scanners **)  | 
| 
5822
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
|
| 27371 | 73  | 
(* context *)  | 
74  | 
||
75  | 
fun context x = (Scan.state >> Context.proof_of) x;  | 
|
76  | 
fun theory x = (Scan.state >> Context.theory_of) x;  | 
|
77  | 
||
78  | 
||
| 27811 | 79  | 
(* basic *)  | 
| 5878 | 80  | 
|
| 56201 | 81  | 
val ident = Parse.token  | 
| 36950 | 82  | 
(Parse.short_ident || Parse.long_ident || Parse.sym_ident || Parse.term_var ||  | 
83  | 
Parse.type_ident || Parse.type_var || Parse.number);  | 
|
| 5878 | 84  | 
|
| 
56500
 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 
wenzelm 
parents: 
56499 
diff
changeset
 | 
85  | 
val string = Parse.token Parse.string;  | 
| 
56499
 
7e0178c84994
allow text cartouches in regular outer syntax categories "text" and "altstring";
 
wenzelm 
parents: 
56333 
diff
changeset
 | 
86  | 
val alt_string = Parse.token (Parse.alt_string || Parse.cartouche);  | 
| 56201 | 87  | 
val symbolic = Parse.token (Parse.keyword_with Token.ident_or_symbolic);  | 
| 27811 | 88  | 
|
| 
55914
 
c5b752d549e3
clarified init_assignable: make double-sure that initial values are reset;
 
wenzelm 
parents: 
55828 
diff
changeset
 | 
89  | 
fun $$$ x =  | 
| 56201 | 90  | 
(ident || Parse.token Parse.keyword) :|-- (fn tok =>  | 
| 56063 | 91  | 
let val y = Token.content_of tok in  | 
92  | 
if x = y  | 
|
| 56202 | 93  | 
then (Token.assign (SOME (Token.Literal (false, Markup.quasi_keyword))) tok; Scan.succeed x)  | 
| 56063 | 94  | 
else Scan.fail  | 
95  | 
end);  | 
|
| 5878 | 96  | 
|
| 10035 | 97  | 
val add = $$$ "add";  | 
98  | 
val del = $$$ "del";  | 
|
| 8803 | 99  | 
val colon = $$$ ":";  | 
| 10035 | 100  | 
val query = $$$ "?";  | 
101  | 
val bang = $$$ "!";  | 
|
| 20111 | 102  | 
val query_colon = $$$ "?" ^^ $$$ ":";  | 
103  | 
val bang_colon = $$$ "!" ^^ $$$ ":";  | 
|
| 10035 | 104  | 
|
| 8803 | 105  | 
fun parens scan = $$$ "(" |-- scan --| $$$ ")";
 | 
| 10150 | 106  | 
fun bracks scan = $$$ "[" |-- scan --| $$$ "]";  | 
| 53168 | 107  | 
fun mode s = Scan.optional (parens ($$$ s) >> K true) false;  | 
| 15703 | 108  | 
fun maybe scan = $$$ "_" >> K NONE || scan >> SOME;  | 
| 5878 | 109  | 
|
| 60577 | 110  | 
val name_token = ident || string;  | 
111  | 
val name = name_token >> Token.content_of;  | 
|
| 
69349
 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 
wenzelm 
parents: 
63120 
diff
changeset
 | 
112  | 
val name_position = name_token >> (Input.source_content o Token.input_of);  | 
| 60577 | 113  | 
|
| 56201 | 114  | 
val cartouche = Parse.token Parse.cartouche;  | 
| 55111 | 115  | 
val cartouche_inner_syntax = cartouche >> Token.inner_syntax_of;  | 
| 59809 | 116  | 
val cartouche_input = cartouche >> Token.input_of;  | 
| 
55045
 
99056d23e05b
cartouche within nested args (attributes, methods, etc.);
 
wenzelm 
parents: 
55033 
diff
changeset
 | 
117  | 
|
| 
63120
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
118  | 
val embedded_token = ident || string || cartouche;  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
119  | 
val embedded_inner_syntax = embedded_token >> Token.inner_syntax_of;  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
120  | 
val embedded_input = embedded_token >> Token.input_of;  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
121  | 
val embedded = embedded_token >> Token.content_of;  | 
| 
69349
 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 
wenzelm 
parents: 
63120 
diff
changeset
 | 
122  | 
val embedded_position = embedded_input >> Input.source_content;  | 
| 
63120
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
123  | 
|
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
124  | 
val text_token = embedded_token || Parse.token Parse.verbatim;  | 
| 59809 | 125  | 
val text_input = text_token >> Token.input_of;  | 
| 
56500
 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 
wenzelm 
parents: 
56499 
diff
changeset
 | 
126  | 
val text = text_token >> Token.content_of;  | 
| 
 
90f17a04567d
proper Args.name vs. Args.text as documented (in contrast to adhoc union in 75aaee32893d, which had to cope with more limited Args.T);
 
wenzelm 
parents: 
56499 
diff
changeset
 | 
127  | 
|
| 
69349
 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 
wenzelm 
parents: 
63120 
diff
changeset
 | 
128  | 
val binding = Parse.input name >> (Binding.make o Input.source_content);  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
129  | 
val alt_name = alt_string >> Token.content_of;  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
130  | 
val liberal_name = (symbolic >> Token.content_of) || name;  | 
| 8233 | 131  | 
|
| 
63120
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
132  | 
val var =  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
133  | 
(ident >> Token.content_of) :|-- (fn x =>  | 
| 
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
134  | 
(case Lexicon.read_variable x of SOME v => Scan.succeed v | NONE => Scan.fail));  | 
| 5878 | 135  | 
|
| 
5822
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
|
| 15703 | 137  | 
(* values *)  | 
138  | 
||
| 27811 | 139  | 
fun value dest = Scan.some (fn arg =>  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
140  | 
(case Token.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE));  | 
| 15703 | 141  | 
|
| 58012 | 142  | 
val internal_source = value (fn Token.Source src => src);  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
143  | 
val internal_name = value (fn Token.Name (a, _) => a);  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
144  | 
val internal_typ = value (fn Token.Typ T => T);  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
145  | 
val internal_term = value (fn Token.Term t => t);  | 
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
56500 
diff
changeset
 | 
146  | 
val internal_fact = value (fn Token.Fact (_, ths) => ths);  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
147  | 
val internal_attribute = value (fn Token.Attribute att => att);  | 
| 58017 | 148  | 
val internal_declaration = value (fn Token.Declaration decl => decl);  | 
| 15703 | 149  | 
|
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
150  | 
fun named_source read =  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
151  | 
internal_source || name_token >> Token.evaluate Token.Source read;  | 
| 58012 | 152  | 
|
| 60577 | 153  | 
fun named_typ read =  | 
| 
63120
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
154  | 
internal_typ || embedded_token >> Token.evaluate Token.Typ (read o Token.inner_syntax_of);  | 
| 60577 | 155  | 
|
156  | 
fun named_term read =  | 
|
| 
63120
 
629a4c5e953e
embedded content may be delimited via cartouches;
 
wenzelm 
parents: 
61814 
diff
changeset
 | 
157  | 
internal_term || embedded_token >> Token.evaluate Token.Term (read o Token.inner_syntax_of);  | 
| 42464 | 158  | 
|
| 
57942
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
56500 
diff
changeset
 | 
159  | 
fun named_fact get =  | 
| 
 
e5bec882fdd0
more informative Token.Fact: retain name of dynamic fact (without selection);
 
wenzelm 
parents: 
56500 
diff
changeset
 | 
160  | 
internal_fact ||  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
161  | 
name_token >> Token.evaluate Token.Fact (get o Token.content_of) >> #2 ||  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
162  | 
alt_string >> Token.evaluate Token.Fact (get o Token.inner_syntax_of) >> #2;  | 
| 42464 | 163  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36950 
diff
changeset
 | 
164  | 
fun named_attribute att =  | 
| 42464 | 165  | 
internal_attribute ||  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
166  | 
name_token >>  | 
| 
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
167  | 
Token.evaluate Token.Attribute (fn tok => att (Token.content_of tok, Token.pos_of tok));  | 
| 15703 | 168  | 
|
| 58017 | 169  | 
fun text_declaration read =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
170  | 
internal_declaration || text_token >> Token.evaluate Token.Declaration (read o Token.input_of);  | 
| 58017 | 171  | 
|
| 59649 | 172  | 
fun cartouche_declaration read =  | 
| 
61814
 
1ca1142e1711
clarified type Token.src: plain token list, with usual implicit value assignment;
 
wenzelm 
parents: 
60577 
diff
changeset
 | 
173  | 
internal_declaration || cartouche >> Token.evaluate Token.Declaration (read o Token.input_of);  | 
| 59649 | 174  | 
|
| 15703 | 175  | 
|
| 5878 | 176  | 
(* terms and types *)  | 
177  | 
||
| 42360 | 178  | 
val typ_abbrev = Scan.peek (named_typ o Proof_Context.read_typ_abbrev o Context.proof_of);  | 
| 25331 | 179  | 
val typ = Scan.peek (named_typ o Syntax.read_typ o Context.proof_of);  | 
| 
24508
 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 
wenzelm 
parents: 
24244 
diff
changeset
 | 
180  | 
val term = Scan.peek (named_term o Syntax.read_term o Context.proof_of);  | 
| 55155 | 181  | 
val term_pattern = Scan.peek (named_term o Proof_Context.read_term_pattern o Context.proof_of);  | 
| 42360 | 182  | 
val term_abbrev = Scan.peek (named_term o Proof_Context.read_term_abbrev o Context.proof_of);  | 
| 
24508
 
c8b82fec6447
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 
wenzelm 
parents: 
24244 
diff
changeset
 | 
183  | 
val prop = Scan.peek (named_term o Syntax.read_prop o Context.proof_of);  | 
| 18635 | 184  | 
|
| 5878 | 185  | 
|
| 15703 | 186  | 
(* type and constant names *)  | 
187  | 
||
| 
55951
 
c07d184aebe9
tuned signature -- more uniform check_type_name/read_type_name;
 
wenzelm 
parents: 
55915 
diff
changeset
 | 
188  | 
fun type_name flags =  | 
| 56002 | 189  | 
Scan.peek (named_typ o Proof_Context.read_type_name flags o Context.proof_of)  | 
| 18998 | 190  | 
>> (fn Type (c, _) => c | TFree (a, _) => a | _ => "");  | 
| 15703 | 191  | 
|
| 55954 | 192  | 
fun const flags =  | 
| 56002 | 193  | 
Scan.peek (named_term o Proof_Context.read_const flags o Context.proof_of)  | 
| 18998 | 194  | 
>> (fn Const (c, _) => c | Free (x, _) => x | _ => "");  | 
| 7553 | 195  | 
|
| 15703 | 196  | 
|
| 27811 | 197  | 
(* improper method arguments *)  | 
| 15703 | 198  | 
|
| 8536 | 199  | 
val from_to =  | 
| 36950 | 200  | 
Parse.nat -- ($$$ "-" |-- Parse.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) ||  | 
201  | 
Parse.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) ||  | 
|
202  | 
Parse.nat >> (fn i => fn tac => tac i) ||  | 
|
| 15703 | 203  | 
$$$ "!" >> K ALLGOALS;  | 
| 8536 | 204  | 
|
| 56202 | 205  | 
val goal = Parse.keyword_improper "[" |-- Parse.!!! (from_to --| Parse.keyword_improper "]");  | 
| 30514 | 206  | 
fun goal_spec x = Scan.lift (Scan.optional goal (fn tac => tac 1)) x;  | 
| 8233 | 207  | 
|
| 
5822
 
3f824514ad88
Concrete argument syntax (for attributes, methods etc.).
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
end;  |