| author | wenzelm | 
| Wed, 15 Oct 1997 15:12:59 +0200 | |
| changeset 3872 | a5839ecee7b8 | 
| parent 3813 | e6142be74e59 | 
| child 3875 | f62a4edb1888 | 
| 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 = | |
| 1512 | 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 | |
| 451 | 45 | val opt_infix: token list -> string * token list | 
| 46 | val opt_mixfix: token list -> string * token list | |
| 636 | 47 | val opt_witness: token list -> string * token list | 
| 389 | 48 | type syntax | 
| 49 | val make_syntax: string list -> | |
| 50 | (string * (token list -> (string * string) * token list)) list -> syntax | |
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 51 | val parse_thy: syntax -> string -> string -> string | 
| 389 | 52 | val section: string -> string -> (token list -> string * token list) | 
| 53 | -> (string * (token list -> (string * string) * token list)) | |
| 54 | val axm_section: string -> string | |
| 55 | -> (token list -> (string * string list) * token list) | |
| 56 | -> (string * (token list -> (string * string) * token list)) | |
| 57 | val pure_keywords: string list | |
| 58 | val pure_sections: | |
| 59 | (string * (token list -> (string * string) * token list)) list | |
| 570 
6333c181a3f7
Pure/Thy/thy_parse/THY_PARSE: deleted duplicate specifications of parens,
 lcp parents: 
558diff
changeset | 60 | (*items for building strings*) | 
| 710 | 61 | val cat: string -> string -> string | 
| 656 | 62 | val parens: string -> string | 
| 63 | val brackets: string -> string | |
| 64 | val mk_list: string list -> string | |
| 65 | val mk_big_list: string list -> string | |
| 66 | val mk_pair: string * string -> string | |
| 67 | val mk_triple: string * string * string -> string | |
| 68 | val strip_quotes: string -> string | |
| 1512 | 69 | end; | 
| 389 | 70 | |
| 1512 | 71 | |
| 72 | structure ThyParse : THY_PARSE= | |
| 389 | 73 | struct | 
| 74 | ||
| 75 | open ThyScan; | |
| 76 | ||
| 77 | ||
| 78 | (** parser toolbox **) | |
| 79 | ||
| 80 | type token = token_kind * string * int; | |
| 81 | ||
| 82 | ||
| 83 | (* errors *) | |
| 84 | ||
| 85 | exception SYNTAX_ERROR of string * string * int; | |
| 86 | ||
| 87 | fun syn_err s1 s2 n = raise SYNTAX_ERROR (s1, s2, n); | |
| 88 | ||
| 89 | fun eof_err () = error "Unexpected end-of-file"; | |
| 90 | ||
| 570 
6333c181a3f7
Pure/Thy/thy_parse/THY_PARSE: deleted duplicate specifications of parens,
 lcp parents: 
