| author | wenzelm | 
| Mon, 27 Jun 2011 22:20:49 +0200 | |
| changeset 43564 | 9864182c6bad | 
| parent 42657 | 6b404fe40877 | 
| child 43775 | b361c7d184e7 | 
| permissions | -rw-r--r-- | 
| 
36949
 
080e85d46108
renamed structure OuterKeyword to Keyword and OuterParse to Parse, keeping the old names as legacy aliases for some time;
 
wenzelm 
parents: 
35838 
diff
changeset
 | 
1  | 
(* Title: Pure/Isar/parse.ML  | 
| 5826 | 2  | 
Author: Markus Wenzel, TU Muenchen  | 
3  | 
||
4  | 
Generic parsers for Isabelle/Isar outer syntax.  | 
|
5  | 
*)  | 
|
6  | 
||
| 
36949
 
080e85d46108
renamed structure OuterKeyword to Keyword and OuterParse to Parse, keeping the old names as legacy aliases for some time;
 
wenzelm 
parents: 
35838 
diff
changeset
 | 
7  | 
signature PARSE =  | 
| 5826 | 8  | 
sig  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
9  | 
type 'a parser = Token.T list -> 'a * Token.T list  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
10  | 
type 'a context_parser = Context.generic * Token.T list -> 'a * (Context.generic * Token.T list)  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
11  | 
val group: string -> (Token.T list -> 'a) -> Token.T list -> 'a  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
12  | 
val !!! : (Token.T list -> 'a) -> Token.T list -> 'a  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
13  | 
val !!!! : (Token.T list -> 'a) -> Token.T list -> 'a  | 
| 12047 | 14  | 
  val triple1: ('a * 'b) * 'c -> 'a * 'b * 'c
 | 
15  | 
  val triple2: 'a * ('b * 'c) -> 'a * 'b * 'c
 | 
|
16  | 
  val triple_swap: ('a * 'b) * 'c -> ('a * 'c) * 'b
 | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
17  | 
val not_eof: Token.T parser  | 
| 42326 | 18  | 
  val position: 'a parser -> ('a * Position.T) parser
 | 
| 42657 | 19  | 
val source_position: 'a parser -> (Symbol_Pos.text * Position.T) parser  | 
| 42326 | 20  | 
val inner_syntax: 'a parser -> string parser  | 
| 29310 | 21  | 
val command: string parser  | 
22  | 
val keyword: string parser  | 
|
23  | 
val short_ident: string parser  | 
|
24  | 
val long_ident: string parser  | 
|
25  | 
val sym_ident: string parser  | 
|
26  | 
val minus: string parser  | 
|
27  | 
val term_var: string parser  | 
|
28  | 
val type_ident: string parser  | 
|
29  | 
val type_var: string parser  | 
|
30  | 
val number: string parser  | 
|
| 
40290
 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 
wenzelm 
parents: 
36959 
diff
changeset
 | 
31  | 
val float_number: string parser  | 
| 29310 | 32  | 
val string: string parser  | 
33  | 
val alt_string: string parser  | 
|
34  | 
val verbatim: string parser  | 
|
35  | 
val sync: string parser  | 
|
36  | 
val eof: string parser  | 
|
37  | 
val keyword_with: (string -> bool) -> string parser  | 
|
38  | 
val keyword_ident_or_symbolic: string parser  | 
|
39  | 
val $$$ : string -> string parser  | 
|
40  | 
val reserved: string -> string parser  | 
|
41  | 
val semicolon: string parser  | 
|
42  | 
val underscore: string parser  | 
|
43  | 
val maybe: 'a parser -> 'a option parser  | 
|
44  | 
val tag_name: string parser  | 
|
45  | 
val tags: string list parser  | 
|
46  | 
val opt_unit: unit parser  | 
|
47  | 
val opt_keyword: string -> bool parser  | 
|
48  | 
val begin: string parser  | 
|
49  | 
val opt_begin: bool parser  | 
|
50  | 
val nat: int parser  | 
|
51  | 
val int: int parser  | 
|
| 
40290
 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 
wenzelm 
parents: 
36959 
diff
changeset
 | 
52  | 
val real: real parser  | 
| 29310 | 53  | 
val enum: string -> 'a parser -> 'a list parser  | 
54  | 
val enum1: string -> 'a parser -> 'a list parser  | 
|
55  | 
val and_list: 'a parser -> 'a list parser  | 
|
56  | 
val and_list1: 'a parser -> 'a list parser  | 
|
| 30511 | 57  | 
val enum': string -> 'a context_parser -> 'a list context_parser  | 
58  | 
val enum1': string -> 'a context_parser -> 'a list context_parser  | 
|
59  | 
val and_list': 'a context_parser -> 'a list context_parser  | 
|
60  | 
val and_list1': 'a context_parser -> 'a list context_parser  | 
|
| 29310 | 61  | 
val list: 'a parser -> 'a list parser  | 
62  | 
val list1: 'a parser -> 'a list parser  | 
|
63  | 
val name: bstring parser  | 
|
| 29581 | 64  | 
val binding: binding parser  | 
| 29310 | 65  | 
val xname: xstring parser  | 
66  | 
val text: string parser  | 
|
67  | 
val path: Path.T parser  | 
|
| 
40800
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
40793 
diff
changeset
 | 
68  | 
val liberal_name: xstring parser  | 
| 29310 | 69  | 
val parname: string parser  | 
| 29581 | 70  | 
val parbinding: binding parser  | 
| 29310 | 71  | 
val sort: string parser  | 
72  | 
val arity: (string * string list * string) parser  | 
|
73  | 
val multi_arity: (string list * string list * string) parser  | 
|
74  | 
val type_args: string list parser  | 
|
| 35838 | 75  | 
val type_args_constrained: (string * string option) list parser  | 
| 29310 | 76  | 
val typ_group: string parser  | 
77  | 
val typ: string parser  | 
|
78  | 
val mixfix: mixfix parser  | 
|
79  | 
val mixfix': mixfix parser  | 
|
80  | 
val opt_mixfix: mixfix parser  | 
|
81  | 
val opt_mixfix': mixfix parser  | 
|
82  | 
val where_: string parser  | 
|
| 42299 | 83  | 
val const_decl: (string * string * mixfix) parser  | 
| 30339 | 84  | 
val const_binding: (binding * string * mixfix) parser  | 
| 29581 | 85  | 
val params: (binding * string option) list parser  | 
86  | 
val simple_fixes: (binding * string option) list parser  | 
|
87  | 
val fixes: (binding * string option * mixfix) list parser  | 
|
88  | 
val for_fixes: (binding * string option * mixfix) list parser  | 
|
89  | 
val for_simple_fixes: (binding * string option) list parser  | 
|
| 30573 | 90  | 
val ML_source: (Symbol_Pos.text * Position.T) parser  | 
91  | 
val doc_source: (Symbol_Pos.text * Position.T) parser  | 
|
| 29310 | 92  | 
val term_group: string parser  | 
93  | 
val prop_group: string parser  | 
|
94  | 
val term: string parser  | 
|
95  | 
val prop: string parser  | 
|
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
96  | 
val type_const: string parser  | 
| 
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
97  | 
val const: string parser  | 
| 
40793
 
d21aedaa91e7
added Parse.literal_fact with proper inner_syntax markup (source position);
 
wenzelm 
parents: 
40296 
diff
changeset
 | 
98  | 
val literal_fact: string parser  | 
| 29310 | 99  | 
val propp: (string * string list) parser  | 
100  | 
val termp: (string * string list) parser  | 
|
101  | 
val target: xstring parser  | 
|
102  | 
val opt_target: xstring option parser  | 
|
| 5826 | 103  | 
end;  | 
104  | 
||
| 
36949
 
080e85d46108
renamed structure OuterKeyword to Keyword and OuterParse to Parse, keeping the old names as legacy aliases for some time;
 
wenzelm 
parents: 
35838 
diff
changeset
 | 
105  | 
structure Parse: PARSE =  | 
| 5826 | 106  | 
struct  | 
107  | 
||
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
108  | 
type 'a parser = Token.T list -> 'a * Token.T list;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
109  | 
type 'a context_parser = Context.generic * Token.T list -> 'a * (Context.generic * Token.T list);  | 
| 29310 | 110  | 
|
| 5826 | 111  | 
|
112  | 
(** error handling **)  | 
|
113  | 
||
114  | 
(* group atomic parsers (no cuts!) *)  | 
|
115  | 
||
116  | 
fun fail_with s = Scan.fail_with  | 
|
117  | 
(fn [] => s ^ " expected (past end-of-file!)"  | 
|
| 42519 | 118  | 
| tok :: _ =>  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
119  | 
(case Token.text_of tok of  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
120  | 
(txt, "") =>  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
121  | 
s ^ " expected,\nbut " ^ txt ^ Token.pos_of tok ^ " was found"  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
122  | 
| (txt1, txt2) =>  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
123  | 
s ^ " expected,\nbut " ^ txt1 ^ Token.pos_of tok ^ " was found:\n" ^ txt2));  | 
| 5826 | 124  | 
|
125  | 
fun group s scan = scan || fail_with s;  | 
|
126  | 
||
127  | 
||
| 5877 | 128  | 
(* cut *)  | 
| 5826 | 129  | 
|
| 
8581
 
5c7ed2af8bfb
!!!! = cut "Corrupted outer syntax in presentation";
 
wenzelm 
parents: 
8350 
diff
changeset
 | 
130  | 
fun cut kind scan =  | 
| 5826 | 131  | 
let  | 
132  | 
fun get_pos [] = " (past end-of-file!)"  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
133  | 
| get_pos (tok :: _) = Token.pos_of tok;  | 
| 5826 | 134  | 
|
| 15531 | 135  | 
fun err (toks, NONE) = kind ^ get_pos toks  | 
| 25625 | 136  | 
| err (toks, SOME msg) =  | 
137  | 
if String.isPrefix kind msg then msg  | 
|
138  | 
else kind ^ get_pos toks ^ ": " ^ msg;  | 
|
| 5826 | 139  | 
in Scan.!! err scan end;  | 
140  | 
||
| 8586 | 141  | 
fun !!! scan = cut "Outer syntax error" scan;  | 
142  | 
fun !!!! scan = cut "Corrupted outer syntax in presentation" scan;  | 
|
| 
8581
 
5c7ed2af8bfb
!!!! = cut "Corrupted outer syntax in presentation";
 
wenzelm 
parents: 
8350 
diff
changeset
 | 
143  | 
|
| 5826 | 144  | 
|
145  | 
||
146  | 
(** basic parsers **)  | 
|
147  | 
||
148  | 
(* utils *)  | 
|
149  | 
||
150  | 
fun triple1 ((x, y), z) = (x, y, z);  | 
|
151  | 
fun triple2 (x, (y, z)) = (x, y, z);  | 
|
| 6430 | 152  | 
fun triple_swap ((x, y), z) = ((x, z), y);  | 
| 5826 | 153  | 
|
154  | 
||
155  | 
(* tokens *)  | 
|
156  | 
||
| 27815 | 157  | 
fun RESET_VALUE atom = (*required for all primitive parsers*)  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
158  | 
Scan.ahead (Scan.one (K true)) -- atom >> (fn (arg, x) => (Token.assign NONE arg; x));  | 
| 27815 | 159  | 
|
160  | 
||
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
161  | 
val not_eof = RESET_VALUE (Scan.one Token.not_eof);  | 
| 15703 | 162  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
163  | 
fun position scan = (Scan.ahead not_eof >> Token.position_of) -- scan >> Library.swap;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
164  | 
fun source_position atom = Scan.ahead atom |-- not_eof >> Token.source_position_of;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
165  | 
fun inner_syntax atom = Scan.ahead atom |-- not_eof >> Token.source_of;  | 
| 5826 | 166  | 
|
167  | 
fun kind k =  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
168  | 
group (Token.str_of_kind k) (RESET_VALUE (Scan.one (Token.is_kind k) >> Token.content_of));  | 
| 5826 | 169  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
170  | 
val command = kind Token.Command;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
171  | 
val keyword = kind Token.Keyword;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
172  | 
val short_ident = kind Token.Ident;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
173  | 
val long_ident = kind Token.LongIdent;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
174  | 
val sym_ident = kind Token.SymIdent;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
175  | 
val term_var = kind Token.Var;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
176  | 
val type_ident = kind Token.TypeIdent;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
177  | 
val type_var = kind Token.TypeVar;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
178  | 
val number = kind Token.Nat;  | 
| 
40290
 
47f572aff50a
support for floating-point tokens in outer syntax (coinciding with inner syntax version);
 
wenzelm 
parents: 
36959 
diff
changeset
 | 
179  | 
val float_number = kind Token.Float;  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
180  | 
val string = kind Token.String;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
181  | 
val alt_string = kind Token.AltString;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
182  | 
val verbatim = kind Token.Verbatim;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
183  | 
val sync = kind Token.Sync;  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
184  | 
val eof = kind Token.EOF;  | 
| 5826 | 185  | 
|
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
186  | 
fun keyword_with pred = RESET_VALUE (Scan.one (Token.keyword_with pred) >> Token.content_of);  | 
| 
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
187  | 
val keyword_ident_or_symbolic = keyword_with Token.ident_or_symbolic;  | 
| 27815 | 188  | 
|
| 5826 | 189  | 
fun $$$ x =  | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
190  | 
group (Token.str_of_kind Token.Keyword ^ " " ^ quote x) (keyword_with (fn y => x = y));  | 
| 9131 | 191  | 
|
| 16030 | 192  | 
fun reserved x =  | 
| 
27753
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
193  | 
  group ("reserved identifier " ^ quote x)
 | 
| 
36959
 
f5417836dbea
renamed structure OuterLex to Token and type token to Token.T, keeping legacy aliases for some time;
 
wenzelm 
parents: 
36955 
diff
changeset
 | 
194  | 
(RESET_VALUE (Scan.one (Token.ident_with (fn y => x = y)) >> Token.content_of));  | 
| 16030 | 195  | 
|
| 9131 | 196  | 
val semicolon = $$$ ";";  | 
| 5826 | 197  | 
|
| 15703 | 198  | 
val minus = sym_ident :-- (fn "-" => Scan.succeed () | _ => Scan.fail) >> #1;  | 
| 
11792
 
311eee3d63b6
parser for underscore (actually a symbolic identifier!);
 
wenzelm 
parents: 
11651 
diff
changeset
 | 
199  | 
val underscore = sym_ident :-- (fn "_" => Scan.succeed () | _ => Scan.fail) >> #1;  | 
| 15703 | 200  | 
fun maybe scan = underscore >> K NONE || scan >> SOME;  | 
| 
11792
 
311eee3d63b6
parser for underscore (actually a symbolic identifier!);
 
wenzelm 
parents: 
11651 
diff
changeset
 | 
201  | 
|
| 14835 | 202  | 
val nat = number >> (#1 o Library.read_int o Symbol.explode);  | 
| 27815 | 203  | 
val int = Scan.optional (minus >> K ~1) 1 -- nat >> op *;  | 
| 40296 | 204  | 
val real = float_number >> (the o Real.fromString) || int >> Real.fromInt;  | 
| 5826 | 205  | 
|
| 17070 | 206  | 
val tag_name = group "tag name" (short_ident || string);  | 
207  | 
val tags = Scan.repeat ($$$ "%" |-- !!! tag_name);  | 
|
208  | 
||
| 7930 | 209  | 
val opt_unit = Scan.optional ($$$ "(" -- $$$ ")" >> (K ())) ();
 | 
| 14646 | 210  | 
fun opt_keyword s = Scan.optional ($$$ "(" |-- !!! (($$$ s >> K true) --| $$$ ")")) false;
 | 
211  | 
||
| 20983 | 212  | 
val begin = $$$ "begin";  | 
213  | 
val opt_begin = Scan.optional (begin >> K true) false;  | 
|
| 20961 | 214  | 
|
| 5826 | 215  | 
|
216  | 
(* enumerations *)  | 
|
217  | 
||
| 25999 | 218  | 
fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan);  | 
| 5826 | 219  | 
fun enum sep scan = enum1 sep scan || Scan.succeed [];  | 
220  | 
||
| 27815 | 221  | 
fun enum1' sep scan = scan ::: Scan.repeat (Scan.lift ($$$ sep) |-- scan);  | 
222  | 
fun enum' sep scan = enum1' sep scan || Scan.succeed [];  | 
|
| 5826 | 223  | 
|
| 6013 | 224  | 
fun and_list1 scan = enum1 "and" scan;  | 
225  | 
fun and_list scan = enum "and" scan;  | 
|
226  | 
||
| 27815 | 227  | 
fun and_list1' scan = enum1' "and" scan;  | 
228  | 
fun and_list' scan = enum' "and" scan;  | 
|
229  | 
||
230  | 
fun list1 scan = enum1 "," scan;  | 
|
231  | 
fun list scan = enum "," scan;  | 
|
232  | 
||
| 5826 | 233  | 
|
| 5960 | 234  | 
(* names and text *)  | 
| 5826 | 235  | 
|
| 8146 | 236  | 
val name = group "name declaration" (short_ident || sym_ident || string || number);  | 
| 
30223
 
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
 
wenzelm 
parents: 
29581 
diff
changeset
 | 
237  | 
val binding = position name >> Binding.make;  | 
| 8146 | 238  | 
val xname = group "name reference" (short_ident || long_ident || sym_ident || string || number);  | 
239  | 
val text = group "text" (short_ident || long_ident || sym_ident || string || number || verbatim);  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21609 
diff
changeset
 | 
240  | 
val path = group "file name/path specification" name >> Path.explode;  | 
| 6553 | 241  | 
|
| 
40800
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
40793 
diff
changeset
 | 
242  | 
val liberal_name = keyword_ident_or_symbolic || xname;  | 
| 
 
330eb65c9469
Parse.liberal_name for document antiquotations and attributes;
 
wenzelm 
parents: 
40793 
diff
changeset
 | 
243  | 
|
| 18898 | 244  | 
val parname = Scan.optional ($$$ "(" |-- name --| $$$ ")") "";
 | 
| 28965 | 245  | 
val parbinding = Scan.optional ($$$ "(" |-- binding --| $$$ ")") Binding.empty;
 | 
| 18898 | 246  | 
|
| 6553 | 247  | 
|
| 6372 | 248  | 
(* sorts and arities *)  | 
| 5826 | 249  | 
|
| 
27753
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
250  | 
val sort = group "sort" (inner_syntax xname);  | 
| 5826 | 251  | 
|
| 22331 | 252  | 
val arity = xname -- ($$$ "::" |-- !!!  | 
253  | 
  (Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- sort)) >> triple2;
 | 
|
| 5826 | 254  | 
|
| 25541 | 255  | 
val multi_arity = and_list1 xname -- ($$$ "::" |-- !!!  | 
256  | 
  (Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- sort)) >> triple2;
 | 
|
257  | 
||
| 5826 | 258  | 
|
259  | 
(* types *)  | 
|
260  | 
||
| 
27753
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
261  | 
val typ_group = group "type"  | 
| 8146 | 262  | 
(short_ident || long_ident || sym_ident || type_ident || type_var || string || number);  | 
| 5826 | 263  | 
|
| 
27753
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
264  | 
val typ = inner_syntax typ_group;  | 
| 
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
265  | 
|
| 35838 | 266  | 
fun type_arguments arg =  | 
267  | 
arg >> single ||  | 
|
268  | 
  $$$ "(" |-- !!! (list1 arg --| $$$ ")") ||
 | 
|
| 5826 | 269  | 
Scan.succeed [];  | 
270  | 
||
| 35838 | 271  | 
val type_args = type_arguments type_ident;  | 
272  | 
val type_args_constrained = type_arguments (type_ident -- Scan.option ($$$ "::" |-- !!! sort));  | 
|
273  | 
||
| 5826 | 274  | 
|
275  | 
(* mixfix annotations *)  | 
|
276  | 
||
| 18669 | 277  | 
val mfix = string --  | 
278  | 
!!! (Scan.optional ($$$ "[" |-- !!! (list nat --| $$$ "]")) [] --  | 
|
| 
42297
 
140f283266b7
discontinued Syntax.max_pri, which is not really a symbolic parameter;
 
wenzelm 
parents: 
42287 
diff
changeset
 | 
279  | 
Scan.optional nat 1000) >> (Mixfix o triple2);  | 
| 18669 | 280  | 
|
| 35130 | 281  | 
val infx = $$$ "infix" |-- !!! (string -- nat >> Infix);  | 
282  | 
val infxl = $$$ "infixl" |-- !!! (string -- nat >> Infixl);  | 
|
283  | 
val infxr = $$$ "infixr" |-- !!! (string -- nat >> Infixr);  | 
|
| 5826 | 284  | 
|
| 18669 | 285  | 
val binder = $$$ "binder" |--  | 
286  | 
!!! (string -- ($$$ "[" |-- nat --| $$$ "]" -- nat || nat >> (fn n => (n, n))))  | 
|
287  | 
>> (Binder o triple2);  | 
|
288  | 
||
289  | 
fun annotation guard fix = $$$ "(" |-- guard (fix --| $$$ ")");
 | 
|
290  | 
fun opt_annotation guard fix = Scan.optional (annotation guard fix) NoSyn;  | 
|
291  | 
||
292  | 
val mixfix = annotation !!! (mfix || binder || infxl || infxr || infx);  | 
|
| 21609 | 293  | 
val mixfix' = annotation I (mfix || binder || infxl || infxr || infx);  | 
| 18669 | 294  | 
val opt_mixfix = opt_annotation !!! (mfix || binder || infxl || infxr || infx);  | 
295  | 
val opt_mixfix' = opt_annotation I (mfix || binder || infxl || infxr || infx);  | 
|
| 5826 | 296  | 
|
297  | 
||
| 18669 | 298  | 
(* fixes *)  | 
| 5826 | 299  | 
|
| 21400 | 300  | 
val where_ = $$$ "where";  | 
301  | 
||
| 42299 | 302  | 
val const_decl = name -- ($$$ "::" |-- !!! typ) -- opt_mixfix >> triple1;  | 
| 30339 | 303  | 
val const_binding = binding -- ($$$ "::" |-- !!! typ) -- opt_mixfix >> triple1;  | 
| 18669 | 304  | 
|
| 
28081
 
d664b2c1dfe6
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
305  | 
val params = Scan.repeat1 binding -- Scan.option ($$$ "::" |-- !!! typ)  | 
| 18669 | 306  | 
>> (fn (xs, T) => map (rpair T) xs);  | 
307  | 
||
| 
19482
 
9f11af8f7ef9
tuned basic list operators (flat, maps, map_filter);
 
wenzelm 
parents: 
19284 
diff
changeset
 | 
308  | 
val simple_fixes = and_list1 params >> flat;  | 
| 18669 | 309  | 
|
310  | 
val fixes =  | 
|
| 
28081
 
d664b2c1dfe6
explicit type Name.binding for higher-specification elements;
 
wenzelm 
parents: 
28017 
diff
changeset
 | 
311  | 
and_list1 (binding -- Scan.option ($$$ "::" |-- typ) -- mixfix >> (single o triple1) ||  | 
| 
42287
 
d98eb048a2e4
discontinued special treatment of structure Mixfix;
 
wenzelm 
parents: 
40800 
diff
changeset
 | 
312  | 
params >> map (fn (x, y) => (x, y, NoSyn))) >> flat;  | 
| 5826 | 313  | 
|
| 19845 | 314  | 
val for_fixes = Scan.optional ($$$ "for" |-- !!! fixes) [];  | 
| 21371 | 315  | 
val for_simple_fixes = Scan.optional ($$$ "for" |-- !!! simple_fixes) [];  | 
| 19845 | 316  | 
|
| 5826 | 317  | 
|
| 27877 | 318  | 
(* embedded source text *)  | 
| 
27872
 
631371a02b8c
P.doc_source and P.ml_sorce for proper SymbolPos.text;
 
wenzelm 
parents: 
27815 
diff
changeset
 | 
319  | 
|
| 27877 | 320  | 
val ML_source = source_position (group "ML source" text);  | 
| 
27872
 
631371a02b8c
P.doc_source and P.ml_sorce for proper SymbolPos.text;
 
wenzelm 
parents: 
27815 
diff
changeset
 | 
321  | 
val doc_source = source_position (group "document source" text);  | 
| 
 
631371a02b8c
P.doc_source and P.ml_sorce for proper SymbolPos.text;
 
wenzelm 
parents: 
27815 
diff
changeset
 | 
322  | 
|
| 
 
631371a02b8c
P.doc_source and P.ml_sorce for proper SymbolPos.text;
 
wenzelm 
parents: 
27815 
diff
changeset
 | 
323  | 
|
| 5826 | 324  | 
(* terms *)  | 
325  | 
||
| 
40793
 
d21aedaa91e7
added Parse.literal_fact with proper inner_syntax markup (source position);
 
wenzelm 
parents: 
40296 
diff
changeset
 | 
326  | 
val tm = short_ident || long_ident || sym_ident || term_var || number || string;  | 
| 5826 | 327  | 
|
| 
40793
 
d21aedaa91e7
added Parse.literal_fact with proper inner_syntax markup (source position);
 
wenzelm 
parents: 
40296 
diff
changeset
 | 
328  | 
val term_group = group "term" tm;  | 
| 
 
d21aedaa91e7
added Parse.literal_fact with proper inner_syntax markup (source position);
 
wenzelm 
parents: 
40296 
diff
changeset
 | 
329  | 
val prop_group = group "proposition" tm;  | 
| 
27753
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
330  | 
|
| 
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
331  | 
val term = inner_syntax term_group;  | 
| 
 
94b672153b49
sort/typ/term/prop: inner_syntax markup encodes original source position;
 
wenzelm 
parents: 
27737 
diff
changeset
 | 
332  | 
val prop = inner_syntax prop_group;  | 
| 5826 | 333  | 
|
| 
42300
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
334  | 
val type_const = inner_syntax (group "type constructor" xname);  | 
| 
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
335  | 
val const = inner_syntax (group "constant" xname);  | 
| 
 
0d1cbc1fe579
notation: proper markup for type constructor / constant;
 
wenzelm 
parents: 
42299 
diff
changeset
 | 
336  | 
|
| 
40793
 
d21aedaa91e7
added Parse.literal_fact with proper inner_syntax markup (source position);
 
wenzelm 
parents: 
40296 
diff
changeset
 | 
337  | 
val literal_fact = inner_syntax (group "literal fact" alt_string);  | 
| 
 
d21aedaa91e7
added Parse.literal_fact with proper inner_syntax markup (source position);
 
wenzelm 
parents: 
40296 
diff
changeset
 | 
338  | 
|
| 5826 | 339  | 
|
| 6949 | 340  | 
(* patterns *)  | 
| 6935 | 341  | 
|
| 6949 | 342  | 
val is_terms = Scan.repeat1 ($$$ "is" |-- term);  | 
| 6935 | 343  | 
val is_props = Scan.repeat1 ($$$ "is" |-- prop);  | 
344  | 
||
| 19585 | 345  | 
val propp = prop -- Scan.optional ($$$ "(" |-- !!! (is_props --| $$$ ")")) [];
 | 
| 6949 | 346  | 
val termp = term -- Scan.optional ($$$ "(" |-- !!! (is_terms --| $$$ ")")) [];
 | 
| 6935 | 347  | 
|
348  | 
||
| 22119 | 349  | 
(* targets *)  | 
| 19811 | 350  | 
|
| 22119 | 351  | 
val target = ($$$ "(" -- $$$ "in") |-- !!! (xname --| $$$ ")");
 | 
352  | 
val opt_target = Scan.option target;  | 
|
| 12272 | 353  | 
|
354  | 
end;  | 
|
| 30511 | 355  | 
|
| 
36949
 
080e85d46108
renamed structure OuterKeyword to Keyword and OuterParse to Parse, keeping the old names as legacy aliases for some time;
 
wenzelm 
parents: 
35838 
diff
changeset
 | 
356  | 
type 'a parser = 'a Parse.parser;  | 
| 
 
080e85d46108
renamed structure OuterKeyword to Keyword and OuterParse to Parse, keeping the old names as legacy aliases for some time;
 
wenzelm 
parents: 
35838 
diff
changeset
 | 
357  | 
type 'a context_parser = 'a Parse.context_parser;  | 
| 30511 | 358  |