author | wenzelm |
Wed, 13 Apr 2005 18:47:01 +0200 | |
changeset 15714 | 9b8da47715c3 |
parent 15703 | 727ef1b8b3ee |
child 15963 | 5b70f789e079 |
permissions | -rw-r--r-- |
5826 | 1 |
(* Title: Pure/Isar/outer_parse.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
Generic parsers for Isabelle/Isar outer syntax. |
|
6 |
*) |
|
7 |
||
8 |
signature OUTER_PARSE = |
|
9 |
sig |
|
10 |
type token |
|
11 |
val group: string -> (token list -> 'a) -> token list -> 'a |
|
12 |
val !!! : (token list -> 'a) -> token list -> 'a |
|
8581
5c7ed2af8bfb
!!!! = cut "Corrupted outer syntax in presentation";
wenzelm
parents:
8350
diff
changeset
|
13 |
val !!!! : (token list -> 'a) -> token 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 |
|
15703 | 17 |
val not_eof: token list -> token * token list |
5826 | 18 |
val position: (token list -> 'a * 'b) -> token list -> ('a * Position.T) * 'b |
7026 | 19 |
val command: token list -> string * token list |
5826 | 20 |
val keyword: token list -> string * token list |
21 |
val short_ident: token list -> string * token list |
|
22 |
val long_ident: token list -> string * token list |
|
23 |
val sym_ident: token list -> string * token list |
|
24 |
val term_var: token list -> string * token list |
|
25 |
val type_ident: token list -> string * token list |
|
26 |
val type_var: token list -> string * token list |
|
27 |
val number: token list -> string * token list |
|
28 |
val string: token list -> string * token list |
|
29 |
val verbatim: token list -> string * token list |
|
6860 | 30 |
val sync: token list -> string * token list |
5826 | 31 |
val eof: token list -> string * token list |
15703 | 32 |
val $$$ : string -> token list -> string * token list |
33 |
val semicolon: token list -> string * token list |
|
34 |
val underscore: token list -> string * token list |
|
35 |
val maybe: (token list -> 'a * token list) -> token list -> 'a option * token list |
|
7930 | 36 |
val opt_unit: token list -> unit * token list |
14646 | 37 |
val opt_keyword: string -> token list -> bool * token list |
5826 | 38 |
val nat: token list -> int * token list |
39 |
val enum: string -> (token list -> 'a * token list) -> token list -> 'a list * token list |
|
40 |
val enum1: string -> (token list -> 'a * token list) -> token list -> 'a list * token list |
|
41 |
val list: (token list -> 'a * token list) -> token list -> 'a list * token list |
|
42 |
val list1: (token list -> 'a * token list) -> token list -> 'a list * token list |
|
6013 | 43 |
val and_list: (token list -> 'a * token list) -> token list -> 'a list * token list |
44 |
val and_list1: (token list -> 'a * token list) -> token list -> 'a list * token list |
|
5826 | 45 |
val name: token list -> bstring * token list |
46 |
val xname: token list -> xstring * token list |
|
47 |
val text: token list -> string * token list |
|
14949 | 48 |
val path: token list -> Path.T * token list |
8897 | 49 |
val sort: token list -> string * token list |
50 |
val arity: token list -> (string list * string) * token list |
|
5826 | 51 |
val type_args: token list -> string list * token list |
52 |
val typ: token list -> string * token list |
|
53 |
val opt_infix: token list -> Syntax.mixfix * token list |
|
54 |
val opt_mixfix: token list -> Syntax.mixfix * token list |
|
9037 | 55 |
val opt_infix': token list -> Syntax.mixfix * token list |
56 |
val opt_mixfix': token list -> Syntax.mixfix * token list |
|
14646 | 57 |
val mixfix': token list -> Syntax.mixfix * token list |
5826 | 58 |
val const: token list -> (bstring * string * Syntax.mixfix) * token list |
59 |
val term: token list -> string * token list |
|
60 |
val prop: token list -> string * token list |
|
6935 | 61 |
val propp: token list -> (string * (string list * string list)) * token list |
6949 | 62 |
val termp: token list -> (string * string list) * token list |
9131 | 63 |
val arguments: token list -> Args.T list * token list |
15703 | 64 |
val attribs: token list -> Attrib.src list * token list |
65 |
val opt_attribs: token list -> Attrib.src list * token list |
|
66 |
val thm_name: string -> token list -> (bstring * Attrib.src list) * token list |
|
67 |
val opt_thm_name: string -> token list -> (bstring * Attrib.src list) * token list |
|
68 |
val spec_name: token list -> ((bstring * string) * Attrib.src list) * token list |
|
69 |
val spec_opt_name: token list -> ((bstring * string) * Attrib.src list) * token list |
|
70 |
val xthm: token list -> (thmref * Attrib.src list) * token list |
|
71 |
val xthms1: token list -> (thmref * Attrib.src list) list * token list |
|
12942 | 72 |
val locale_target: token list -> xstring option * token list |
12272 | 73 |
val locale_expr: token list -> Locale.expr * token list |
12942 | 74 |
val locale_keyword: token list -> string * token list |
15703 | 75 |
val locale_element: token list -> Locale.element * token list |
76 |
val locale_elem_or_expr: token list -> Locale.element Locale.elem_expr * token list |
|
5826 | 77 |
val method: token list -> Method.text * token list |
78 |
end; |
|
79 |
||
80 |
structure OuterParse: OUTER_PARSE = |
|
81 |
struct |
|
82 |
||
9131 | 83 |
structure T = OuterLex; |
84 |
type token = T.token; |
|
5826 | 85 |
|
86 |
||
87 |
(** error handling **) |
|
88 |
||
89 |
(* group atomic parsers (no cuts!) *) |
|
90 |
||
91 |
fun fail_with s = Scan.fail_with |
|
92 |
(fn [] => s ^ " expected (past end-of-file!)" |
|
9155 | 93 |
| (tok :: _) => s ^ " expected,\nbut " ^ T.name_of tok ^ T.pos_of tok ^ " was found"); |
5826 | 94 |
|
95 |
fun group s scan = scan || fail_with s; |
|
96 |
||
97 |
||
5877 | 98 |
(* cut *) |
5826 | 99 |
|
8581
5c7ed2af8bfb
!!!! = cut "Corrupted outer syntax in presentation";
wenzelm
parents:
8350
diff
changeset
|
100 |
fun cut kind scan = |
5826 | 101 |
let |
102 |
fun get_pos [] = " (past end-of-file!)" |
|
9131 | 103 |
| get_pos (tok :: _) = T.pos_of tok; |
5826 | 104 |
|
15531 | 105 |
fun err (toks, NONE) = kind ^ get_pos toks |
106 |
| err (toks, SOME msg) = kind ^ get_pos toks ^ ": " ^ msg; |
|
5826 | 107 |
in Scan.!! err scan end; |
108 |
||
8586 | 109 |
fun !!! scan = cut "Outer syntax error" scan; |
110 |
fun !!!! scan = cut "Corrupted outer syntax in presentation" scan; |
|
8581
5c7ed2af8bfb
!!!! = cut "Corrupted outer syntax in presentation";
wenzelm
parents:
8350
diff
changeset
|
111 |
|
5826 | 112 |
|
113 |
||
114 |
(** basic parsers **) |
|
115 |
||
116 |
(* utils *) |
|
117 |
||
118 |
fun triple1 ((x, y), z) = (x, y, z); |
|
119 |
fun triple2 (x, (y, z)) = (x, y, z); |
|
6430 | 120 |
fun triple_swap ((x, y), z) = ((x, z), y); |
5826 | 121 |
|
122 |
||
123 |
(* tokens *) |
|
124 |
||
15703 | 125 |
val not_eof = Scan.one T.not_eof; |
126 |
||
127 |
fun position scan = (Scan.ahead not_eof >> T.position_of) -- scan >> Library.swap; |
|
5826 | 128 |
|
129 |
fun kind k = |
|
9131 | 130 |
group (T.str_of_kind k) (Scan.one (T.is_kind k) >> T.val_of); |
5826 | 131 |
|
9131 | 132 |
val command = kind T.Command; |
133 |
val keyword = kind T.Keyword; |
|
134 |
val short_ident = kind T.Ident; |
|
135 |
val long_ident = kind T.LongIdent; |
|
136 |
val sym_ident = kind T.SymIdent; |
|
137 |
val term_var = kind T.Var; |
|
138 |
val type_ident = kind T.TypeIdent; |
|
139 |
val type_var = kind T.TypeVar; |
|
140 |
val number = kind T.Nat; |
|
141 |
val string = kind T.String; |
|
142 |
val verbatim = kind T.Verbatim; |
|
143 |
val sync = kind T.Sync; |
|
144 |
val eof = kind T.EOF; |
|
5826 | 145 |
|
146 |
fun $$$ x = |
|
9131 | 147 |
group (T.str_of_kind T.Keyword ^ " " ^ quote x) |
148 |
(Scan.one (T.keyword_with (equal x)) >> T.val_of); |
|
149 |
||
150 |
val semicolon = $$$ ";"; |
|
5826 | 151 |
|
15703 | 152 |
val minus = sym_ident :-- (fn "-" => Scan.succeed () | _ => Scan.fail) >> #1; |
11792
311eee3d63b6
parser for underscore (actually a symbolic identifier!);
wenzelm
parents:
11651
diff
changeset
|
153 |
val underscore = sym_ident :-- (fn "_" => Scan.succeed () | _ => Scan.fail) >> #1; |
15703 | 154 |
fun maybe scan = underscore >> K NONE || scan >> SOME; |
11792
311eee3d63b6
parser for underscore (actually a symbolic identifier!);
wenzelm
parents:
11651
diff
changeset
|
155 |
|
14835 | 156 |
val nat = number >> (#1 o Library.read_int o Symbol.explode); |
5826 | 157 |
|
7930 | 158 |
val opt_unit = Scan.optional ($$$ "(" -- $$$ ")" >> (K ())) (); |
159 |
||
14646 | 160 |
fun opt_keyword s = Scan.optional ($$$ "(" |-- !!! (($$$ s >> K true) --| $$$ ")")) false; |
161 |
||
5826 | 162 |
|
163 |
(* enumerations *) |
|
164 |
||
6003 | 165 |
fun enum1 sep scan = scan -- Scan.repeat ($$$ sep |-- !!! scan) >> op ::; |
5826 | 166 |
fun enum sep scan = enum1 sep scan || Scan.succeed []; |
167 |
||
168 |
fun list1 scan = enum1 "," scan; |
|
169 |
fun list scan = enum "," scan; |
|
170 |
||
6013 | 171 |
fun and_list1 scan = enum1 "and" scan; |
172 |
fun and_list scan = enum "and" scan; |
|
173 |
||
5826 | 174 |
|
5960 | 175 |
(* names and text *) |
5826 | 176 |
|
8146 | 177 |
val name = group "name declaration" (short_ident || sym_ident || string || number); |
178 |
val xname = group "name reference" (short_ident || long_ident || sym_ident || string || number); |
|
179 |
val text = group "text" (short_ident || long_ident || sym_ident || string || number || verbatim); |
|
14949 | 180 |
val path = group "file name/path specification" name >> Path.unpack; |
6553 | 181 |
|
182 |
||
6372 | 183 |
(* sorts and arities *) |
5826 | 184 |
|
8897 | 185 |
val sort = group "sort" xname; |
5826 | 186 |
|
6372 | 187 |
fun gen_arity cod = |
7352 | 188 |
Scan.optional ($$$ "(" |-- !!! (list1 sort --| $$$ ")")) [] -- cod; |
6372 | 189 |
|
190 |
val arity = gen_arity sort; |
|
5826 | 191 |
|
192 |
||
193 |
(* types *) |
|
194 |
||
8146 | 195 |
val typ = group "type" |
196 |
(short_ident || long_ident || sym_ident || type_ident || type_var || string || number); |
|
5826 | 197 |
|
198 |
val type_args = |
|
199 |
type_ident >> single || |
|
200 |
$$$ "(" |-- !!! (list1 type_ident --| $$$ ")") || |
|
201 |
Scan.succeed []; |
|
202 |
||
203 |
||
204 |
(* mixfix annotations *) |
|
205 |
||
11651 | 206 |
val infx = $$$ "infix" |-- !!! (nat >> Syntax.Infix || string -- nat >> Syntax.InfixName); |
5826 | 207 |
val infxl = $$$ "infixl" |-- !!! (nat >> Syntax.Infixl || string -- nat >> Syntax.InfixlName); |
208 |
val infxr = $$$ "infixr" |-- !!! (nat >> Syntax.Infixr || string -- nat >> Syntax.InfixrName); |
|
209 |
||
210 |
val binder = |
|
211 |
$$$ "binder" |-- |
|
212 |
!!! (string -- ($$$ "[" |-- nat --| $$$ "]" -- nat || nat >> (fn n => (n, n)))) |
|
213 |
>> (Syntax.Binder o triple2); |
|
214 |
||
215 |
||
216 |
val opt_pris = Scan.optional ($$$ "[" |-- !!! (list nat --| $$$ "]")) []; |
|
217 |
||
218 |
val mixfix = |
|
8648 | 219 |
(string -- !!! (opt_pris -- Scan.optional nat Syntax.max_pri)) |
220 |
>> (Syntax.Mixfix o triple2); |
|
5826 | 221 |
|
14646 | 222 |
fun fix_decl guard fix = $$$ "(" |-- guard (fix --| $$$ ")"); |
223 |
fun opt_fix guard fix = Scan.optional (fix_decl guard fix) Syntax.NoSyn; |
|
5826 | 224 |
|
11651 | 225 |
val opt_infix = opt_fix !!! (infxl || infxr || infx); |
226 |
val opt_mixfix = opt_fix !!! (mixfix || binder || infxl || infxr || infx); |
|
227 |
val opt_infix' = opt_fix I (infxl || infxr || infx); |
|
228 |
val opt_mixfix' = opt_fix I (mixfix || binder || infxl || infxr || infx); |
|
14646 | 229 |
val mixfix' = fix_decl !!! (mixfix || binder || infxl || infxr || infx); |
5826 | 230 |
|
231 |
||
232 |
(* consts *) |
|
233 |
||
234 |
val const = |
|
235 |
name -- ($$$ "::" |-- !!! (typ -- opt_mixfix)) >> triple2; |
|
236 |
||
237 |
||
238 |
(* terms *) |
|
239 |
||
7477 | 240 |
val trm = short_ident || long_ident || sym_ident || term_var || number || string; |
5826 | 241 |
|
242 |
val term = group "term" trm; |
|
243 |
val prop = group "proposition" trm; |
|
244 |
||
245 |
||
6949 | 246 |
(* patterns *) |
6935 | 247 |
|
6949 | 248 |
val is_terms = Scan.repeat1 ($$$ "is" |-- term); |
6935 | 249 |
val is_props = Scan.repeat1 ($$$ "is" |-- prop); |
250 |
val concl_props = $$$ "concl" |-- !!! is_props; |
|
7418 | 251 |
val any_props = concl_props >> pair [] || is_props -- Scan.optional concl_props []; |
6935 | 252 |
|
253 |
val propp = prop -- Scan.optional ($$$ "(" |-- !!! (any_props --| $$$ ")")) ([], []); |
|
12047 | 254 |
val propp' = prop -- Scan.optional ($$$ "(" |-- !!! (is_props --| $$$ ")")) []; |
6949 | 255 |
val termp = term -- Scan.optional ($$$ "(" |-- !!! (is_terms --| $$$ ")")) []; |
6935 | 256 |
|
257 |
||
5826 | 258 |
(* arguments *) |
259 |
||
9131 | 260 |
fun keyword_symid is_symid = Scan.one (T.keyword_with is_symid) >> T.val_of; |
261 |
val keyword_sid = keyword_symid T.is_sid; |
|
5826 | 262 |
|
6983 | 263 |
fun atom_arg is_symid blk = |
5826 | 264 |
group "argument" |
7477 | 265 |
(position (short_ident || long_ident || sym_ident || term_var || |
15703 | 266 |
type_ident || type_var || number) >> Args.mk_ident || |
267 |
position (keyword_symid is_symid) >> Args.mk_keyword || |
|
268 |
position (string || verbatim) >> Args.mk_string || |
|
269 |
position (if blk then $$$ "," else Scan.fail) >> Args.mk_keyword); |
|
5826 | 270 |
|
5877 | 271 |
fun paren_args l r scan = position ($$$ l) -- !!! (scan true -- position ($$$ r)) |
15703 | 272 |
>> (fn (x, (ys, z)) => Args.mk_keyword x :: ys @ [Args.mk_keyword z]); |
5826 | 273 |
|
6983 | 274 |
fun args is_symid blk x = Scan.optional (args1 is_symid blk) [] x |
275 |
and args1 is_symid blk x = |
|
5826 | 276 |
((Scan.repeat1 |
6983 | 277 |
(Scan.repeat1 (atom_arg is_symid blk) || |
278 |
paren_args "(" ")" (args is_symid) || |
|
15570 | 279 |
paren_args "[" "]" (args is_symid))) >> List.concat) x; |
5826 | 280 |
|
9131 | 281 |
val arguments = args T.is_sid false; |
282 |
||
5826 | 283 |
|
6372 | 284 |
(* theorem specifications *) |
5826 | 285 |
|
9131 | 286 |
val attrib = position ((keyword_sid || xname) -- !!! arguments) >> Args.src; |
5826 | 287 |
val attribs = $$$ "[" |-- !!! (list attrib --| $$$ "]"); |
288 |
val opt_attribs = Scan.optional attribs []; |
|
289 |
||
6398 | 290 |
fun thm_name s = name -- opt_attribs --| $$$ s; |
6511 | 291 |
fun opt_thm_name s = |
292 |
Scan.optional ((name -- opt_attribs || (attribs >> pair "")) --| $$$ s) ("", []);; |
|
5917 | 293 |
|
6372 | 294 |
val spec_name = thm_name ":" -- prop >> (fn ((x, y), z) => ((x, z), y)); |
295 |
val spec_opt_name = opt_thm_name ":" -- prop >> (fn ((x, y), z) => ((x, z), y)); |
|
296 |
||
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15206
diff
changeset
|
297 |
val thm_sel = $$$ "(" |-- list1 |
15703 | 298 |
(nat --| minus -- nat >> PureThy.FromTo || |
299 |
nat --| minus >> PureThy.From || |
|
300 |
nat >> PureThy.Single) --| $$$ ")"; |
|
15456
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15206
diff
changeset
|
301 |
|
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
berghofe
parents:
15206
diff
changeset
|
302 |
val xthm = xname -- Scan.option thm_sel -- opt_attribs; |
6372 | 303 |
val xthms1 = Scan.repeat1 xthm; |
5826 | 304 |
|
305 |
||
12047 | 306 |
(* locale elements *) |
307 |
||
12272 | 308 |
local |
309 |
||
15531 | 310 |
val loc_mixfix = $$$ "(" -- $$$ "structure" -- !!! ($$$ ")") >> K NONE || opt_mixfix >> SOME; |
12955 | 311 |
val loc_keyword = $$$ "fixes" || $$$ "assumes" || $$$ "defines" || $$$ "notes" || $$$ "includes"; |
12272 | 312 |
|
15127 | 313 |
val loc_element = |
314 |
$$$ "fixes" |-- !!! (and_list1 (name -- Scan.option ($$$ "::" |-- typ) -- loc_mixfix |
|
315 |
>> triple1)) >> Locale.Fixes || |
|
316 |
$$$ "assumes" |-- !!! (and_list1 (opt_thm_name ":" -- Scan.repeat1 propp)) |
|
317 |
>> Locale.Assumes || |
|
318 |
$$$ "defines" |-- !!! (and_list1 (opt_thm_name ":" -- propp')) |
|
319 |
>> Locale.Defines || |
|
320 |
$$$ "notes" |-- !!! (and_list1 (opt_thm_name "=" -- xthms1)) |
|
321 |
>> Locale.Notes; |
|
322 |
||
12272 | 323 |
fun plus1 scan = |
324 |
scan -- Scan.repeat ($$$ "+" |-- Scan.unless loc_keyword (!!! scan)) >> op ::; |
|
325 |
||
12268 | 326 |
fun expr2 x = (xname >> Locale.Locale || $$$ "(" |-- !!! (expr0 --| $$$ ")")) x |
15703 | 327 |
and expr1 x = (expr2 -- Scan.repeat1 (maybe name) >> Locale.Rename || expr2) x |
12272 | 328 |
and expr0 x = (plus1 expr1 >> (fn [e] => e | es => Locale.Merge es)) x; |
12268 | 329 |
|
12272 | 330 |
in |
331 |
||
12942 | 332 |
val locale_target = Scan.option (($$$ "(" -- $$$ "in") |-- !!! (xname --| $$$ ")")); |
12272 | 333 |
val locale_expr = expr0; |
12942 | 334 |
val locale_keyword = loc_keyword; |
12047 | 335 |
|
15127 | 336 |
val locale_element = group "locale element" loc_element; |
337 |
||
338 |
val locale_elem_or_expr = group "locale element or includes" |
|
339 |
(loc_element >> Locale.Elem || |
|
12955 | 340 |
$$$ "includes" |-- !!! locale_expr >> Locale.Expr); |
12272 | 341 |
|
342 |
end; |
|
12047 | 343 |
|
344 |
||
5826 | 345 |
(* proof methods *) |
346 |
||
6983 | 347 |
fun is_symid_meth s = |
9131 | 348 |
s <> "|" andalso s <> "?" andalso s <> "+" andalso T.is_sid s; |
6983 | 349 |
|
5826 | 350 |
fun meth4 x = |
5877 | 351 |
(position (xname >> rpair []) >> (Method.Source o Args.src) || |
6983 | 352 |
$$$ "(" |-- !!! (meth0 --| $$$ ")")) x |
5826 | 353 |
and meth3 x = |
354 |
(meth4 --| $$$ "?" >> Method.Try || |
|
355 |
meth4 --| $$$ "+" >> Method.Repeat1 || |
|
5940 | 356 |
meth4) x |
357 |
and meth2 x = |
|
6983 | 358 |
(position (xname -- args1 is_symid_meth false) >> (Method.Source o Args.src) || |
5826 | 359 |
meth3) x |
360 |
and meth1 x = (enum1 "," meth2 >> (fn [m] => m | ms => Method.Then ms)) x |
|
361 |
and meth0 x = (enum1 "|" meth1 >> (fn [m] => m | ms => Method.Orelse ms)) x; |
|
362 |
||
6558 | 363 |
val method = meth3; |
5826 | 364 |
|
365 |
||
366 |
end; |
|
15714 | 367 |