558diff
changeset | 91 | (*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: 
558diff
changeset | 92 | through a superior || *) | 
| 389 | 93 | fun !! parse toks = parse toks | 
| 94 |   handle SYNTAX_ERROR (s1, s2, n) => error ("Syntax error on line " ^
 | |
| 95 | string_of_int n ^ ": " ^ s1 ^ " expected and " ^ s2 ^ " was found"); | |
| 96 | ||
| 97 | ||
| 98 | (* parser combinators *) | |
| 99 | ||
| 100 | fun (parse >> f) toks = apfst f (parse toks); | |
| 101 | ||
| 102 | fun (parse1 || parse2) toks = | |
| 103 | parse1 toks handle SYNTAX_ERROR _ => parse2 toks; | |
| 104 | ||
| 105 | fun (parse1 -- parse2) toks = | |
| 106 | let | |
| 107 | val (x, toks') = parse1 toks; | |
| 108 | val (y, toks'') = parse2 toks'; | |
| 109 | in | |
| 110 | ((x, y), toks'') | |
| 111 | end; | |
| 112 | ||
| 113 | fun (parse1 ^^ parse2) = parse1 -- parse2 >> op ^; | |
| 114 | ||
| 115 | ||
| 116 | (* generic parsers *) | |
| 117 | ||
| 118 | fun $$ a ((k, b, n) :: toks) = | |
| 119 | if k = Keyword andalso a = b then (a, toks) | |
| 120 | else syn_err (quote a) (quote b) n | |
| 121 | | $$ _ [] = eof_err (); | |
| 122 | ||
| 123 | fun (a $$-- parse) = $$ a -- parse >> #2; | |
| 124 | ||
| 125 | fun (parse --$$ a) = parse -- $$ a >> #1; | |
| 126 | ||
| 127 | ||
| 128 | fun kind k1 ((k2, s, n) :: toks) = | |
| 129 | if k1 = k2 then (s, toks) | |
| 130 | else syn_err (name_of_kind k1) (quote s) n | |
| 131 | | kind _ [] = eof_err (); | |
| 132 | ||
| 133 | val ident = kind Ident; | |
| 134 | val long_ident = kind LongIdent; | |
| 135 | val long_id = ident || long_ident; | |
| 136 | val type_var = kind TypeVar >> quote; | |
| 137 | val nat = kind Nat; | |
| 138 | val string = kind String; | |
| 139 | val verbatim = kind Verbatim; | |
| 140 | val eof = kind EOF; | |
| 141 | ||
| 142 | fun empty toks = ([], toks); | |
| 143 | ||
| 144 | fun optional parse def = parse || empty >> K def; | |
| 145 | ||
| 146 | fun repeat parse toks = (parse -- repeat parse >> op :: || empty) toks; | |
| 147 | fun repeat1 parse = parse -- repeat parse >> op ::; | |
| 148 | ||
| 149 | fun enum1 sep parse = parse -- repeat (sep $$-- parse) >> op ::; | |
| 150 | fun enum sep parse = enum1 sep parse || empty; | |
| 151 | ||
| 2231 
71385461570a
Eta-expansion of a function definition, for value polymorphism
 paulson parents: 
2203diff
changeset | 152 | fun list parse = enum "," parse; | 
| 
71385461570a
Eta-expansion of a function definition, for value polymorphism
 paulson parents: 
2203diff
changeset | 153 | fun list1 parse = enum1 "," parse; | 
| 389 | 154 | |
| 155 | ||
| 156 | (** theory parsers **) | |
| 157 | ||
| 158 | (* misc utilities *) | |
| 159 | ||
| 160 | fun cat s1 s2 = s1 ^ " " ^ s2; | |
| 161 | ||
| 558 | 162 | val parens = enclose "(" ")";
 | 
| 163 | val brackets = enclose "[" "]"; | |
| 389 | 164 | |
| 165 | val mk_list = brackets o commas; | |
| 166 | val mk_big_list = brackets o space_implode ",\n "; | |
| 167 | ||
| 558 | 168 | fun mk_pair (x, y) = parens (commas [x, y]); | 
| 169 | fun mk_triple (x, y, z) = parens (commas [x, y, z]); | |
| 389 | 170 | fun mk_triple1 ((x, y), z) = mk_triple (x, y, z); | 
| 171 | fun mk_triple2 (x, (y, z)) = mk_triple (x, y, z); | |
| 172 | ||
| 2231 
71385461570a
Eta-expansion of a function definition, for value polymorphism
 paulson parents: 
2203diff
changeset | 173 | fun split_decls l = flat (map (fn (xs, y) => map (rpair y) xs) l); | 
| 389 | 174 | |
| 175 | fun strip_quotes str = | |
| 176 | implode (tl (take (size str - 1, explode str))); | |
| 177 | ||
| 178 | ||
| 179 | (* names *) | |
| 180 | ||
| 181 | val name = ident >> quote || string; | |
| 182 | val names = list name; | |
| 183 | val names1 = list1 name; | |
| 184 | val name_list = names >> mk_list; | |
| 185 | val name_list1 = names1 >> mk_list; | |
| 186 | ||
| 187 | ||
| 188 | (* classes *) | |
| 189 | ||
| 190 | val subclass = name -- optional ("<" $$-- !! name_list1) "[]";
 | |
| 191 | ||
| 558 | 192 | val class_decls = repeat1 (subclass >> mk_pair) >> mk_big_list; | 
| 389 | 193 | |
| 194 | ||
| 195 | (* arities *) | |
| 196 | ||
| 197 | val sort = | |
| 198 | name >> brackets || | |
| 199 |   "{" $$-- name_list --$$ "}";
 | |
| 200 | ||
| 201 | val sort_list1 = list1 sort >> mk_list; | |
| 202 | ||
| 203 | ||
| 204 | val arity = optional ("(" $$-- !! (sort_list1 --$$")")) "[]" -- sort;
 | |
| 205 | ||
| 206 | val arity_decls = repeat1 (names1 --$$ "::" -- !! arity) | |
| 207 | >> (mk_big_list o map mk_triple2 o split_decls); | |
| 208 | ||
| 209 | ||
| 210 | (* mixfix annotations *) | |
| 211 | ||
| 2203 | 212 | val infxl = | 
| 213 | "infixl" $$-- !! (nat >> cat "Infixl" || string -- nat >> (cat "InfixlName" o mk_pair)); | |
| 214 | val infxr = | |
| 215 | "infixr" $$-- !! (nat >> cat "Infixr" || string -- nat >> (cat "InfixrName" o mk_pair)); | |
| 389 | 216 | |
| 889 | 217 | val binder = "binder" $$-- | 
| 2203 | 218 |   !! (string -- (("[" $$-- nat --$$ "]") -- nat || nat >> (fn n => (n, n))))
 | 
| 219 | >> (cat "Binder" o mk_triple2); | |
| 389 | 220 | |
| 221 | val opt_pris = optional ("[" $$-- !! (list nat --$$ "]")) [] >> mk_list;
 | |
| 222 | ||
| 223 | val mixfix = string -- !! (opt_pris -- optional nat "max_pri") | |
| 224 | >> (cat "Mixfix" o mk_triple2); | |
| 225 | ||
| 558 | 226 | fun opt_syn fx = optional ("(" $$-- fx --$$ ")") "NoSyn";
 | 
| 389 | 227 | |
| 228 | val opt_infix = opt_syn (infxl || infxr); | |
| 229 | val opt_mixfix = opt_syn (mixfix || infxl || infxr || binder); | |
| 230 | ||
| 231 | ||
| 232 | (* types *) | |
| 233 | ||
| 1705 
19fe0ab646b4
changed ident_no_colon so that it forbids postfix "=", too
 clasohm parents: 
1555diff
changeset | 234 | (*Parse an identifier, but only if it is not followed by "::", "=" or ","; | 
| 1377 | 235 | the exclusion of a postfix comma can be controlled to allow expressions | 
| 236 | like "(id, id)" but disallow ones like "'a => id id,id :: ..."*) | |
| 237 | fun ident_no_colon _ [] = eof_err() | |
| 238 | | ident_no_colon allow_comma ((Ident, s, n) :: (rest as (Keyword, s2, n2) :: | |
| 239 | toks)) = | |
| 1705 
19fe0ab646b4
changed ident_no_colon so that it forbids postfix "=", too
 clasohm parents: 
1555diff
changeset | 240 | if s2 = "::" orelse s2 = "=" orelse (not allow_comma andalso s2 = ",") | 
| 
19fe0ab646b4
changed ident_no_colon so that it forbids postfix "=", too
 clasohm parents: 
1555diff
changeset | 241 | then syn_err (name_of_kind Ident) (quote s2) n2 | 
| 1377 | 242 | else (s, rest) | 
| 243 | | ident_no_colon _ ((Ident, s, n) :: toks) = (s, toks) | |
| 244 | | ident_no_colon _ ((k, s, n) :: _) = | |
| 245 | 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: 
1250diff
changeset | 246 | |
| 3110 | 247 | (*type used in types, consts and syntax sections*) | 
| 1377 | 248 | fun const_type allow_comma toks = | 
| 249 | let val simple_type = | |
| 250 | (ident || | |
| 251 |          kind TypeVar -- optional ("::" $$-- ident >> cat "::") "" >>
 | |
| 252 | (fn (tv, cl) => cat tv cl)) -- | |
| 253 | repeat (ident_no_colon allow_comma) >> | |
| 254 | (fn (args, ts) => cat args (space_implode " " ts)) || | |
| 255 |          ("(" $$-- (list1 (const_type true)) --$$ ")" >> (parens o commas)) --
 | |
| 256 | repeat1 (ident_no_colon allow_comma) >> | |
| 257 | (fn (args, ts) => cat args (space_implode " " ts)); | |
| 258 | ||
| 259 | val appl_param = | |
| 260 |         simple_type || "(" $$-- const_type true --$$ ")" >> parens || 
 | |
| 261 | "[" $$-- (list1 (const_type true)) --$$ "]" --$$ "=>" -- | |
| 262 | const_type allow_comma >> | |
| 263 | (fn (src, dest) => mk_list src ^ " => " ^ dest); | |
| 264 |   in ("[" $$-- (list1 (const_type true)) --$$ "]" --$$ "=>" --
 | |
| 265 | const_type allow_comma >> | |
| 266 | (fn (src, dest) => mk_list src ^ " => " ^ dest) || | |
| 267 | repeat1 (appl_param --$$ "=>") -- const_type allow_comma >> | |
| 268 | (fn (src, dest) => space_implode " => " (src@[dest])) || | |
| 269 | simple_type || | |
| 270 |       "(" $$-- const_type true --$$ ")" >> parens) toks
 | |
| 1321 
9a6e7bd2bfaf
types in consts section of .thy files can now be specified without quotes
 clasohm parents: 
1250diff
changeset | 271 | end; | 
| 
9a6e7bd2bfaf
types in consts section of .thy files can now be specified without quotes
 clasohm parents: 
1250diff
changeset | 272 | |
| 1383 | 273 | fun mk_old_type_decl ((ts, n), syn) = | 
| 274 | map (fn t => (mk_triple (t, n, syn), false)) ts; | |
| 275 | ||
| 276 | fun mk_type_decl (((xs, t), None), syn) = | |
| 277 | [(mk_triple (t, string_of_int (length xs), syn), false)] | |
| 278 | | mk_type_decl (((xs, t), Some rhs), syn) = | |
| 279 | [(parens (commas [t, mk_list xs, rhs, syn]), true)]; | |
| 280 | ||
| 281 | fun mk_type_decls tys = | |
| 3764 | 282 | "|> Theory.add_types\n" ^ mk_big_list (keyfilter tys false) ^ "\n\n\ | 
| 283 | \|> Theory.add_tyabbrs\n" ^ mk_big_list (keyfilter tys true); | |
| 1383 | 284 | |
| 285 | ||
| 286 | val old_type_decl = names1 -- nat -- opt_infix >> mk_old_type_decl; | |
| 287 | ||
| 288 | val type_args = | |
| 289 | type_var >> (fn x => [x]) || | |
| 290 |   "(" $$-- !! (list1 type_var --$$ ")") ||
 | |
| 291 | empty >> K []; | |
| 292 | ||
| 293 | val type_decl = type_args -- name -- | |
| 294 |   optional ("=" $$-- (string || (const_type false >> quote)) >> Some) None
 | |
| 295 | -- opt_infix >> mk_type_decl; | |
| 296 | ||
| 2360 | 297 | val type_decls = | 
| 298 | repeat1 (old_type_decl || type_decl) >> (mk_type_decls o flat); | |
| 1383 | 299 | |
| 300 | ||
| 301 | (* consts *) | |
| 302 | ||
| 2360 | 303 | val const_decls = | 
| 304 | repeat1 | |
| 305 | (names1 --$$ "::" -- !! ((string || const_type false >> quote) -- opt_mixfix)) | |
| 306 | >> (mk_big_list o map mk_triple2 o split_decls); | |
| 389 | 307 | |
| 2360 | 308 | val opt_mode = | 
| 309 | optional | |
| 310 |     ("(" $$-- !! (name -- optional ($$ "output" >> K "false") "true" --$$ ")"))
 | |
| 311 |     ("\"\"", "true")
 | |
| 312 | >> mk_pair; | |
| 313 | ||
| 314 | val syntax_decls = opt_mode -- const_decls >> (fn (mode, txt) => mode ^ "\n" ^ txt); | |
| 2203 | 315 | |
| 389 | 316 | |
| 317 | (* translations *) | |
| 318 | ||
| 319 | val trans_pat = | |
| 320 |   optional ("(" $$-- !! (name --$$ ")")) "\"logic\"" -- string >> mk_pair;
 | |
| 321 | ||
| 322 | val trans_arrow = | |
| 3528 | 323 | $$ "=>" >> K "Syntax.ParseRule " || | 
| 324 | $$ "<=" >> K "Syntax.PrintRule " || | |
| 325 | $$ "==" >> K "Syntax.ParsePrintRule "; | |
| 389 | 326 | |
| 1810 
0eef167ebe1b
Translation infixes <->, etc., no longer available at top-level
 paulson parents: 
1705diff
changeset | 327 | val trans_line = | 
| 2203 | 328 | trans_pat -- !! (trans_arrow -- trans_pat) | 
| 329 | >> (fn (left, (arr, right)) => arr ^ mk_pair (left, right)); | |
| 1810 
0eef167ebe1b
Translation infixes <->, etc., no longer available at top-level
 paulson parents: 
1705diff
changeset | 330 | |
| 
0eef167ebe1b
Translation infixes <->, etc., no longer available at top-level
 paulson parents: 
1705diff
changeset | 331 | val trans_decls = repeat1 trans_line >> mk_big_list; | 
| 389 | 332 | |
| 333 | ||
| 334 | (* ML translations *) | |
| 335 | ||
| 336 | val trfun_defs = | |
| 337 | " val parse_ast_translation = [];\n\ | |
| 338 | \ val parse_translation = [];\n\ | |
| 339 | \ val print_translation = [];\n\ | |
| 2385 | 340 | \ val typed_print_translation = [];\n\ | 
| 2694 | 341 | \ val print_ast_translation = [];\n\ | 
| 342 | \ val token_translation = [];"; | |
| 389 | 343 | |
| 344 | val trfun_args = | |
| 345 | "(parse_ast_translation, parse_translation, \ | |
| 346 | \print_translation, print_ast_translation)"; | |
| 347 | ||
| 348 | ||
| 349 | (* axioms *) | |
| 350 | ||
| 351 | val mk_axms = mk_big_list o map (mk_pair o apfst quote); | |
| 352 | ||
| 353 | fun mk_axiom_decls axms = (mk_axms axms, map fst axms); | |
| 354 | ||
| 355 | val axiom_decls = repeat1 (ident -- !! string) >> mk_axiom_decls; | |
| 356 | ||
| 357 | ||
| 3813 | 358 | (* oracle *) | 
| 359 | ||
| 360 | val oracle_decl = (name --$$ "=" -- long_id) >> mk_pair; | |
| 361 | ||
| 362 | ||
| 1555 | 363 | (* combined consts and axioms *) | 
| 364 | ||
| 365 | fun mk_constaxiom_decls x = | |
| 366 | let | |
| 367 | val (axms_defs, axms_names) = | |
| 368 | mk_axiom_decls (map (fn ((id, _), def) => (id ^ "_def", def)) x); | |
| 369 | in ((mk_big_list o map mk_triple2 o map (apfst quote o fst)) x ^ | |
| 3764 | 370 | "\n\n|> Theory.add_defs\n" ^ axms_defs, axms_names) | 
| 1555 | 371 | end; | 
| 372 | ||
| 373 | val constaxiom_decls = | |
| 374 | repeat1 (ident --$$ "::" -- !! | |
| 375 | ((string || const_type false >> quote) -- opt_mixfix) -- !! | |
| 376 | string) >> mk_constaxiom_decls; | |
| 377 | ||
| 378 | ||
| 389 | 379 | (* axclass *) | 
| 380 | ||
| 381 | fun mk_axclass_decl ((c, cs), axms) = | |
| 382 | (mk_pair (c, cs) ^ "\n" ^ mk_axms axms, | |
| 383 | (strip_quotes c ^ "I") :: map fst axms); | |
| 384 | ||
| 385 | val axclass_decl = subclass -- repeat (ident -- !! string) >> mk_axclass_decl; | |
| 386 | ||
| 387 | ||
| 451 | 388 | (* instance *) | 
| 389 | 389 | |
| 425 | 390 | fun mk_witness (axths, opt_tac) = | 
| 389 | 391 | mk_list (keyfilter axths false) ^ "\n" ^ | 
| 392 | mk_list (keyfilter axths true) ^ "\n" ^ | |
| 393 | opt_tac; | |
| 394 | ||
| 395 | val axm_or_thm = | |
| 396 | string >> rpair false || | |
| 397 | long_id >> rpair true; | |
| 398 | ||
| 451 | 399 | |
| 425 | 400 | val opt_witness = | 
| 389 | 401 |   optional ("(" $$-- list1 axm_or_thm --$$ ")") [] --
 | 
