author | clasohm |
Fri, 15 Jul 1994 13:30:42 +0200 | |
changeset 476 | 836cad329311 |
parent 451 | 9ebdead316e0 |
child 558 | c4092ae47210 |
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. |
|
451 | 6 |
|
7 |
TODO: |
|
8 |
remove sigclass (?) |
|
389 | 9 |
*) |
10 |
||
11 |
infix 5 -- --$$ $$-- ^^; |
|
12 |
infix 3 >>; |
|
13 |
infix 0 ||; |
|
14 |
||
15 |
signature THY_PARSE = |
|
16 |
sig |
|
17 |
type token |
|
18 |
val !! : ('a -> 'b * 'c) -> 'a -> 'b * 'c |
|
19 |
val >> : ('a -> 'b * 'c) * ('b -> 'd) -> 'a -> 'd * 'c |
|
20 |
val || : ('a -> 'b) * ('a -> 'b) -> 'a -> 'b |
|
21 |
val -- : ('a -> 'b * 'c) * ('c -> 'd * 'e) -> 'a -> ('b * 'd) * 'e |
|
22 |
val ^^ : ('a -> string * 'b) * ('b -> string * 'c) -> 'a -> string * 'c |
|
23 |
val $$ : string -> token list -> string * token list |
|
24 |
val $$-- : string * (token list -> 'b * 'c) -> token list -> 'b * 'c |
|
25 |
val --$$ : ('a -> 'b * token list) * string -> 'a -> 'b * token list |
|
26 |
val ident: token list -> string * token list |
|
27 |
val long_ident: token list -> string * token list |
|
28 |
val long_id: token list -> string * token list |
|
29 |
val type_var: token list -> string * token list |
|
30 |
val nat: token list -> string * token list |
|
31 |
val string: token list -> string * token list |
|
32 |
val verbatim: token list -> string * token list |
|
33 |
val empty: 'a -> 'b list * 'a |
|
34 |
val optional: ('a -> 'b * 'a) -> 'b -> 'a -> 'b * 'a |
|
35 |
val repeat: ('a -> 'b * 'a) -> 'a -> 'b list * 'a |
|
36 |
val repeat1: ('a -> 'b * 'a) -> 'a -> 'b list * 'a |
|
37 |
val enum: string -> (token list -> 'a * token list) |
|
38 |
-> token list -> 'a list * token list |
|
39 |
val enum1: string -> (token list -> 'a * token list) |
|
40 |
-> token list -> 'a list * token list |
|
41 |
val list: (token list -> 'a * token list) |
|
42 |
-> token list -> 'a list * token list |
|
43 |
val list1: (token list -> 'a * token list) |
|
44 |
-> token list -> 'a list * token list |
|
45 |
val name: token list -> string * token list |
|
46 |
val sort: token list -> string * token list |
|
451 | 47 |
val opt_infix: token list -> string * token list |
48 |
val opt_mixfix: token list -> string * token list |
|
389 | 49 |
type syntax |
50 |
val make_syntax: string list -> |
|
51 |
(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
|
52 |
val parse_thy: syntax -> string -> string -> string |
389 | 53 |
val section: string -> string -> (token list -> string * token list) |
54 |
-> (string * (token list -> (string * string) * token list)) |
|
55 |
val axm_section: string -> string |
|
56 |
-> (token list -> (string * string list) * token list) |
|
57 |
-> (string * (token list -> (string * string) * token list)) |
|
58 |
val pure_keywords: string list |
|
59 |
val pure_sections: |
|
60 |
(string * (token list -> (string * string) * token list)) list |
|
61 |
end; |
|
62 |
||
414 | 63 |
functor ThyParseFun(structure Symtab: SYMTAB and ThyScan: THY_SCAN): THY_PARSE = |
389 | 64 |
struct |
65 |
||
66 |
open ThyScan; |
|
67 |
||
68 |
||
69 |
(** parser toolbox **) |
|
70 |
||
71 |
type token = token_kind * string * int; |
|
72 |
||
73 |
||
74 |
(* errors *) |
|
75 |
||
76 |
exception SYNTAX_ERROR of string * string * int; |
|
77 |
||
78 |
fun syn_err s1 s2 n = raise SYNTAX_ERROR (s1, s2, n); |
|
79 |
||
80 |
fun eof_err () = error "Unexpected end-of-file"; |
|
81 |
||
82 |
fun !! parse toks = parse toks |
|
83 |
handle SYNTAX_ERROR (s1, s2, n) => error ("Syntax error on line " ^ |
|
84 |
string_of_int n ^ ": " ^ s1 ^ " expected and " ^ s2 ^ " was found"); |
|
85 |
||
86 |
||
87 |
(* parser combinators *) |
|
88 |
||
89 |
fun (parse >> f) toks = apfst f (parse toks); |
|
90 |
||
91 |
fun (parse1 || parse2) toks = |
|
92 |
parse1 toks handle SYNTAX_ERROR _ => parse2 toks; |
|
93 |
||
94 |
fun (parse1 -- parse2) toks = |
|
95 |
let |
|
96 |
val (x, toks') = parse1 toks; |
|
97 |
val (y, toks'') = parse2 toks'; |
|
98 |
in |
|
99 |
((x, y), toks'') |
|
100 |
end; |
|
101 |
||
102 |
fun (parse1 ^^ parse2) = parse1 -- parse2 >> op ^; |
|
103 |
||
104 |
||
105 |
(* generic parsers *) |
|
106 |
||
107 |
fun $$ a ((k, b, n) :: toks) = |
|
108 |
if k = Keyword andalso a = b then (a, toks) |
|
109 |
else syn_err (quote a) (quote b) n |
|
110 |
| $$ _ [] = eof_err (); |
|
111 |
||
112 |
fun (a $$-- parse) = $$ a -- parse >> #2; |
|
113 |
||
114 |
fun (parse --$$ a) = parse -- $$ a >> #1; |
|
115 |
||
116 |
||
117 |
fun kind k1 ((k2, s, n) :: toks) = |
|
118 |
if k1 = k2 then (s, toks) |
|
119 |
else syn_err (name_of_kind k1) (quote s) n |
|
120 |
| kind _ [] = eof_err (); |
|
121 |
||
122 |
val ident = kind Ident; |
|
123 |
val long_ident = kind LongIdent; |
|
124 |
val long_id = ident || long_ident; |
|
125 |
val type_var = kind TypeVar >> quote; |
|
126 |
val nat = kind Nat; |
|
127 |
val string = kind String; |
|
128 |
val verbatim = kind Verbatim; |
|
129 |
val eof = kind EOF; |
|
130 |
||
131 |
fun empty toks = ([], toks); |
|
132 |
||
133 |
fun optional parse def = parse || empty >> K def; |
|
134 |
||
135 |
fun repeat parse toks = (parse -- repeat parse >> op :: || empty) toks; |
|
136 |
fun repeat1 parse = parse -- repeat parse >> op ::; |
|
137 |
||
138 |
fun enum1 sep parse = parse -- repeat (sep $$-- parse) >> op ::; |
|
139 |
fun enum sep parse = enum1 sep parse || empty; |
|
140 |
||
141 |
val list = enum ","; |
|
142 |
val list1 = enum1 ","; |
|
143 |
||
144 |
||
145 |
||
146 |
(** theory parsers **) |
|
147 |
||
148 |
(* misc utilities *) |
|
149 |
||
150 |
fun cat s1 s2 = s1 ^ " " ^ s2; |
|
151 |
||
152 |
val pars = parents "(" ")"; |
|
153 |
val brackets = parents "[" "]"; |
|
154 |
||
155 |
val mk_list = brackets o commas; |
|
156 |
val mk_big_list = brackets o space_implode ",\n "; |
|
157 |
||
158 |
fun mk_pair (x, y) = pars (commas [x, y]); |
|
159 |
fun mk_triple (x, y, z) = pars (commas [x, y, z]); |
|
160 |
fun mk_triple1 ((x, y), z) = mk_triple (x, y, z); |
|
161 |
fun mk_triple2 (x, (y, z)) = mk_triple (x, y, z); |
|
162 |
||
163 |
val split_decls = flat o map (fn (xs, y) => map (rpair y) xs); |
|
164 |
||
165 |
fun strip_quotes str = |
|
166 |
implode (tl (take (size str - 1, explode str))); |
|
167 |
||
168 |
||
169 |
(* names *) |
|
170 |
||
171 |
val name = ident >> quote || string; |
|
172 |
val names = list name; |
|
173 |
val names1 = list1 name; |
|
174 |
val name_list = names >> mk_list; |
|
175 |
val name_list1 = names1 >> mk_list; |
|
176 |
||
177 |
||
178 |
(* classes *) |
|
179 |
||
180 |
fun mk_subclass (c, cs) = mk_triple ("[]", c, cs); |
|
181 |
||
182 |
val subclass = name -- optional ("<" $$-- !! name_list1) "[]"; |
|
183 |
||
184 |
val class_decls = repeat1 (subclass >> mk_subclass) >> mk_big_list; |
|
185 |
||
186 |
||
187 |
(* arities *) |
|
188 |
||
189 |
val sort = |
|
190 |
name >> brackets || |
|
191 |
"{" $$-- name_list --$$ "}"; |
|
192 |
||
193 |
val sort_list1 = list1 sort >> mk_list; |
|
194 |
||
195 |
||
196 |
val arity = optional ("(" $$-- !! (sort_list1 --$$")")) "[]" -- sort; |
|
197 |
||
198 |
val arity_decls = repeat1 (names1 --$$ "::" -- !! arity) |
|
199 |
>> (mk_big_list o map mk_triple2 o split_decls); |
|
200 |
||
201 |
||
202 |
(* mixfix annotations *) |
|
203 |
||
204 |
val infxl = "infixl" $$-- !! nat >> cat "Infixl"; |
|
205 |
val infxr = "infixr" $$-- !! nat >> cat "Infixr"; |
|
206 |
||
207 |
val binder = "binder" $$-- !! (string -- nat) >> (cat "Binder" o mk_pair); |
|
208 |
||
209 |
val opt_pris = optional ("[" $$-- !! (list nat --$$ "]")) [] >> mk_list; |
|
210 |
||
211 |
val mixfix = string -- !! (opt_pris -- optional nat "max_pri") |
|
212 |
>> (cat "Mixfix" o mk_triple2); |
|
213 |
||
214 |
fun opt_syn fx = |
|
215 |
"(" $$-- fx --$$ ")" || |
|
216 |
empty >> K "NoSyn"; |
|
217 |
||
218 |
val opt_infix = opt_syn (infxl || infxr); |
|
219 |
val opt_mixfix = opt_syn (mixfix || infxl || infxr || binder); |
|
220 |
||
221 |
||
222 |
(* types *) |
|
223 |
||
224 |
fun mk_old_type_decl ((ts, n), syn) = |
|
225 |
map (fn t => (mk_triple (t, n, syn), false)) ts; |
|
226 |
||
227 |
fun mk_type_decl (((xs, t), None), syn) = |
|
228 |
[(mk_triple (t, string_of_int (length xs), syn), false)] |
|
229 |
| mk_type_decl (((xs, t), Some rhs), syn) = |
|
230 |
[(pars (commas [t, mk_list xs, rhs, syn]), true)]; |
|
231 |
||
232 |
fun mk_type_decls tys = |
|
414 | 233 |
"|> add_types\n" ^ mk_big_list (keyfilter tys false) ^ "\n\n\ |
234 |
\|> add_tyabbrs\n" ^ mk_big_list (keyfilter tys true); |
|
389 | 235 |
|
236 |
||
425 | 237 |
val old_type_decl = names1 -- nat -- opt_infix >> mk_old_type_decl; |
389 | 238 |
|
239 |
val type_args = |
|
240 |
type_var >> (fn x => [x]) || |
|
241 |
"(" $$-- !! (list1 type_var --$$ ")") || |
|
242 |
empty >> K []; |
|
243 |
||
244 |
val type_decl = type_args -- name -- optional ("=" $$-- !! string >> Some) None |
|
245 |
-- opt_infix >> mk_type_decl; |
|
246 |
||
247 |
val type_decls = repeat1 (old_type_decl || type_decl) |
|
248 |
>> (rpair "" o mk_type_decls o flat); |
|
249 |
||
250 |
||
251 |
(* consts *) |
|
252 |
||
253 |
val const_decls = repeat1 (names1 --$$ "::" -- !! (string -- opt_mixfix)) |
|
254 |
>> (mk_big_list o map mk_triple2 o split_decls); |
|
255 |
||
256 |
||
257 |
(* translations *) |
|
258 |
||
259 |
val trans_pat = |
|
260 |
optional ("(" $$-- !! (name --$$ ")")) "\"logic\"" -- string >> mk_pair; |
|
261 |
||
262 |
val trans_arrow = |
|
263 |
$$ "=>" >> K " |-> " || |
|
264 |
$$ "<=" >> K " <-| " || |
|
265 |
$$ "==" >> K " <-> "; |
|
266 |
||
267 |
val trans_decls = repeat1 (trans_pat ^^ !! (trans_arrow ^^ trans_pat)) |
|
268 |
>> mk_big_list; |
|
269 |
||
270 |
||
271 |
(* ML translations *) |
|
272 |
||
273 |
val trfun_defs = |
|
274 |
" val parse_ast_translation = [];\n\ |
|
275 |
\ val parse_translation = [];\n\ |
|
276 |
\ val print_translation = [];\n\ |
|
277 |
\ val print_ast_translation = [];"; |
|
278 |
||
279 |
val trfun_args = |
|
280 |
"(parse_ast_translation, parse_translation, \ |
|
281 |
\print_translation, print_ast_translation)"; |
|
282 |
||
283 |
fun mk_mltrans txt = |
|
284 |
"let\n" |
|
285 |
^ trfun_defs ^ "\n" |
|
286 |
^ txt ^ "\n\ |
|
287 |
\in\n\ |
|
288 |
\ " ^ trfun_args ^ "\n\ |
|
289 |
\end"; |
|
290 |
||
291 |
val mltrans = verbatim >> mk_mltrans; |
|
292 |
||
293 |
||
294 |
(* axioms *) |
|
295 |
||
296 |
val mk_axms = mk_big_list o map (mk_pair o apfst quote); |
|
297 |
||
298 |
fun mk_axiom_decls axms = (mk_axms axms, map fst axms); |
|
299 |
||
300 |
val axiom_decls = repeat1 (ident -- !! string) >> mk_axiom_decls; |
|
301 |
||
302 |
||
303 |
(* axclass *) |
|
304 |
||
305 |
fun mk_axclass_decl ((c, cs), axms) = |
|
306 |
(mk_pair (c, cs) ^ "\n" ^ mk_axms axms, |
|
307 |
(strip_quotes c ^ "I") :: map fst axms); |
|
308 |
||
309 |
val axclass_decl = subclass -- repeat (ident -- !! string) >> mk_axclass_decl; |
|
310 |
||
311 |
||
414 | 312 |
(* sigclass *) |
313 |
||
425 | 314 |
fun mk_sigclass_decl ((c, cs), consts) = |
414 | 315 |
(mk_pair (c, cs) ^ "\n" ^ consts, [strip_quotes c ^ "I"]); |
316 |
||
317 |
val sigclass_decl = subclass -- optional const_decls "[]" >> mk_sigclass_decl; |
|
318 |
||
319 |
||
451 | 320 |
(* instance *) |
389 | 321 |
|
425 | 322 |
fun mk_witness (axths, opt_tac) = |
389 | 323 |
mk_list (keyfilter axths false) ^ "\n" ^ |
324 |
mk_list (keyfilter axths true) ^ "\n" ^ |
|
325 |
opt_tac; |
|
326 |
||
327 |
val axm_or_thm = |
|
328 |
string >> rpair false || |
|
329 |
long_id >> rpair true; |
|
330 |
||
451 | 331 |
|
425 | 332 |
val opt_witness = |
389 | 333 |
optional ("(" $$-- list1 axm_or_thm --$$ ")") [] -- |
334 |
optional (verbatim >> (pars o cat "Some" o pars)) "None" |
|
425 | 335 |
>> mk_witness; |
336 |
||
337 |
val instance_decl = |
|
451 | 338 |
(name --$$ "<" -- name >> (pair "|> add_inst_subclass" o mk_pair) || |
339 |
name --$$ "::" -- arity >> (pair "|> add_inst_arity" o mk_triple2)) |
|
340 |
-- opt_witness |
|
341 |
>> (fn ((x, y), z) => (cat_lines [x, y, z], "")); |
|
389 | 342 |
|
343 |
||
344 |
||
345 |
(** theory syntax **) |
|
346 |
||
347 |
type syntax = |
|
348 |
lexicon * (token list -> (string * string) * token list) Symtab.table; |
|
349 |
||
350 |
fun make_syntax keywords sects = |
|
451 | 351 |
(make_lexicon (map fst sects @ keywords), |
352 |
Symtab.make sects handle Symtab.DUPS dups => |
|
353 |
error ("Duplicate sections in theory file syntax: " ^ commas_quote dups)); |
|
389 | 354 |
|
355 |
||
356 |
(* header *) |
|
357 |
||
358 |
fun mk_header (thy_name, bases) = |
|
425 | 359 |
(thy_name, "base_on " ^ mk_list bases ^ " " ^ quote thy_name); |
389 | 360 |
|
361 |
val base = |
|
362 |
ident >> (cat "Thy" o quote) || |
|
363 |
string >> cat "File"; |
|
364 |
||
365 |
val header = ident --$$ "=" -- enum1 "+" base >> mk_header; |
|
366 |
||
367 |
||
368 |
(* extension *) |
|
369 |
||
370 |
fun mk_extension (txts, mltxt) = |
|
371 |
let |
|
372 |
val cat_sects = space_implode "\n\n" o filter_out (equal ""); |
|
373 |
val (extxts, postxts) = split_list txts; |
|
374 |
in |
|
375 |
(cat_sects extxts, cat_sects postxts, mltxt) |
|
376 |
end; |
|
377 |
||
378 |
fun sect tab ((Keyword, s, n) :: toks) = |
|
379 |
(case Symtab.lookup (tab, s) of |
|
380 |
Some parse => !! parse toks |
|
381 |
| None => syn_err "section" s n) |
|
382 |
| sect _ ((_, s, n) :: _) = syn_err "section" s n |
|
383 |
| sect _ [] = eof_err (); |
|
384 |
||
385 |
fun extension sectab = "+" $$-- !! (repeat (sect sectab) --$$ "end") -- |
|
386 |
optional ("ML" $$-- verbatim) "" >> mk_extension; |
|
387 |
||
388 |
||
389 |
(* theory definition *) |
|
390 |
||
476
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
391 |
fun mk_structure tname ((thy_name, old_thys), Some (extxt, postxt, mltxt)) = |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
392 |
if (thy_name = tname) then |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
393 |
"structure " ^ thy_name ^ " =\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
394 |
\struct\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
395 |
\\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
396 |
\local\n" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
397 |
^ trfun_defs ^ "\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
398 |
\in\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
399 |
\\n" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
400 |
^ mltxt ^ "\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
401 |
\\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
402 |
\val thy = (" ^ old_thys ^ " true)\n\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
403 |
\|> add_trfuns\n" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
404 |
^ trfun_args ^ "\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
405 |
\\n" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
406 |
^ extxt ^ "\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
407 |
\\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
408 |
\|> add_thyname " ^ quote thy_name ^ ";\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
409 |
\\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
410 |
\\n" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
411 |
^ postxt ^ "\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
412 |
\\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
413 |
\end;\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
414 |
\end;\n" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
415 |
else error ("use_thy: Filename \"" ^ tname ^ ".thy\" and theory name \"" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
416 |
^ thy_name ^ "\" are different") |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
417 |
| mk_structure tname ((thy_name, old_thys), None) = |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
418 |
if (thy_name = tname) then |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
419 |
"structure " ^ thy_name ^ " =\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
420 |
\struct\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
421 |
\\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
422 |
\val thy = (" ^ old_thys ^ " false);\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
423 |
\\n\ |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
424 |
\end;\n" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
425 |
else error ("use_thy: Filename \"" ^ tname ^ ".thy\" and theory name \"" |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
426 |
^ thy_name ^ "\" are different"); |
389 | 427 |
|
476
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
428 |
fun theory_defn sectab tname = |
389 | 429 |
header -- optional (extension sectab >> Some) None -- eof |
476
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
430 |
>> (mk_structure tname o #1); |
389 | 431 |
|
476
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
432 |
fun parse_thy (lex, sectab) tname str = |
836cad329311
added check for concistency of filename and theory name;
clasohm
parents:
451
diff
changeset
|
433 |
#1 (!! (theory_defn sectab tname) (tokenize lex str)); |
389 | 434 |
|
435 |
||
436 |
(* standard sections *) |
|
437 |
||
438 |
fun mk_val ax = "val " ^ ax ^ " = get_axiom thy " ^ quote ax ^ ";"; |
|
439 |
||
440 |
fun mk_axm_sect pretxt (txt, axs) = |
|
441 |
(pretxt ^ "\n" ^ txt, cat_lines (map mk_val axs)); |
|
442 |
||
443 |
fun axm_section name pretxt parse = |
|
444 |
(name, parse >> mk_axm_sect pretxt); |
|
445 |
||
446 |
fun section name pretxt parse = |
|
447 |
axm_section name pretxt (parse >> rpair []); |
|
448 |
||
449 |
||
450 |
val pure_keywords = |
|
451 | 451 |
["end", "ML", "mixfix", "infixr", "infixl", "binder", "=", "+", ",", "<", |
452 |
"{", "}", "(", ")", "[", "]", "::", "==", "=>", "<="]; |
|
389 | 453 |
|
454 |
val pure_sections = |
|
414 | 455 |
[section "classes" "|> add_classes" class_decls, |
456 |
section "default" "|> add_defsort" sort, |
|
389 | 457 |
("types", type_decls), |
414 | 458 |
section "arities" "|> add_arities" arity_decls, |
459 |
section "consts" "|> add_consts" const_decls, |
|
460 |
section "syntax" "|> add_syntax" const_decls, |
|
461 |
section "translations" "|> add_trrules" trans_decls, |
|
462 |
section "MLtrans" "|> add_trfuns" mltrans, |
|
389 | 463 |
("MLtext", verbatim >> rpair ""), |
414 | 464 |
axm_section "rules" "|> add_axioms" axiom_decls, |
465 |
axm_section "defns" "|> add_defns" axiom_decls, |
|
466 |
axm_section "axclass" "|> add_axclass" axclass_decl, |
|
467 |
axm_section "sigclass" "|> add_sigclass" sigclass_decl, |
|
451 | 468 |
("instance", instance_decl)]; |
389 | 469 |
|
470 |
||
471 |
end; |
|
472 |