1 (* Title: Pure/Thy/thy_parse.ML
3 Author: Markus Wenzel, TU Muenchen
5 The parser for theory files.
8 infix 5 -- --$$ $$-- ^^;
15 val !! : ('a -> 'b * 'c) -> 'a -> 'b * 'c
16 val >> : ('a -> 'b * 'c) * ('b -> 'd) -> 'a -> 'd * 'c
17 val || : ('a -> 'b) * ('a -> 'b) -> 'a -> 'b
18 val -- : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e
19 val ^^ : ('a -> string * 'b) * ('b -> string * 'c) -> 'a -> string * 'c
20 val $$ : string -> token list -> string * token list
21 val $$-- : string * (token list -> 'b * 'c) -> token list -> 'b * 'c
22 val --$$ : ('a -> 'b * token list) * string -> 'a -> 'b * token list
23 val ident: token list -> string * token list
24 val long_ident: token list -> string * token list
25 val long_id: token list -> string * token list
26 val type_var: token list -> string * token list
27 val type_args: token list -> string list * token list
28 val nat: token list -> string * token list
29 val string: token list -> string * token list
30 val verbatim: token list -> string * token list
31 val empty: 'a -> 'b list * 'a
32 val optional: ('a -> 'b * 'a) -> 'b -> 'a -> 'b * 'a
33 val repeat: ('a -> 'b * 'a) -> 'a -> 'b list * 'a
34 val repeat1: ('a -> 'b * 'a) -> 'a -> 'b list * 'a
35 val enum: string -> (token list -> 'a * token list)
36 -> token list -> 'a list * token list
37 val enum1: string -> (token list -> 'a * token list)
38 -> token list -> 'a list * token list
39 val list: (token list -> 'a * token list)
40 -> token list -> 'a list * token list
41 val list1: (token list -> 'a * token list)
42 -> token list -> 'a list * token list
43 val name: token list -> string * token list
44 val sort: token list -> string * token list
45 val typ: token list -> string * token list
46 val opt_infix: token list -> string * token list
47 val opt_mixfix: token list -> string * token list
48 val opt_witness: token list -> string * token list
49 val const_decls: token list -> string * token list
51 val make_syntax: string list ->
52 (string * (token list -> (string * string) * token list)) list -> syntax
53 val parse_thy: syntax -> string -> string -> string
54 val section: string -> string -> (token list -> string * token list)
55 -> (string * (token list -> (string * string) * token list))
56 val axm_section: string -> string
57 -> (token list -> (string * string list) * token list)
58 -> (string * (token list -> (string * string) * token list))
59 val pure_keywords: string list
61 (string * (token list -> (string * string) * token list)) list
62 (*items for building strings*)
63 val cat: string -> string -> string
64 val parens: string -> string
65 val brackets: string -> string
66 val mk_list: string list -> string
67 val mk_big_list: string list -> string
68 val mk_pair: string * string -> string
69 val mk_triple: string * string * string -> string
70 val mk_triple1: (string * string) * string -> string
71 val mk_triple2: string * (string * string) -> string
72 val strip_quotes: string -> string
76 structure ThyParse : THY_PARSE=
82 (** parser toolbox **)
84 type token = token_kind * string * int;
89 exception SYNTAX_ERROR of string * string * int;
91 fun syn_err s1 s2 n = raise SYNTAX_ERROR (s1, s2, n);
93 fun eof_err () = error "Unexpected end-of-file";
95 (*Similar to Prolog's cut: reports any syntax error instead of backtracking
96 through a superior || *)
97 fun !! parse toks = parse toks
98 handle SYNTAX_ERROR (s1, s2, n) => error ("Syntax error on line " ^
99 string_of_int n ^ ": " ^ s1 ^ " expected and " ^ s2 ^ " was found");
102 (* parser combinators *)
104 fun (parse >> f) toks = apfst f (parse toks);
106 fun (parse1 || parse2) toks =
107 parse1 toks handle SYNTAX_ERROR _ => parse2 toks;
109 fun (parse1 -- parse2) toks =
111 val (x, toks') = parse1 toks;
112 val (y, toks'') = parse2 toks';
117 fun (parse1 ^^ parse2) = parse1 -- parse2 >> op ^;
120 (* generic parsers *)
122 fun $$ a ((k, b, n) :: toks) =
123 if k = Keyword andalso a = b then (a, toks)
124 else syn_err (quote a) (quote b) n
125 | $$ _ [] = eof_err ();
127 fun (a $$-- parse) = $$ a -- parse >> #2;
129 fun (parse --$$ a) = parse -- $$ a >> #1;
132 fun kind k1 ((k2, s, n) :: toks) =
133 if k1 = k2 then (s, toks)
134 else syn_err (name_of_kind k1) (quote s) n
135 | kind _ [] = eof_err ();
137 val ident = kind Ident;
138 val long_ident = kind LongIdent;
139 val long_id = ident || long_ident;
140 val type_var = kind TypeVar >> quote;
142 val string = kind String;
143 val verbatim = kind Verbatim;
146 fun empty toks = ([], toks);
148 fun optional parse def = parse || empty >> K def;
150 fun repeat parse toks = (parse -- repeat parse >> op :: || empty) toks;
151 fun repeat1 parse = parse -- repeat parse >> op ::;
153 fun enum1 sep parse = parse -- repeat (sep $$-- parse) >> op ::;
154 fun enum sep parse = enum1 sep parse || empty;
156 fun list1 parse = enum1 "," parse;
157 fun list parse = enum "," parse;
161 (** theory parsers **)
165 fun cat s1 s2 = s1 ^ " " ^ s2;
167 val parens = enclose "(" ")";
168 val brackets = enclose "[" "]";
170 val mk_list = brackets o commas;
171 val mk_big_list = brackets o space_implode ",\n ";
173 fun mk_pair (x, y) = parens (commas [x, y]);
174 fun mk_triple (x, y, z) = parens (commas [x, y, z]);
175 fun mk_triple1 ((x, y), z) = mk_triple (x, y, z);
176 fun mk_triple2 (x, (y, z)) = mk_triple (x, y, z);
178 fun split_decls l = flat (map (fn (xs, y) => map (rpair y) xs) l);
180 (*Remove the leading and trailing chararacters. Actually called to
181 remove quotation marks.*)
182 fun strip_quotes s = String.substring (s, 1, size s - 2);
187 val name = ident >> quote || string;
188 val names = list name;
189 val names1 = list1 name;
190 val name_list = names >> mk_list;
191 val name_list1 = names1 >> mk_list;
196 fun empty_decl toks = (empty >> K "") toks;
201 val subclass = name -- optional ("<" $$-- !! name_list1) "[]";
203 val class_decls = repeat1 (subclass >> mk_pair) >> mk_big_list;
210 "{" $$-- name_list --$$ "}";
212 val sort_list1 = list1 sort >> mk_list;
215 val arity = optional ("(" $$-- !! (sort_list1 --$$")")) "[]" -- sort;
217 val arity_decls = repeat1 (names1 --$$ "::" -- !! arity)
218 >> (mk_big_list o map mk_triple2 o split_decls);
221 (* mixfix annotations *)
224 "infixl" $$-- !! (nat >> cat "Infixl" || string -- nat >> (cat "InfixlName" o mk_pair));
226 "infixr" $$-- !! (nat >> cat "Infixr" || string -- nat >> (cat "InfixrName" o mk_pair));
228 val binder = "binder" $$--
229 !! (string -- (("[" $$-- nat --$$ "]") -- nat || nat >> (fn n => (n, n))))
230 >> (cat "Binder" o mk_triple2);
232 val opt_pris = optional ("[" $$-- !! (list nat --$$ "]")) [] >> mk_list;
234 val mixfix = string -- !! (opt_pris -- optional nat "Syntax.max_pri")
235 >> (cat "Mixfix" o mk_triple2);
237 fun opt_syn fx = optional ("(" $$-- fx --$$ ")") "NoSyn";
239 val opt_infix = opt_syn (infxl || infxr);
240 val opt_mixfix = opt_syn (mixfix || infxl || infxr || binder);
247 (*Parse an identifier, but only if it is not followed by "::", "=" or ",";
248 the exclusion of a postfix comma can be controlled to allow expressions
249 like "(id, id)" but disallow ones like "'a => id id,id :: ..."*)
250 fun ident_no_colon _ [] = eof_err()
251 | ident_no_colon allow_comma ((Ident, s, n) :: (rest as (Keyword, s2, n2) ::
253 if s2 = "::" orelse s2 = "=" orelse (not allow_comma andalso s2 = ",")
254 then syn_err (name_of_kind Ident) (quote s2) n2
256 | ident_no_colon _ ((Ident, s, n) :: toks) = (s, toks)
257 | ident_no_colon _ ((k, s, n) :: _) =
258 syn_err (name_of_kind Ident) (quote s) n;
260 (*type used in types, consts and syntax sections*)
261 fun const_type allow_comma toks =
264 (ident || kind TypeVar ^^ optional ($$ "::" ^^ ident) "") --
265 repeat (ident_no_colon allow_comma)
266 >> (fn (args, ts) => cat args (space_implode " " ts)) ||
267 ("(" $$-- (list1 (const_type true)) --$$ ")" >> (parens o commas)) --
268 repeat1 (ident_no_colon allow_comma)
269 >> (fn (args, ts) => cat args (space_implode " " ts));
272 simple_type || "(" $$-- const_type true --$$ ")" >> parens ||
273 "[" $$-- (list1 (const_type true)) --$$ "]" --$$ "=>" --
274 const_type allow_comma >>
275 (fn (src, dest) => mk_list src ^ " => " ^ dest);
276 in ("[" $$-- (list1 (const_type true)) --$$ "]" --$$ "=>" --
277 const_type allow_comma >>
278 (fn (src, dest) => mk_list src ^ " => " ^ dest) ||
279 repeat1 (appl_param --$$ "=>") -- const_type allow_comma >>
280 (fn (src, dest) => space_implode " => " (src@[dest])) ||
282 "(" $$-- const_type true --$$ ")" >> parens) toks
285 val typ = string || (const_type false >> quote);
288 fun mk_old_type_decl ((ts, n), syn) =
289 map (fn t => (mk_triple (t, n, syn), false)) ts;
291 fun mk_type_decl (((xs, t), None), syn) =
292 [(mk_triple (t, string_of_int (length xs), syn), false)]
293 | mk_type_decl (((xs, t), Some rhs), syn) =
294 [(parens (commas [t, mk_list xs, rhs, syn]), true)];
296 fun mk_type_decls tys =
297 "|> Theory.add_types\n" ^ mk_big_list (keyfilter tys false) ^ "\n\n\
298 \|> Theory.add_tyabbrs\n" ^ mk_big_list (keyfilter tys true);
301 val old_type_decl = names1 -- nat -- opt_infix >> mk_old_type_decl;
304 type_var >> (fn x => [x]) ||
305 "(" $$-- !! (list1 type_var --$$ ")") ||
308 val type_decl = type_args -- name --
309 optional ("=" $$-- typ >> Some) None -- opt_infix >> mk_type_decl;
312 repeat1 (old_type_decl || type_decl) >> (mk_type_decls o flat);
318 repeat1 (names1 --$$ "::" -- !! (typ -- opt_mixfix))
319 >> (mk_big_list o map mk_triple2 o split_decls);
323 ("(" $$-- !! (name -- optional ($$ "output" >> K "false") "true" --$$ ")"))
327 val syntax_decls = opt_mode -- const_decls >> (fn (mode, txt) => mode ^ "\n" ^ txt);
333 optional ("(" $$-- !! (name --$$ ")")) "\"logic\"" -- string >> mk_pair;
336 $$ "=>" >> K "Syntax.ParseRule " ||
337 $$ "<=" >> K "Syntax.PrintRule " ||
338 $$ "==" >> K "Syntax.ParsePrintRule ";
341 trans_pat -- !! (trans_arrow -- trans_pat)
342 >> (fn (left, (arr, right)) => arr ^ mk_pair (left, right));
344 val trans_decls = repeat1 trans_line >> mk_big_list;
347 (* ML translations *)
350 " val parse_ast_translation = [];\n\
351 \ val parse_translation = [];\n\
352 \ val print_translation = [];\n\
353 \ val typed_print_translation = [];\n\
354 \ val print_ast_translation = [];\n\
355 \ val token_translation = [];";
358 "(parse_ast_translation, parse_translation, \
359 \print_translation, print_ast_translation)";
364 val mk_axms = mk_big_list o map (mk_pair o apfst quote);
366 fun mk_axiom_decls axms = (mk_axms axms, map fst axms);
368 val axiom_decls = repeat1 (ident -- !! string) >> mk_axiom_decls;
373 val oracle_decl = (name --$$ "=" -- long_id) >> mk_pair;
376 (* combined consts and axioms *)
378 fun mk_constaxiom_decls x =
380 val (axms_defs, axms_names) =
381 mk_axiom_decls (map (fn ((id, _), def) => (id ^ "_def", def)) x);
382 in ((mk_big_list o map mk_triple2 o map (apfst quote o fst)) x ^
383 "\n\n|> (PureThy.add_defs o map Thm.no_attributes)\n" ^ axms_defs, axms_names)
386 val constaxiom_decls =
387 repeat1 (ident --$$ "::" -- !! (typ -- opt_mixfix) -- !! string)
388 >> mk_constaxiom_decls;
393 fun mk_axclass_decl ((c, cs), axms) =
394 (mk_pair (c, cs) ^ "\n" ^ mk_axms axms,
395 (strip_quotes c ^ "I") :: map fst axms);
397 val axclass_decl = subclass -- repeat (ident -- !! string) >> mk_axclass_decl;
402 fun mk_witness (axths, opt_tac) =
403 mk_list (keyfilter axths false) ^ "\n" ^
404 mk_list (keyfilter axths true) ^ "\n" ^
408 string >> rpair false ||
409 long_id >> rpair true;
413 optional ("(" $$-- list1 axm_or_thm --$$ ")") [] --
414 optional (verbatim >> (parens o cat "Some" o parens)) "None"
418 (name --$$ "<" -- name >> (pair "|> AxClass.add_inst_subclass" o mk_pair) ||
419 name --$$ "::" -- arity >> (pair "|> AxClass.add_inst_arity" o mk_triple2))
421 >> (fn ((x, y), z) => (cat_lines [x, y, z]));
428 (optional ((ident >> (fn x => parens ("Some" ^ quote x))) --$$ "+") ("None")) --
430 (repeat (name --$$ "::" -- !! (typ -- opt_mixfix))
431 >> (mk_big_list o map mk_triple2))) --
433 ("assumes" $$-- (repeat ((ident >> quote) -- !! string)
434 >> (mk_list o map mk_pair)))
437 ("defines" $$-- (repeat ((ident >> quote) -- !! string)
438 >> (mk_list o map mk_pair)))
440 >> (fn ((((nm, ext), cs), asms), defs) => cat_lines [nm, ext, cs, asms, defs]);
444 (** theory syntax **)
447 Scan.lexicon * (token list -> (string * string) * token list) Symtab.table;
449 fun make_syntax keywords sects =
451 val dups = duplicates (map fst sects);
452 val sects' = gen_distinct eq_fst sects;
453 val keys = map Symbol.explode (map fst sects' @ keywords);
456 else warning ("Duplicate declaration of theory file sections:\n" ^ commas_quote dups);
457 (Scan.make_lexicon keys, Symtab.make sects')
463 fun mk_header (thy_name, bases) =
464 (thy_name, "mk_base " ^ mk_list bases ^ " " ^ quote thy_name);
467 ident >> (cat "Thy" o quote) ||
468 string >> cat "File";
470 val header = ident --$$ "=" -- enum1 "+" base >> mk_header;
475 fun mk_extension (txts, mltxt) =
477 val cat_sects = space_implode "\n\n" o filter_out (equal "");
478 val (extxts, postxts) = split_list txts;
480 (cat_sects extxts, cat_sects postxts, mltxt)
483 fun sect tab ((Keyword, s, n) :: toks) =
484 (case Symtab.lookup (tab, s) of
485 Some parse => !! parse toks
486 | None => syn_err "section" s n)
487 | sect _ ((_, s, n) :: _) = syn_err "section" s n
488 | sect _ [] = eof_err ();
490 fun extension sectab = "+" $$-- !!
491 (repeat (sect sectab) --$$ "end" -- optional verbatim "")
494 fun opt_extension sectab = optional (extension sectab) ("", "", "");
497 (* theory definition *)
499 fun mk_structure tname ((thy_name, old_thys), (extxt, postxt, mltxt)) =
500 if thy_name <> tname then
501 error ("Filename \"" ^ tname ^ ".thy\" and theory name "
502 ^ quote thy_name ^ " are different")
504 "val thy = " ^ old_thys ^ ";\n\n\
505 \structure " ^ thy_name ^ " =\n\
515 \|> PureThy.put_name " ^ quote thy_name ^ "\n\
516 \|> PureThy.local_path\n\
517 \|> Theory.add_trfuns\n"
519 \|> Theory.add_trfunsT typed_print_translation\n\
520 \|> Theory.add_tokentrfuns token_translation\n\
524 \|> PureThy.end_theory\n\
526 \val _ = store_theory thy;\n\
527 \val _ = context thy;\n\
535 \open " ^ thy_name ^ ";\n\
538 fun theory_defn sectab tname =
539 header -- opt_extension sectab -- eof >> (mk_structure tname o #1);
541 fun parse_thy (lex, sectab) tname str =
542 #1 (!! (theory_defn sectab tname) (tokenize lex str));
545 (* standard sections *)
547 fun mk_val ax = "val " ^ ax ^ " = get_axiom thy " ^ quote ax ^ ";";
548 val mk_vals = cat_lines o map mk_val;
550 fun mk_axm_sect "" (txt, axs) = (txt, mk_vals axs)
551 | mk_axm_sect pretxt (txt, axs) = (pretxt ^ "\n" ^ txt, mk_vals axs);
553 fun axm_section name pretxt parse =
554 (name, parse >> mk_axm_sect pretxt);
556 fun section name pretxt parse =
557 axm_section name pretxt (parse >> rpair []);
561 ["end", "ML", "mixfix", "infixr", "infixl", "binder", "output", "=",
562 "+", ",", "<", "{", "}", "(", ")", "[", "]", "::", "==", "=>",
563 "<=", "fixes", "assumes", "defines"];
566 [section "classes" "|> Theory.add_classes" class_decls,
567 section "default" "|> Theory.add_defsort" sort,
568 section "types" "" type_decls,
569 section "nonterminals" "|> Theory.add_nonterminals" (repeat1 name >> mk_list),
570 section "arities" "|> Theory.add_arities" arity_decls,
571 section "consts" "|> Theory.add_consts" const_decls,
572 section "syntax" "|> Theory.add_modesyntax" syntax_decls,
573 section "translations" "|> Theory.add_trrules" trans_decls,
574 axm_section "rules" "|> (PureThy.add_axioms o map Thm.no_attributes)" axiom_decls,
575 axm_section "defs" "|> (PureThy.add_defs o map Thm.no_attributes)" axiom_decls,
576 section "oracle" "|> Theory.add_oracle" oracle_decl,
577 axm_section "constdefs" "|> Theory.add_consts" constaxiom_decls,
578 axm_section "axclass" "|> AxClass.add_axclass" axclass_decl,
579 section "instance" "" instance_decl,
580 section "path" "|> Theory.add_path" name,
581 section "global" "|> PureThy.global_path" empty_decl,
582 section "local" "|> PureThy.local_path" empty_decl,
583 section "setup" "|> Library.apply" long_id,
584 section "MLtext" "" verbatim,
585 section "locale" "|> Locale.add_locale" locale_decl];