| 558 | 402 | optional (verbatim >> (parens o cat "Some" o parens)) "None" | 
| 425 | 403 | >> mk_witness; | 
| 404 | ||
| 405 | val instance_decl = | |
| 636 | 406 | (name --$$ "<" -- name >> (pair "|> AxClass.add_inst_subclass" o mk_pair) || | 
| 407 | name --$$ "::" -- arity >> (pair "|> AxClass.add_inst_arity" o mk_triple2)) | |
| 451 | 408 | -- opt_witness | 
| 777 | 409 | >> (fn ((x, y), z) => (cat_lines [x, y, z])); | 
| 389 | 410 | |
| 411 | ||
| 412 | ||
| 413 | (** theory syntax **) | |
| 414 | ||
| 415 | type syntax = | |
| 416 | lexicon * (token list -> (string * string) * token list) Symtab.table; | |
| 417 | ||
| 418 | fun make_syntax keywords sects = | |
| 451 | 419 | (make_lexicon (map fst sects @ keywords), | 
| 420 | Symtab.make sects handle Symtab.DUPS dups => | |
| 421 |       error ("Duplicate sections in theory file syntax: " ^ commas_quote dups));
 | |
| 389 | 422 | |
| 423 | ||
| 424 | (* header *) | |
| 425 | ||
| 426 | 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: 
570diff
changeset | 427 | (thy_name, "mk_base " ^ mk_list bases ^ " " ^ quote thy_name); | 
| 389 | 428 | |
| 429 | val base = | |
| 430 | ident >> (cat "Thy" o quote) || | |
| 431 | string >> cat "File"; | |
| 432 | ||
| 433 | val header = ident --$$ "=" -- enum1 "+" base >> mk_header; | |
| 434 | ||
| 435 | ||
| 436 | (* extension *) | |
| 437 | ||
| 3813 | 438 | fun mk_extension ((begin, txts), mltxt) = | 
| 389 | 439 | let | 
| 440 | val cat_sects = space_implode "\n\n" o filter_out (equal ""); | |
| 441 | val (extxts, postxts) = split_list txts; | |
| 442 | in | |
| 3813 | 443 | (begin, cat_sects extxts, cat_sects postxts, mltxt) | 
| 389 | 444 | end; | 
| 445 | ||
| 446 | fun sect tab ((Keyword, s, n) :: toks) = | |
| 447 | (case Symtab.lookup (tab, s) of | |
| 448 | Some parse => !! parse toks | |
| 449 | | None => syn_err "section" s n) | |
| 450 | | sect _ ((_, s, n) :: _) = syn_err "section" s n | |
| 451 | | sect _ [] = eof_err (); | |
| 452 | ||
| 3813 | 453 | val opt_begin = | 
| 454 |   optional ("begin" $$-- optional name "" >> Some) None;
 | |
