author | paulson |
Mon, 28 Dec 1998 16:50:37 +0100 | |
changeset 6043 | 3eecc7fbfad8 |
parent 6022 | 259e4f2114e1 |
child 6090 | 78c068b838ff |
permissions | -rw-r--r-- |
389 | 1 |
(* Title: Pure/Thy/thy_parse.ML |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel, TU Muenchen |
|
4 |
||
5 |
The parser for theory files. |
|
6 |
*) |
|
7 |
||
8 |
infix 5 -- --$$ $$-- ^^; |
|
9 |
infix 3 >>; |
|
10 |
infix 0 ||; |
|
11 |
||
12 |
signature THY_PARSE = |
|
3977 | 13 |
sig |
389 | 14 |
type token |
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 |
|
636 | 27 |
val type_args: token list -> string list * token list |
389 | 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 |
|
3977 | 45 |
val typ: token list -> string * token list |
451 | 46 |
val opt_infix: token list -> string * token list |
47 |
val opt_mixfix: token list -> string * token list |
|
636 | 48 |
val opt_witness: token list -> string * token list |
4099 | 49 |
val const_decls: token list -> string * token list |
389 | 50 |
type syntax |
51 |
val make_syntax: string list -> |
|
52 |
(string * (token list -> (string * string) * token list)) list -> syntax |
|
476
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
53 |
val parse_thy: syntax -> string -> string -> string |
389 | 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 |
|
60 |
val pure_sections: |
|
61 |
(string * (token list -> (string * string) * token list)) list |
|
570
6333c181a3f7
Pure/Thy/thy_parse/THY_PARSE: deleted duplicate specifications of parens,
lcp
parents:
558
diff
changeset
|
62 |
(*items for building strings*) |
710 | 63 |
val cat: string -> string -> string |
656 | 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 |
|
5058 | 70 |
val mk_triple1: (string * string) * string -> string |
71 |
val mk_triple2: string * (string * string) -> string |
|
656 | 72 |
val strip_quotes: string -> string |
3977 | 73 |
end; |
389 | 74 |
|
1512 | 75 |
|
76 |
structure ThyParse : THY_PARSE= |
|
389 | 77 |
struct |
78 |
||
79 |
open ThyScan; |
|
80 |
||
81 |
||
82 |
(** parser toolbox **) |
|
83 |
||
84 |
type token = token_kind * string * int; |
|
85 |
||
86 |
||
87 |
(* errors *) |
|
88 |
||
89 |
exception SYNTAX_ERROR of string * string * int; |
|
90 |
||
91 |
fun syn_err s1 s2 n = raise SYNTAX_ERROR (s1, s2, n); |
|
92 |
||
93 |
fun eof_err () = error "Unexpected end-of-file"; |
|
94 |
||
570
6333c181a3f7
Pure/Thy/thy_parse/THY_PARSE: deleted duplicate specifications of parens,
lcp
parents:
558
diff
changeset
|
95 |
(*Similar to Prolog's cut: reports any syntax error instead of backtracking |
6333c181a3f7
Pure/Thy/thy_parse/THY_PARSE: deleted duplicate specifications of parens,
lcp
parents:
558
diff
changeset
|
96 |
through a superior || *) |
389 | 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"); |
|
100 |
||
101 |
||
102 |
(* parser combinators *) |
|
103 |
||
104 |
fun (parse >> f) toks = apfst f (parse toks); |
|
105 |
||
106 |
fun (parse1 || parse2) toks = |
|
107 |
parse1 toks handle SYNTAX_ERROR _ => parse2 toks; |
|
108 |
||
109 |
fun (parse1 -- parse2) toks = |
|
110 |
let |
|
111 |
val (x, toks') = parse1 toks; |
|
112 |
val (y, toks'') = parse2 toks'; |
|
113 |
in |
|
114 |
((x, y), toks'') |
|
115 |
end; |
|
116 |
||
117 |
fun (parse1 ^^ parse2) = parse1 -- parse2 >> op ^; |
|
118 |
||
119 |
||
120 |
(* generic parsers *) |
|
121 |
||
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 (); |
|
126 |
||
127 |
fun (a $$-- parse) = $$ a -- parse >> #2; |
|
128 |
||
129 |
fun (parse --$$ a) = parse -- $$ a >> #1; |
|
130 |
||
131 |
||
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 (); |
|
136 |
||
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; |
|
141 |
val nat = kind Nat; |
|
142 |
val string = kind String; |
|
143 |
val verbatim = kind Verbatim; |
|
144 |
val eof = kind EOF; |
|
145 |
||
146 |
fun empty toks = ([], toks); |
|
147 |
||
148 |
fun optional parse def = parse || empty >> K def; |
|
149 |
||
150 |
fun repeat parse toks = (parse -- repeat parse >> op :: || empty) toks; |
|
151 |
fun repeat1 parse = parse -- repeat parse >> op ::; |
|
152 |
||
153 |
fun enum1 sep parse = parse -- repeat (sep $$-- parse) >> op ::; |
|
154 |
fun enum sep parse = enum1 sep parse || empty; |
|
155 |
||
2231
71385461570a
Eta-expansion of a function definition, for value polymorphism
paulson
parents:
2203
diff
changeset
|
156 |
fun list1 parse = enum1 "," parse; |
3977 | 157 |
fun list parse = enum "," parse; |
158 |
||
389 | 159 |
|
160 |
||
161 |
(** theory parsers **) |
|
162 |
||
163 |
(* misc utilities *) |
|
164 |
||
165 |
fun cat s1 s2 = s1 ^ " " ^ s2; |
|
166 |
||
558 | 167 |
val parens = enclose "(" ")"; |
168 |
val brackets = enclose "[" "]"; |
|
389 | 169 |
|
170 |
val mk_list = brackets o commas; |
|
171 |
val mk_big_list = brackets o space_implode ",\n "; |
|
172 |
||
558 | 173 |
fun mk_pair (x, y) = parens (commas [x, y]); |
174 |
fun mk_triple (x, y, z) = parens (commas [x, y, z]); |
|
389 | 175 |
fun mk_triple1 ((x, y), z) = mk_triple (x, y, z); |
176 |
fun mk_triple2 (x, (y, z)) = mk_triple (x, y, z); |
|
177 |
||
2231
71385461570a
Eta-expansion of a function definition, for value polymorphism
paulson
parents:
2203
diff
changeset
|
178 |
fun split_decls l = flat (map (fn (xs, y) => map (rpair y) xs) l); |
389 | 179 |
|
6043 | 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); |
|
389 | 183 |
|
184 |
||
185 |
(* names *) |
|
186 |
||
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; |
|
192 |
||
193 |
||
4965 | 194 |
(* empty *) |
195 |
||
196 |
fun empty_decl toks = (empty >> K "") toks; |
|
197 |
||
198 |
||
389 | 199 |
(* classes *) |
200 |
||
201 |
val subclass = name -- optional ("<" $$-- !! name_list1) "[]"; |
|
202 |
||
558 | 203 |
val class_decls = repeat1 (subclass >> mk_pair) >> mk_big_list; |
389 | 204 |
|
205 |
||
206 |
(* arities *) |
|
207 |
||
208 |
val sort = |
|
209 |
name >> brackets || |
|
210 |
"{" $$-- name_list --$$ "}"; |
|
211 |
||
212 |
val sort_list1 = list1 sort >> mk_list; |
|
213 |
||
214 |
||
215 |
val arity = optional ("(" $$-- !! (sort_list1 --$$")")) "[]" -- sort; |
|
216 |
||
217 |
val arity_decls = repeat1 (names1 --$$ "::" -- !! arity) |
|
218 |
>> (mk_big_list o map mk_triple2 o split_decls); |
|
219 |
||
220 |
||
221 |
(* mixfix annotations *) |
|
222 |
||
2203 | 223 |
val infxl = |
224 |
"infixl" $$-- !! (nat >> cat "Infixl" || string -- nat >> (cat "InfixlName" o mk_pair)); |
|
225 |
val infxr = |
|
226 |
"infixr" $$-- !! (nat >> cat "Infixr" || string -- nat >> (cat "InfixrName" o mk_pair)); |
|
389 | 227 |
|
889 | 228 |
val binder = "binder" $$-- |
2203 | 229 |
!! (string -- (("[" $$-- nat --$$ "]") -- nat || nat >> (fn n => (n, n)))) |
230 |
>> (cat "Binder" o mk_triple2); |
|
389 | 231 |
|
232 |
val opt_pris = optional ("[" $$-- !! (list nat --$$ "]")) [] >> mk_list; |
|
233 |
||
5687 | 234 |
val mixfix = string -- !! (opt_pris -- optional nat "Syntax.max_pri") |
389 | 235 |
>> (cat "Mixfix" o mk_triple2); |
236 |
||
558 | 237 |
fun opt_syn fx = optional ("(" $$-- fx --$$ ")") "NoSyn"; |
389 | 238 |
|
239 |
val opt_infix = opt_syn (infxl || infxr); |
|
240 |
val opt_mixfix = opt_syn (mixfix || infxl || infxr || binder); |
|
241 |
||
242 |
||
243 |
(* types *) |
|
244 |
||
3977 | 245 |
(* FIXME clean!! *) |
246 |
||
1705
19fe0ab646b4
changed ident_no_colon so that it forbids postfix "=", too
clasohm
parents:
1555
diff
changeset
|
247 |
(*Parse an identifier, but only if it is not followed by "::", "=" or ","; |
1377 | 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) :: |
|
252 |
toks)) = |
|
1705
19fe0ab646b4
changed ident_no_colon so that it forbids postfix "=", too
clasohm
parents:
1555
diff
changeset
|
253 |
if s2 = "::" orelse s2 = "=" orelse (not allow_comma andalso s2 = ",") |
19fe0ab646b4
changed ident_no_colon so that it forbids postfix "=", too
clasohm
parents:
1555
diff
changeset
|
254 |
then syn_err (name_of_kind Ident) (quote s2) n2 |
1377 | 255 |
else (s, rest) |
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; |
|
1321
9a6e7bd2bfaf
types in consts section of .thy files can now be specified without quotes
clasohm
parents:
1250
diff
changeset
|
259 |
|
3110 | 260 |
(*type used in types, consts and syntax sections*) |
1377 | 261 |
fun const_type allow_comma toks = |
3977 | 262 |
let |
263 |
val simple_type = |
|
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)); |
|
1377 | 270 |
|
271 |
val appl_param = |
|
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])) || |
|
281 |
simple_type || |
|
282 |
"(" $$-- const_type true --$$ ")" >> parens) toks |
|
1321
9a6e7bd2bfaf
types in consts section of .thy files can now be specified without quotes
clasohm
parents:
1250
diff
changeset
|
283 |
end; |
9a6e7bd2bfaf
types in consts section of .thy files can now be specified without quotes
clasohm
parents:
1250
diff
changeset
|
284 |
|
3977 | 285 |
val typ = string || (const_type false >> quote); |
286 |
||
287 |
||
1383 | 288 |
fun mk_old_type_decl ((ts, n), syn) = |
289 |
map (fn t => (mk_triple (t, n, syn), false)) ts; |
|
290 |
||
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)]; |
|
295 |
||
296 |
fun mk_type_decls tys = |
|
3764 | 297 |
"|> Theory.add_types\n" ^ mk_big_list (keyfilter tys false) ^ "\n\n\ |
298 |
\|> Theory.add_tyabbrs\n" ^ mk_big_list (keyfilter tys true); |
|
1383 | 299 |
|
300 |
||
301 |
val old_type_decl = names1 -- nat -- opt_infix >> mk_old_type_decl; |
|
302 |
||
303 |
val type_args = |
|
304 |
type_var >> (fn x => [x]) || |
|
305 |
"(" $$-- !! (list1 type_var --$$ ")") || |
|
306 |
empty >> K []; |
|
307 |
||
308 |
val type_decl = type_args -- name -- |
|
3977 | 309 |
optional ("=" $$-- typ >> Some) None -- opt_infix >> mk_type_decl; |
1383 | 310 |
|
2360 | 311 |
val type_decls = |
312 |
repeat1 (old_type_decl || type_decl) >> (mk_type_decls o flat); |
|
1383 | 313 |
|
314 |
||
315 |
(* consts *) |
|
316 |
||
2360 | 317 |
val const_decls = |
3977 | 318 |
repeat1 (names1 --$$ "::" -- !! (typ -- opt_mixfix)) |
2360 | 319 |
>> (mk_big_list o map mk_triple2 o split_decls); |
389 | 320 |
|
2360 | 321 |
val opt_mode = |
322 |
optional |
|
323 |
("(" $$-- !! (name -- optional ($$ "output" >> K "false") "true" --$$ ")")) |
|
324 |
("\"\"", "true") |
|
325 |
>> mk_pair; |
|
326 |
||
327 |
val syntax_decls = opt_mode -- const_decls >> (fn (mode, txt) => mode ^ "\n" ^ txt); |
|
2203 | 328 |
|
389 | 329 |
|
330 |
(* translations *) |
|
331 |
||
332 |
val trans_pat = |
|
333 |
optional ("(" $$-- !! (name --$$ ")")) "\"logic\"" -- string >> mk_pair; |
|
334 |
||
335 |
val trans_arrow = |
|
3528 | 336 |
$$ "=>" >> K "Syntax.ParseRule " || |
337 |
$$ "<=" >> K "Syntax.PrintRule " || |
|
338 |
$$ "==" >> K "Syntax.ParsePrintRule "; |
|
389 | 339 |
|
1810
0eef167ebe1b
Translation infixes <->, etc., no longer available at top-level
paulson
parents:
1705
diff
changeset
|
340 |
val trans_line = |
2203 | 341 |
trans_pat -- !! (trans_arrow -- trans_pat) |
342 |
>> (fn (left, (arr, right)) => arr ^ mk_pair (left, right)); |
|
1810
0eef167ebe1b
Translation infixes <->, etc., no longer available at top-level
paulson
parents:
1705
diff
changeset
|
343 |
|
0eef167ebe1b
Translation infixes <->, etc., no longer available at top-level
paulson
parents:
1705
diff
changeset
|
344 |
val trans_decls = repeat1 trans_line >> mk_big_list; |
389 | 345 |
|
346 |
||
347 |
(* ML translations *) |
|
348 |
||
4047 | 349 |
val local_defs = |
389 | 350 |
" val parse_ast_translation = [];\n\ |
351 |
\ val parse_translation = [];\n\ |
|
352 |
\ val print_translation = [];\n\ |
|
2385 | 353 |
\ val typed_print_translation = [];\n\ |
2694 | 354 |
\ val print_ast_translation = [];\n\ |
4852 | 355 |
\ val token_translation = [];"; |
389 | 356 |
|
357 |
val trfun_args = |
|
358 |
"(parse_ast_translation, parse_translation, \ |
|
359 |
\print_translation, print_ast_translation)"; |
|
360 |
||
361 |
||
362 |
(* axioms *) |
|
363 |
||
364 |
val mk_axms = mk_big_list o map (mk_pair o apfst quote); |
|
365 |
||
366 |
fun mk_axiom_decls axms = (mk_axms axms, map fst axms); |
|
367 |
||
368 |
val axiom_decls = repeat1 (ident -- !! string) >> mk_axiom_decls; |
|
369 |
||
370 |
||
3813 | 371 |
(* oracle *) |
372 |
||
373 |
val oracle_decl = (name --$$ "=" -- long_id) >> mk_pair; |
|
374 |
||
375 |
||
1555 | 376 |
(* combined consts and axioms *) |
377 |
||
378 |
fun mk_constaxiom_decls x = |
|
379 |
let |
|
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 ^ |
|
4852 | 383 |
"\n\n|> (PureThy.add_defs o map Attribute.none)\n" ^ axms_defs, axms_names) |
1555 | 384 |
end; |
385 |
||
386 |
val constaxiom_decls = |
|
3977 | 387 |
repeat1 (ident --$$ "::" -- !! (typ -- opt_mixfix) -- !! string) |
388 |
>> mk_constaxiom_decls; |
|
1555 | 389 |
|
390 |
||
389 | 391 |
(* axclass *) |
392 |
||
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); |
|
396 |
||
397 |
val axclass_decl = subclass -- repeat (ident -- !! string) >> mk_axclass_decl; |
|
398 |
||
399 |
||
451 | 400 |
(* instance *) |
389 | 401 |
|
425 | 402 |
fun mk_witness (axths, opt_tac) = |
389 | 403 |
mk_list (keyfilter axths false) ^ "\n" ^ |
404 |
mk_list (keyfilter axths true) ^ "\n" ^ |
|
405 |
opt_tac; |
|
406 |
||
407 |
val axm_or_thm = |
|
408 |
string >> rpair false || |
|
409 |
long_id >> rpair true; |
|
410 |
||
451 | 411 |
|
425 | 412 |
val opt_witness = |
389 | 413 |
optional ("(" $$-- list1 axm_or_thm --$$ ")") [] -- |
558 | 414 |
optional (verbatim >> (parens o cat "Some" o parens)) "None" |
425 | 415 |
>> mk_witness; |
416 |
||
417 |
val instance_decl = |
|
636 | 418 |
(name --$$ "<" -- name >> (pair "|> AxClass.add_inst_subclass" o mk_pair) || |
419 |
name --$$ "::" -- arity >> (pair "|> AxClass.add_inst_arity" o mk_triple2)) |
|
451 | 420 |
-- opt_witness |
777 | 421 |
>> (fn ((x, y), z) => (cat_lines [x, y, z])); |
389 | 422 |
|
423 |
||
5248 | 424 |
(* locale *) |
425 |
||
426 |
val locale_decl = |
|
6022 | 427 |
(name --$$ "=") -- |
428 |
(optional ((ident >> (fn x => parens ("Some" ^ quote x))) --$$ "+") ("None")) -- |
|
5248 | 429 |
("fixes" $$-- |
6015 | 430 |
(repeat (name --$$ "::" -- !! (typ -- opt_mixfix)) |
431 |
>> (mk_big_list o map mk_triple2))) -- |
|
432 |
(optional |
|
433 |
("assumes" $$-- (repeat ((ident >> quote) -- !! string) |
|
434 |
>> (mk_list o map mk_pair))) |
|
435 |
"[]") -- |
|
436 |
(optional |
|
437 |
("defines" $$-- (repeat ((ident >> quote) -- !! string) |
|
438 |
>> (mk_list o map mk_pair))) |
|
439 |
"[]") |
|
6022 | 440 |
>> (fn ((((nm, ext), cs), asms), defs) => cat_lines [nm, ext, cs, asms, defs]); |
5248 | 441 |
|
442 |
||
389 | 443 |
|
444 |
(** theory syntax **) |
|
445 |
||
446 |
type syntax = |
|
4707 | 447 |
Scan.lexicon * (token list -> (string * string) * token list) Symtab.table; |
389 | 448 |
|
449 |
fun make_syntax keywords sects = |
|
4056 | 450 |
let |
451 |
val dups = duplicates (map fst sects); |
|
4496 | 452 |
val sects' = gen_distinct eq_fst sects; |
4707 | 453 |
val keys = map Symbol.explode (map fst sects' @ keywords); |
4056 | 454 |
in |
455 |
if null dups then () |
|
456 |
else warning ("Duplicate declaration of theory file sections:\n" ^ commas_quote dups); |
|
4707 | 457 |
(Scan.make_lexicon keys, Symtab.make sects') |
4056 | 458 |
end; |
389 | 459 |
|
460 |
||
461 |
(* header *) |
|
462 |
||
463 |
fun mk_header (thy_name, bases) = |
|
586
201e115d8031
renamed base_on into mk_base and moved it to the beginning of the generated
clasohm
parents:
570
diff
changeset
|
464 |
(thy_name, "mk_base " ^ mk_list bases ^ " " ^ quote thy_name); |
389 | 465 |
|
466 |
val base = |
|
467 |
ident >> (cat "Thy" o quote) || |
|
468 |
string >> cat "File"; |
|
469 |
||
470 |
val header = ident --$$ "=" -- enum1 "+" base >> mk_header; |
|
471 |
||
472 |
||
473 |
(* extension *) |
|
474 |
||
3900 | 475 |
fun mk_extension (txts, mltxt) = |
389 | 476 |
let |
477 |
val cat_sects = space_implode "\n\n" o filter_out (equal ""); |
|
478 |
val (extxts, postxts) = split_list txts; |
|
479 |
in |
|
3900 | 480 |
(cat_sects extxts, cat_sects postxts, mltxt) |
389 | 481 |
end; |
482 |
||
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 (); |
|
489 |
||
3900 | 490 |
fun extension sectab = "+" $$-- !! |
4707 | 491 |
(repeat (sect sectab) --$$ "end" -- optional verbatim "") |
3900 | 492 |
>> mk_extension; |
3813 | 493 |
|
3900 | 494 |
fun opt_extension sectab = optional (extension sectab) ("", "", ""); |
3875 | 495 |
|
389 | 496 |
|
497 |
(* theory definition *) |
|
498 |
||
3900 | 499 |
fun mk_structure tname ((thy_name, old_thys), (extxt, postxt, mltxt)) = |
558 | 500 |
if thy_name <> tname then |
501 |
error ("Filename \"" ^ tname ^ ".thy\" and theory name " |
|
502 |
^ quote thy_name ^ " are different") |
|
503 |
else |
|
3875 | 504 |
"val thy = " ^ old_thys ^ ";\n\n\ |
505 |
\structure " ^ thy_name ^ " =\n\ |
|
506 |
\struct\n\ |
|
507 |
\\n\ |
|
4965 | 508 |
\local\n" |
4047 | 509 |
^ local_defs ^ "\n\ |
3875 | 510 |
\in\n\ |
511 |
\\n" |
|
512 |
^ mltxt ^ "\n\ |
|
513 |
\\n\ |
|
514 |
\val thy = thy\n\ |
|
4965 | 515 |
\|> PureThy.put_name " ^ quote thy_name ^ "\n\ |
516 |
\|> PureThy.local_path\n\ |
|
3875 | 517 |
\|> Theory.add_trfuns\n" |
518 |
^ trfun_args ^ "\n\ |
|
4047 | 519 |
\|> Theory.add_trfunsT typed_print_translation\n\ |
520 |
\|> Theory.add_tokentrfuns token_translation\n\ |
|
3875 | 521 |
\\n" |
522 |
^ extxt ^ "\n\ |
|
523 |
\\n\ |
|
4965 | 524 |
\|> PureThy.end_theory\n\ |
3875 | 525 |
\\n\ |
4965 | 526 |
\val _ = store_theory thy;\n\ |
4056 | 527 |
\val _ = context thy;\n\ |
3875 | 528 |
\\n\ |
529 |
\\n" |
|
530 |
^ postxt ^ "\n\ |
|
531 |
\\n\ |
|
532 |
\end;\n\ |
|
533 |
\end;\n\ |
|
534 |
\\n\ |
|
535 |
\open " ^ thy_name ^ ";\n\ |
|
536 |
\\n"; |
|
389 | 537 |
|
476
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
538 |
fun theory_defn sectab tname = |
3875 | 539 |
header -- opt_extension sectab -- eof >> (mk_structure tname o #1); |
389 | 540 |
|
476
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
541 |
fun parse_thy (lex, sectab) tname str = |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
542 |
#1 (!! (theory_defn sectab tname) (tokenize lex str)); |
389 | 543 |
|
544 |
||
545 |
(* standard sections *) |
|
546 |
||
547 |
fun mk_val ax = "val " ^ ax ^ " = get_axiom thy " ^ quote ax ^ ";"; |
|
777 | 548 |
val mk_vals = cat_lines o map mk_val; |
389 | 549 |
|
777 | 550 |
fun mk_axm_sect "" (txt, axs) = (txt, mk_vals axs) |
551 |
| mk_axm_sect pretxt (txt, axs) = (pretxt ^ "\n" ^ txt, mk_vals axs); |
|
389 | 552 |
|
553 |
fun axm_section name pretxt parse = |
|
554 |
(name, parse >> mk_axm_sect pretxt); |
|
555 |
||
556 |
fun section name pretxt parse = |
|
557 |
axm_section name pretxt (parse >> rpair []); |
|
558 |
||
559 |
||
560 |
val pure_keywords = |
|
4952 | 561 |
["end", "ML", "mixfix", "infixr", "infixl", "binder", "output", "=", |
562 |
"+", ",", "<", "{", "}", "(", ")", "[", "]", "::", "==", "=>", |
|
5248 | 563 |
"<=", "fixes", "assumes", "defines"]; |
389 | 564 |
|
565 |
val pure_sections = |
|
3813 | 566 |
[section "classes" "|> Theory.add_classes" class_decls, |
3764 | 567 |
section "default" "|> Theory.add_defsort" sort, |
777 | 568 |
section "types" "" type_decls, |
4852 | 569 |
section "nonterminals" "|> Theory.add_nonterminals" (repeat1 name >> mk_list), |
3764 | 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, |
|
4852 | 574 |
axm_section "rules" "|> (PureThy.add_axioms o map Attribute.none)" axiom_decls, |
575 |
axm_section "defs" "|> (PureThy.add_defs o map Attribute.none)" axiom_decls, |
|
3813 | 576 |
section "oracle" "|> Theory.add_oracle" oracle_decl, |
3764 | 577 |
axm_section "constdefs" "|> Theory.add_consts" constaxiom_decls, |
636 | 578 |
axm_section "axclass" "|> AxClass.add_axclass" axclass_decl, |
3780 | 579 |
section "instance" "" instance_decl, |
3900 | 580 |
section "path" "|> Theory.add_path" name, |
4965 | 581 |
section "global" "|> PureThy.global_path" empty_decl, |
582 |
section "local" "|> PureThy.local_path" empty_decl, |
|
5905 | 583 |
section "setup" "|> Library.apply" long_id, |
5248 | 584 |
section "MLtext" "" verbatim, |
585 |
section "locale" "|> Locale.add_locale" locale_decl]; |
|
389 | 586 |
|
3813 | 587 |
|
389 | 588 |
end; |