src/Pure/Thy/thy_parse.ML
changeset 17927 4b42562ec171
parent 17926 8e12d3a4b890
child 17928 c567e5f885bf
equal deleted inserted replaced
17926:8e12d3a4b890 17927:4b42562ec171
     1 (*  Title:      Pure/Thy/thy_parse.ML
       
     2     ID:         $Id$
       
     3     Author:     Markus Wenzel, TU Muenchen
       
     4 
       
     5 The parser for *old-style* theory files.
       
     6 *)
       
     7 
       
     8 infix 5 -- --$$ $$-- ^^;
       
     9 infix 3 >>;
       
    10 infix 0 ||;
       
    11 
       
    12 signature THY_PARSE =
       
    13 sig
       
    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
       
    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
       
    50   type syntax
       
    51   val get_lexicon: syntax -> Scan.lexicon;
       
    52   val make_syntax: string list ->
       
    53     (string * (token list -> (string * string) * token list)) list -> syntax
       
    54   val parse_thy: syntax -> string list -> string
       
    55   val section: string -> string -> (token list -> string * token list)
       
    56     -> (string * (token list -> (string * string) * token list))
       
    57   val axm_section: string -> string
       
    58     -> (token list -> (string * string list) * token list)
       
    59     -> (string * (token list -> (string * string) * token list))
       
    60   val pure_keywords: string list
       
    61   val pure_sections:
       
    62     (string * (token list -> (string * string) * token list)) list
       
    63   (*items for building strings*)
       
    64   val cat: string -> string -> string
       
    65   val parens: string -> string
       
    66   val brackets: string -> string
       
    67   val mk_list: string list -> string
       
    68   val mk_big_list: string list -> string
       
    69   val mk_pair: string * string -> string
       
    70   val mk_triple: string * string * string -> string
       
    71   val mk_triple1: (string * string) * string -> string
       
    72   val mk_triple2: string * (string * string) -> string
       
    73 end;
       
    74 
       
    75 
       
    76 structure ThyParse : THY_PARSE=
       
    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 
       
    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");
       
   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 >> Library.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 
       
   156 fun list1 parse = enum1 "," parse;
       
   157 fun list parse = enum "," parse;
       
   158 
       
   159 
       
   160 
       
   161 (** theory parsers **)
       
   162 
       
   163 (* misc utilities *)
       
   164 
       
   165 fun cat s1 s2 = s1 ^ " " ^ s2;
       
   166 
       
   167 val parens = enclose "(" ")";
       
   168 val brackets = enclose "[" "]";
       
   169 
       
   170 val mk_list = brackets o commas;
       
   171 val mk_big_list = brackets o space_implode ",\n ";
       
   172 
       
   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);
       
   177 
       
   178 fun split_decls l = List.concat (map (fn (xs, y) => map (rpair y) xs) l);
       
   179 
       
   180 
       
   181 (* names *)
       
   182 
       
   183 val name = ident >> Library.quote || string;
       
   184 val names = list name;
       
   185 val names1 = list1 name;
       
   186 val name_list = names >> mk_list;
       
   187 val name_list1 = names1 >> mk_list;
       
   188 
       
   189 
       
   190 (* empty *)
       
   191 
       
   192 fun empty_decl toks = (empty >> K "") toks;
       
   193 
       
   194 
       
   195 (* classes *)
       
   196 
       
   197 val subclass = name -- optional ("<" $$-- !! name_list1) "[]";
       
   198 
       
   199 val class_decls = repeat1 (subclass >> mk_pair) >> mk_big_list;
       
   200 
       
   201 
       
   202 (* arities *)
       
   203 
       
   204 val sort = name || "{" $$-- list ident --$$ "}" >> (Library.quote o enclose "{" "}" o commas);
       
   205 val sort_list1 = list1 sort >> mk_list;
       
   206 
       
   207 val arity = optional ("(" $$-- !! (sort_list1 --$$")")) "[]" -- sort;
       
   208 
       
   209 val arity_decls = repeat1 (names1 --$$ "::" -- !! arity)
       
   210   >> (mk_big_list o map mk_triple2 o split_decls);
       
   211 
       
   212 
       
   213 (* mixfix annotations *)
       
   214 
       
   215 val infxl =
       
   216   "infixl" $$-- !! (nat >> cat "Infixl" || string -- nat >> (cat "InfixlName" o mk_pair));
       
   217 val infxr =
       
   218   "infixr" $$-- !! (nat >> cat "Infixr" || string -- nat >> (cat "InfixrName" o mk_pair));
       
   219 
       
   220 val binder = "binder" $$--
       
   221   !! (string -- (("[" $$-- nat --$$ "]") -- nat || nat >> (fn n => (n, n))))
       
   222     >> (cat "Binder" o mk_triple2);
       
   223 
       
   224 val opt_pris = optional ("[" $$-- !! (list nat --$$ "]")) [] >> mk_list;
       
   225 
       
   226 val mixfix = string -- !! (opt_pris -- optional nat "Syntax.max_pri")
       
   227   >> (cat "Mixfix" o mk_triple2);
       
   228 
       
   229 fun opt_syn fx = optional ("(" $$-- fx --$$ ")") "NoSyn";
       
   230 
       
   231 val opt_infix = opt_syn (infxl || infxr);
       
   232 val opt_mixfix = opt_syn (mixfix || infxl || infxr || binder);
       
   233 
       
   234 
       
   235 (* types *)
       
   236 
       
   237 (*Parse an identifier, but only if it is not followed by "::", "=" or ",";
       
   238   the exclusion of a postfix comma can be controlled to allow expressions
       
   239   like "(id, id)" but disallow ones like "'a => id id,id :: ..."*)
       
   240 fun ident_no_colon _ [] = eof_err()
       
   241   | ident_no_colon allow_comma ((Ident, s, n) :: (rest as (Keyword, s2, n2) ::
       
   242                                 toks)) =
       
   243       if s2 = "::" orelse s2 = "=" orelse (not allow_comma andalso s2 = ",")
       
   244       then syn_err (name_of_kind Ident) (quote s2) n2
       
   245       else (s, rest)
       
   246   | ident_no_colon _ ((Ident, s, n) :: toks) = (s, toks)
       
   247   | ident_no_colon _ ((k, s, n) :: _) =
       
   248       syn_err (name_of_kind Ident) (quote s) n;
       
   249 
       
   250 (*type used in types, consts and syntax sections*)
       
   251 fun const_type allow_comma toks =
       
   252   let
       
   253     val simple_type =
       
   254       (ident || kind TypeVar ^^ optional ($$ "::" ^^ ident) "") --
       
   255           repeat (ident_no_colon allow_comma)
       
   256           >> (fn (args, ts) => cat args (space_implode " " ts)) ||
       
   257         ("(" $$-- (list1 (const_type true)) --$$ ")" >> (parens o commas)) --
       
   258           repeat1 (ident_no_colon allow_comma)
       
   259           >> (fn (args, ts) => cat args (space_implode " " ts));
       
   260 
       
   261       val appl_param =
       
   262         simple_type || "(" $$-- const_type true --$$ ")" >> parens || 
       
   263         "[" $$-- (list1 (const_type true)) --$$ "]" --$$ "=>" --
       
   264           const_type allow_comma >>
       
   265           (fn (src, dest) => mk_list src ^ " => " ^ dest);
       
   266   in ("[" $$-- (list1 (const_type true)) --$$ "]" --$$ "=>" --
       
   267         const_type allow_comma >>
       
   268         (fn (src, dest) => mk_list src ^ " => " ^ dest) ||
       
   269       repeat1 (appl_param --$$ "=>") -- const_type allow_comma >>
       
   270         (fn (src, dest) => space_implode " => " (src@[dest])) ||
       
   271       simple_type ||
       
   272       "(" $$-- const_type true --$$ ")" >> parens) toks
       
   273   end;
       
   274 
       
   275 val typ = string || (const_type false >> Library.quote);
       
   276 
       
   277 
       
   278 fun mk_old_type_decl ((ts, n), syn) =
       
   279   map (fn t => (mk_triple (t, n, syn), false)) ts;
       
   280 
       
   281 fun mk_type_decl (((xs, t), NONE), syn) =
       
   282       [(mk_triple (t, string_of_int (length xs), syn), false)]
       
   283   | mk_type_decl (((xs, t), SOME rhs), syn) =
       
   284       [(parens (commas [t, mk_list xs, rhs, syn]), true)];
       
   285 
       
   286 fun mk_type_decls tys =
       
   287   "|> Theory.add_types\n" ^ mk_big_list (AList.find (op =) tys false) ^ "\n\n\
       
   288   \|> Theory.add_tyabbrs\n" ^ mk_big_list (AList.find (op =) tys true);
       
   289 
       
   290 
       
   291 val old_type_decl = names1 -- nat -- opt_infix >> mk_old_type_decl;
       
   292 
       
   293 val type_args =
       
   294   type_var >> (fn x => [x]) ||
       
   295   "(" $$-- !! (list1 type_var --$$ ")") ||
       
   296   empty >> K [];
       
   297 
       
   298 val type_decl = type_args -- name --
       
   299   optional ("=" $$-- typ >> SOME) NONE -- opt_infix >> mk_type_decl;
       
   300 
       
   301 val type_decls =
       
   302   repeat1 (old_type_decl || type_decl) >> (mk_type_decls o List.concat);
       
   303 
       
   304 
       
   305 (* consts *)
       
   306 
       
   307 val const_decls =
       
   308   repeat1 (names1 --$$ "::" -- !! (typ -- opt_mixfix))
       
   309   >> (mk_big_list o map mk_triple2 o split_decls);
       
   310 
       
   311 val opt_mode =
       
   312   optional
       
   313     ("(" $$-- !! (name -- optional ($$ "output" >> K "false") "true" --$$ ")"))
       
   314     ("\"\"", "true")
       
   315   >> mk_pair;
       
   316 
       
   317 val syntax_decls = opt_mode -- const_decls >> (fn (mode, txt) => mode ^ "\n" ^ txt);
       
   318 
       
   319 
       
   320 (* translations *)
       
   321 
       
   322 val trans_pat =
       
   323   optional ("(" $$-- !! (name --$$ ")")) "\"logic\"" -- string >> mk_pair;
       
   324 
       
   325 val trans_arrow =
       
   326   $$ "=>" >> K "Syntax.ParseRule " ||
       
   327   $$ "<=" >> K "Syntax.PrintRule " ||
       
   328   $$ "==" >> K "Syntax.ParsePrintRule ";
       
   329 
       
   330 val trans_line =
       
   331   trans_pat -- !! (trans_arrow -- trans_pat)
       
   332     >> (fn (left, (arr, right)) => arr ^ mk_pair (left, right));
       
   333 
       
   334 val trans_decls = repeat1 trans_line >> mk_big_list;
       
   335 
       
   336 
       
   337 (* ML translations *)
       
   338 
       
   339 val local_defs =
       
   340   " val parse_ast_translation = [];\n\
       
   341   \ val parse_translation = [];\n\
       
   342   \ val print_translation = [];\n\
       
   343   \ val typed_print_translation = [];\n\
       
   344   \ val print_ast_translation = [];\n\
       
   345   \ val token_translation = [];";
       
   346 
       
   347 val trfun_args =
       
   348   "(parse_ast_translation, parse_translation, \
       
   349   \print_translation, print_ast_translation)";
       
   350 
       
   351 
       
   352 (* axioms *)
       
   353 
       
   354 val mk_axms = mk_big_list o map (mk_pair o apfst Library.quote);
       
   355 val mk_axms' = mk_big_list o map (mk_pair o rpair "[]" o mk_pair o apfst Library.quote);
       
   356 
       
   357 fun mk_axiom_decls axms = (mk_axms axms, map fst axms);
       
   358 
       
   359 val axiom_decls = repeat1 (ident -- !! string) >> mk_axiom_decls;
       
   360 
       
   361 
       
   362 (* oracle *)
       
   363 
       
   364 val oracle_decl = (name --$$ "=" -- long_id) >> mk_pair;
       
   365 
       
   366 
       
   367 (* combined consts and axioms *)
       
   368 
       
   369 fun mk_constaxiom_decls x =
       
   370   let
       
   371     val (axms_defs, axms_names) =
       
   372       mk_axiom_decls (map (fn ((id, _), def) => (id ^ "_def", def)) x);
       
   373   in ((mk_big_list o map mk_triple2 o map (apfst Library.quote o fst)) x ^
       
   374        "\n\n|> (#1 oo (PureThy.add_defs false o map Thm.no_attributes))\n" ^ axms_defs, axms_names)
       
   375   end;
       
   376 
       
   377 val constaxiom_decls =
       
   378   repeat1 (ident --$$ "::" -- !! (typ -- opt_mixfix) -- !! string)
       
   379   >> mk_constaxiom_decls;
       
   380 
       
   381 
       
   382 (* axclass *)
       
   383 
       
   384 fun mk_axclass_decl ((c, cs), axms) =
       
   385   (mk_pair (c, cs) ^ "\n" ^ mk_axms' axms, map fst axms);
       
   386 
       
   387 val axclass_decl = subclass -- repeat (ident -- !! string) >> mk_axclass_decl;
       
   388 
       
   389 
       
   390 (* instance *)
       
   391 
       
   392 fun mk_witness (axths, opt_tac) =
       
   393   mk_list (AList.find (op =) axths false) ^ "\n" ^
       
   394   mk_list (AList.find (op =) axths true) ^ "\n" ^
       
   395   opt_tac;
       
   396 
       
   397 val axm_or_thm =
       
   398   string >> rpair false ||
       
   399   long_id >> rpair true;
       
   400 
       
   401 
       
   402 val opt_witness =
       
   403   optional ("(" $$-- list1 axm_or_thm --$$ ")") [] --
       
   404   optional (verbatim >> (parens o cat "SOME" o parens)) "NONE"
       
   405   >> mk_witness;
       
   406 
       
   407 val instance_decl =
       
   408   (name --$$ "<" -- name >> (pair "|> AxClass.add_inst_subclass_x" o mk_pair) ||
       
   409     name --$$ "::" -- arity >> (pair "|> AxClass.add_inst_arity_x" o mk_triple2))
       
   410   -- opt_witness
       
   411   >> (fn ((x, y), z) => (cat_lines [x, y, z]));
       
   412 
       
   413 
       
   414 (* locale *)
       
   415 
       
   416 val locale_decl =
       
   417   (name --$$ "=") -- 
       
   418     (optional ((ident >> (fn x => parens ("SOME" ^ Library.quote x))) --$$ "+") ("NONE")) --
       
   419     ("fixes" $$--
       
   420       (repeat (name --$$ "::" -- !! (typ -- opt_mixfix)) 
       
   421        >> (mk_big_list o map mk_triple2))) --
       
   422     (optional 
       
   423      ("assumes" $$-- (repeat ((ident >> Library.quote) -- !! string) 
       
   424 		     >> (mk_list o map mk_pair)))
       
   425      "[]") --
       
   426     (optional 
       
   427      ("defines" $$-- (repeat ((ident >> Library.quote) -- !! string) 
       
   428 		      >> (mk_list o map mk_pair)))
       
   429      "[]")
       
   430   >> (fn ((((nm, ext), cs), asms), defs) => cat_lines [nm, ext, cs, asms, defs]);
       
   431 
       
   432 
       
   433 
       
   434 (** theory syntax **)
       
   435 
       
   436 type syntax =
       
   437   Scan.lexicon * (token list -> (string * string) * token list) Symtab.table;
       
   438 
       
   439 fun get_lexicon (lex, _) = lex;
       
   440 
       
   441 fun make_syntax keywords sects =
       
   442   let
       
   443     val dups = duplicates (map fst sects);
       
   444     val sects' = gen_distinct (eq_fst (op =)) sects;
       
   445     val keys = map Symbol.explode (map fst sects' @ keywords);
       
   446   in
       
   447     if null dups then ()
       
   448     else warning ("Duplicate declaration of theory file section(s): " ^ commas_quote dups);
       
   449     (Scan.make_lexicon keys, Symtab.make sects')
       
   450   end;
       
   451 
       
   452 
       
   453 (* header *)
       
   454 
       
   455 fun mk_header (thy_name, parents) =
       
   456   (thy_name, "IsarThy.begin_theory " ^ cat (Library.quote thy_name) (mk_list parents) ^ " [] false");
       
   457 
       
   458 val header = ident --$$ "=" -- enum1 "+" name >> mk_header;
       
   459 
       
   460 
       
   461 (* extension *)
       
   462 
       
   463 fun mk_extension (txts, mltxt) =
       
   464   let
       
   465     val cat_sects = space_implode "\n\n" o filter_out (equal "");
       
   466     val (extxts, postxts) = split_list txts;
       
   467   in
       
   468     (cat_sects extxts, cat_sects postxts, mltxt)
       
   469   end;
       
   470 
       
   471 fun sect tab ((Keyword, s, n) :: toks) =
       
   472       (case Symtab.lookup tab s of
       
   473         SOME parse => !! parse toks
       
   474       | NONE => syn_err "section" s n)
       
   475   | sect _ ((_, s, n) :: _) = syn_err "section" s n
       
   476   | sect _ [] = eof_err ();
       
   477 
       
   478 fun extension sectab = "+" $$-- !!
       
   479   (repeat (sect sectab) --$$ "end" -- optional verbatim "")
       
   480     >> mk_extension;
       
   481 
       
   482 fun opt_extension sectab = optional (extension sectab) ("", "", "");
       
   483 
       
   484 
       
   485 (* theory definition *)
       
   486 
       
   487 fun mk_structure ((thy_name, bg_theory), (extxt, postxt, mltxt)) =
       
   488   "structure " ^ thy_name ^ " =\n\
       
   489   \struct\n\
       
   490   \\n\
       
   491   \local\n"
       
   492   ^ local_defs ^ "\n\
       
   493   \in\n\
       
   494   \\n"
       
   495   ^ mltxt ^ "\n\
       
   496   \\n\
       
   497   \val thy = " ^ bg_theory ^ "\n\
       
   498   \\n\
       
   499   \|> Theory.add_trfuns\n"
       
   500   ^ trfun_args ^ "\n\
       
   501   \|> Theory.add_trfunsT typed_print_translation\n\
       
   502   \|> Theory.add_tokentrfuns token_translation\n\
       
   503   \\n"
       
   504   ^ extxt ^ "\n\
       
   505   \\n\
       
   506   \|> IsarThy.end_theory;\n\
       
   507   \\n\
       
   508   \val _ = context thy;\n\
       
   509   \\n\
       
   510   \\n"
       
   511   ^ postxt ^ "\n\
       
   512   \\n\
       
   513   \end;\n\
       
   514   \end;\n\
       
   515   \\n\
       
   516   \open " ^ thy_name ^ ";\n";
       
   517 
       
   518 fun theory_defn sectab =
       
   519   header -- opt_extension sectab -- eof >> (mk_structure o #1);
       
   520 
       
   521 fun parse_thy (lex, sectab) chs = #1 (!! (theory_defn sectab) (tokenize lex chs));
       
   522 
       
   523 
       
   524 (* standard sections *)
       
   525 
       
   526 fun mk_val ax = "val " ^ ax ^ " = PureThy.get_thm thy (PureThy.Name " ^ Library.quote ax ^ ");";
       
   527 val mk_vals = cat_lines o map mk_val;
       
   528 
       
   529 fun mk_axm_sect "" (txt, axs) = (txt, mk_vals axs)
       
   530   | mk_axm_sect pretxt (txt, axs) = (pretxt ^ "\n" ^ txt, mk_vals axs);
       
   531 
       
   532 fun axm_section name pretxt parse =
       
   533   (name, parse >> mk_axm_sect pretxt);
       
   534 
       
   535 fun section name pretxt parse =
       
   536   axm_section name pretxt (parse >> rpair []);
       
   537 
       
   538 
       
   539 val pure_keywords =
       
   540  ["end", "ML", "mixfix", "infixr", "infixl", "binder", "output", "=",
       
   541   "+", ",", "<", "{", "}", "(", ")", "[", "]", "::", "==", "=>",
       
   542   "<=", "fixes", "assumes", "defines"];
       
   543 
       
   544 val pure_sections =
       
   545  [section "classes" "|> Theory.add_classes" class_decls,
       
   546   section "default" "|> Theory.add_defsort" sort,
       
   547   section "types" "" type_decls,
       
   548   section "nonterminals" "|> Theory.add_nonterminals" (repeat1 name >> mk_list),
       
   549   section "arities" "|> Theory.add_arities" arity_decls,
       
   550   section "consts" "|> Theory.add_consts" const_decls,
       
   551   section "syntax" "|> Theory.add_modesyntax" syntax_decls,
       
   552   section "translations" "|> Theory.add_trrules" trans_decls,
       
   553   axm_section "rules" "|> (#1 oo (PureThy.add_axioms o map Thm.no_attributes))" axiom_decls,
       
   554   axm_section "defs" "|> (#1 oo (PureThy.add_defs false o map Thm.no_attributes))" axiom_decls,
       
   555   section "oracle" "|> Theory.add_oracle" oracle_decl,
       
   556   axm_section "constdefs" "|> Theory.add_consts" constaxiom_decls,
       
   557   axm_section "axclass" "|> (#1 ooo AxClass.add_axclass)" axclass_decl,
       
   558   section "instance" "" instance_decl,
       
   559   section "path" "|> Theory.add_path" name,
       
   560   section "global" "|> Sign.root_path" empty_decl,
       
   561   section "local" "|> Sign.local_path" empty_decl,
       
   562   section "setup" "|> Library.apply" long_id,
       
   563   section "MLtext" "" verbatim,
       
   564   section "locale" "|> Goals.add_locale" locale_decl];
       
   565 
       
   566 
       
   567 end;