| 455 | ||
| 3798 | 456 | fun extension sectab = "+" $$-- !! | 
| 3813 | 457 | (opt_begin -- (repeat (sect sectab) --$$ "end") -- | 
| 3798 | 458 |     optional ("ML" $$-- verbatim) "") >> mk_extension;
 | 
| 389 | 459 | |
| 460 | ||
| 461 | (* theory definition *) | |
| 462 | ||
| 3798 | 463 | fun mk_structure tname ((thy_name, old_thys), opt_ext) = | 
| 558 | 464 | if thy_name <> tname then | 
| 465 |     error ("Filename \"" ^ tname ^ ".thy\" and theory name "
 | |
| 466 | ^ quote thy_name ^ " are different") | |
| 467 | else | |
| 3798 | 468 | (case opt_ext of | 
| 3813 | 469 | Some (begin, extxt, postxt, mltxt) => | 
| 587 
3ba470399605
renamed temporary variable 'base' to 'thy' in mk_structure
 clasohm parents: 
586diff
changeset | 470 | "val thy = " ^ old_thys ^ " true;\n\n\ | 
| 586 
201e115d8031
renamed base_on into mk_base and moved it to the beginning of the generated
 clasohm parents: 
570diff
changeset | 471 | \structure " ^ thy_name ^ " =\n\ | 
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 472 | \struct\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 473 | \\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 474 | \local\n" | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 475 | ^ trfun_defs ^ "\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 476 | \in\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 477 | \\n" | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 478 | ^ mltxt ^ "\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 479 | \\n\ | 
| 777 | 480 | \val thy = thy\n\ | 
| 3813 | 481 | \\n\ | 
| 482 | \|> Theory.add_path \"/\"\n" ^ | |
| 483 | (case begin of | |
| 484 |           None => (warning ("Flat name space for theory " ^ tname ^ "?  \
 | |
| 485 | \Consider begin ..."); "") | |
| 486 | | Some "" => "|> Theory.add_path " ^ quote tname ^ "\n" | |
| 487 | | Some path => "|> Theory.add_path " ^ path ^ "\n") ^ | |
| 3797 | 488 | "\n\ | 
| 3764 | 489 | \|> Theory.add_trfuns\n" | 
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 490 | ^ trfun_args ^ "\n\ | 
| 3764 | 491 | \|> Theory.add_trfunsT typed_print_translation \n\ | 
| 492 | \|> Theory.add_tokentrfuns token_translation \n\ | |
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 493 | \\n" | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 494 | ^ extxt ^ "\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 495 | \\n\ | 
| 3764 | 496 | \|> Theory.add_name " ^ quote thy_name ^ ";\n\ | 
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 497 | \\n\ | 
| 1235 | 498 | \val _ = store_theory (thy, " ^ quote thy_name ^ ");\n\ | 
| 777 | 499 | \\n\ | 
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 500 | \\n" | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 501 | ^ postxt ^ "\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 502 | \\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 503 | \end;\n\ | 
| 2253 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 504 | \end;\n\ | 
| 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 505 | \\n\ | 
| 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 506 | \open " ^ thy_name ^ ";\n\ | 
| 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 507 | \\n" | 
| 558 | 508 | | None => | 
| 777 | 509 | "val thy = " ^ old_thys ^ " false;\n\ | 
| 510 | \\n\ | |
| 586 
201e115d8031
renamed base_on into mk_base and moved it to the beginning of the generated
 clasohm parents: 
570diff
changeset | 511 | \structure " ^ thy_name ^ " =\n\ | 
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 512 | \struct\n\ | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 513 | \\n\ | 
| 587 
3ba470399605
renamed temporary variable 'base' to 'thy' in mk_structure
 clasohm parents: 
586diff
changeset | 514 | \val thy = thy\n\ | 
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 515 | \\n\ | 
| 1235 | 516 | \val _ = store_theory (thy, " ^ quote thy_name ^ ");\n\ | 
| 777 | 517 | \\n\ | 
| 2253 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 518 | \end;\n\ | 
| 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 519 | \\n\ | 
| 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 520 | \open " ^ thy_name ^ ";\n\ | 
| 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 521 | \\n"); | 
| 
95b615550b8b
use_thy now automatically opens theory structures;
 wenzelm parents: 
2231diff
changeset | 522 | |
| 389 | 523 | |
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 524 | fun theory_defn sectab tname = | 
| 3798 | 525 | header -- optional (extension sectab >> Some) None -- eof | 
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 526 | >> (mk_structure tname o #1); | 
| 389 | 527 | |
| 476 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 528 | fun parse_thy (lex, sectab) tname str = | 
| 
836cad329311
added check for concistency of filename and theory name;
 clasohm parents: 
451diff
changeset | 529 | #1 (!! (theory_defn sectab tname) (tokenize lex str)); | 
| 389 | 530 | |
| 531 | ||
| 532 | (* standard sections *) | |
| 533 | ||
| 534 | fun mk_val ax = "val " ^ ax ^ " = get_axiom thy " ^ quote ax ^ ";"; | |
| 777 | 535 | val mk_vals = cat_lines o map mk_val; | 
| 389 | 536 | |
| 777 | 537 | fun mk_axm_sect "" (txt, axs) = (txt, mk_vals axs) | 
| 538 | | mk_axm_sect pretxt (txt, axs) = (pretxt ^ "\n" ^ txt, mk_vals axs); | |
| 389 | 539 | |
| 540 | fun axm_section name pretxt parse = | |
| 541 | (name, parse >> mk_axm_sect pretxt); | |
| 542 | ||
| 543 | fun section name pretxt parse = | |
| 544 | axm_section name pretxt (parse >> rpair []); | |
| 545 | ||
| 546 | ||
| 547 | val pure_keywords = | |
| 3797 | 548 | ["end", "ML", "mixfix", "infixr", "infixl", "begin", "binder", "output", "=", | 
| 2360 | 549 |   "+", ",", "<", "{", "}", "(", ")", "[", "]", "::", "==", "=>", "<="];
 | 
| 389 | 550 | |
| 551 | val pure_sections = | |
| 3813 | 552 | [section "classes" "|> Theory.add_classes" class_decls, | 
| 3764 | 553 | section "default" "|> Theory.add_defsort" sort, | 
| 777 | 554 | section "types" "" type_decls, | 
| 3764 | 555 | section "arities" "|> Theory.add_arities" arity_decls, | 
| 556 | section "consts" "|> Theory.add_consts" const_decls, | |
| 557 | section "syntax" "|> Theory.add_modesyntax" syntax_decls, | |
| 558 | section "translations" "|> Theory.add_trrules" trans_decls, | |
| 559 | axm_section "rules" "|> Theory.add_axioms" axiom_decls, | |
| 560 | axm_section "defs" "|> Theory.add_defs" axiom_decls, | |
| 3813 | 561 | section "oracle" "|> Theory.add_oracle" oracle_decl, | 
| 3764 | 562 | axm_section "constdefs" "|> Theory.add_consts" constaxiom_decls, | 
| 636 | 563 | axm_section "axclass" "|> AxClass.add_axclass" axclass_decl, | 
| 3780 | 564 | section "instance" "" instance_decl, | 
| 565 | section "path" "|> Theory.add_path" name]; | |
| 389 | 566 | |
| 3813 | 567 | |
| 389 | 568 | end; |