src/Pure/Syntax/syntax.ML
author wenzelm
Tue, 27 Nov 2007 16:48:38 +0100
changeset 25476 03da46cfab9e
parent 25394 db25c98f32e1
child 25993 d542486e39e7
permissions -rw-r--r--
standard_parse_term: check ambiguous results without changing the result yet;
     1 (*  Title:      Pure/Syntax/syntax.ML
     2     ID:         $Id$
     3     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
     4 
     5 Standard Isabelle syntax, based on arbitrary context-free grammars
     6 (specified by mixfix declarations).
     7 *)
     8 
     9 signature BASIC_SYNTAX =
    10 sig
    11   include AST0
    12   include SYN_TRANS0
    13   include MIXFIX0
    14   include PRINTER0
    15 end;
    16 
    17 signature SYNTAX =
    18 sig
    19   include AST1
    20   include LEXICON0
    21   include SYN_EXT0
    22   include TYPE_EXT0
    23   include SYN_TRANS1
    24   include MIXFIX1
    25   include PRINTER0
    26   datatype 'a trrule =
    27     ParseRule of 'a * 'a |
    28     PrintRule of 'a * 'a |
    29     ParsePrintRule of 'a * 'a
    30   type syntax
    31   val eq_syntax: syntax * syntax -> bool
    32   val is_keyword: syntax -> string -> bool
    33   type mode
    34   val mode_default: mode
    35   val mode_input: mode
    36   val update_type_gram: (string * int * mixfix) list -> syntax -> syntax
    37   val update_consts: string list -> syntax -> syntax
    38   val update_trfuns:
    39     (string * ((ast list -> ast) * stamp)) list *
    40     (string * ((term list -> term) * stamp)) list *
    41     (string * ((bool -> typ -> term list -> term) * stamp)) list *
    42     (string * ((ast list -> ast) * stamp)) list -> syntax -> syntax
    43   val update_advanced_trfuns:
    44     (string * ((Proof.context -> ast list -> ast) * stamp)) list *
    45     (string * ((Proof.context -> term list -> term) * stamp)) list *
    46     (string * ((Proof.context -> bool -> typ -> term list -> term) * stamp)) list *
    47     (string * ((Proof.context -> ast list -> ast) * stamp)) list -> syntax -> syntax
    48   val extend_tokentrfuns: (string * string * (string -> output * int)) list -> syntax -> syntax
    49   val update_const_gram: (string -> bool) ->
    50     mode -> (string * typ * mixfix) list -> syntax -> syntax
    51   val remove_const_gram: (string -> bool) ->
    52     mode -> (string * typ * mixfix) list -> syntax -> syntax
    53   val update_trrules: Proof.context -> (string -> bool) -> syntax ->
    54     (string * string) trrule list -> syntax -> syntax
    55   val remove_trrules: Proof.context -> (string -> bool) -> syntax ->
    56     (string * string) trrule list -> syntax -> syntax
    57   val update_trrules_i: ast trrule list -> syntax -> syntax
    58   val remove_trrules_i: ast trrule list -> syntax -> syntax
    59   val map_trrule: ('a -> 'b) -> 'a trrule -> 'b trrule
    60   val merge_syntaxes: syntax -> syntax -> syntax
    61   val basic_syn: syntax
    62   val basic_nonterms: string list
    63   val print_gram: syntax -> unit
    64   val print_trans: syntax -> unit
    65   val print_syntax: syntax -> unit
    66   val read: Proof.context -> (string -> bool) -> syntax -> typ -> string -> term list
    67   val standard_parse_term: Pretty.pp -> (term -> string option) ->
    68     (((string * int) * sort) list -> string * int -> Term.sort) ->
    69     (string -> string option) -> (string -> string option) ->
    70     (typ -> typ) -> (sort -> sort) -> Proof.context ->
    71     (string -> bool) -> syntax -> typ -> string -> term
    72   val standard_parse_typ: Proof.context -> syntax ->
    73     ((indexname * sort) list -> indexname -> sort) -> (sort -> sort) -> string -> typ
    74   val standard_parse_sort: Proof.context -> syntax -> (sort -> sort) -> string -> sort
    75   val standard_unparse_term: (string -> xstring) ->
    76     Proof.context -> syntax -> bool -> term -> Pretty.T
    77   val standard_unparse_typ: Proof.context -> syntax -> typ -> Pretty.T
    78   val standard_unparse_sort: Proof.context -> syntax -> sort -> Pretty.T
    79   val ambiguity_level: int ref
    80   val ambiguity_is_error: bool ref
    81   val parse_sort: Proof.context -> string -> sort
    82   val parse_typ: Proof.context -> string -> typ
    83   val parse_term: Proof.context -> string -> term
    84   val parse_prop: Proof.context -> string -> term
    85   val unparse_sort: Proof.context -> sort -> Pretty.T
    86   val unparse_classrel: Proof.context -> class list -> Pretty.T
    87   val unparse_arity: Proof.context -> arity -> Pretty.T
    88   val unparse_typ: Proof.context -> typ -> Pretty.T
    89   val unparse_term: Proof.context -> term -> Pretty.T
    90   val install_operations:
    91    {parse_sort: Proof.context -> string -> sort,
    92     parse_typ: Proof.context -> string -> typ,
    93     parse_term: Proof.context -> string -> term,
    94     parse_prop: Proof.context -> string -> term,
    95     unparse_sort: Proof.context -> sort -> Pretty.T,
    96     unparse_typ: Proof.context -> typ -> Pretty.T,
    97     unparse_term: Proof.context -> term -> Pretty.T} -> unit
    98   val print_checks: Proof.context -> unit
    99   val add_typ_check: int -> string ->
   100     (typ list -> Proof.context -> (typ list * Proof.context) option) ->
   101     Context.generic -> Context.generic
   102   val add_term_check: int -> string ->
   103     (term list -> Proof.context -> (term list * Proof.context) option) ->
   104     Context.generic -> Context.generic
   105   val add_typ_uncheck: int -> string ->
   106     (typ list -> Proof.context -> (typ list * Proof.context) option) ->
   107     Context.generic -> Context.generic
   108   val add_term_uncheck: int -> string ->
   109     (term list -> Proof.context -> (term list * Proof.context) option) ->
   110     Context.generic -> Context.generic
   111   val check_sort: Proof.context -> sort -> sort
   112   val check_typ: Proof.context -> typ -> typ
   113   val check_term: Proof.context -> term -> term
   114   val check_prop: Proof.context -> term -> term
   115   val check_typs: Proof.context -> typ list -> typ list
   116   val check_terms: Proof.context -> term list -> term list
   117   val check_props: Proof.context -> term list -> term list
   118   val uncheck_sort: Proof.context -> sort -> sort
   119   val uncheck_arity: Proof.context -> arity -> arity
   120   val uncheck_classrel: Proof.context -> class list -> class list
   121   val uncheck_typs: Proof.context -> typ list -> typ list
   122   val uncheck_terms: Proof.context -> term list -> term list
   123   val read_sort: Proof.context -> string -> sort
   124   val read_typ: Proof.context -> string -> typ
   125   val read_term: Proof.context -> string -> term
   126   val read_prop: Proof.context -> string -> term
   127   val read_terms: Proof.context -> string list -> term list
   128   val read_props: Proof.context -> string list -> term list
   129   val read_sort_global: theory -> string -> sort
   130   val read_typ_global: theory -> string -> typ
   131   val read_term_global: theory -> string -> term
   132   val read_prop_global: theory -> string -> term
   133   val pretty_term: Proof.context -> term -> Pretty.T
   134   val pretty_typ: Proof.context -> typ -> Pretty.T
   135   val pretty_sort: Proof.context -> sort -> Pretty.T
   136   val pretty_classrel: Proof.context -> class list -> Pretty.T
   137   val pretty_arity: Proof.context -> arity -> Pretty.T
   138   val string_of_term: Proof.context -> term -> string
   139   val string_of_typ: Proof.context -> typ -> string
   140   val string_of_sort: Proof.context -> sort -> string
   141   val string_of_classrel: Proof.context -> class list -> string
   142   val string_of_arity: Proof.context -> arity -> string
   143   val pp: Proof.context -> Pretty.pp
   144 end;
   145 
   146 structure Syntax: SYNTAX =
   147 struct
   148 
   149 (** tables of translation functions **)
   150 
   151 (* parse (ast) translations *)
   152 
   153 fun lookup_tr tab c = Option.map fst (Symtab.lookup tab c);
   154 
   155 fun err_dup_trfun name c =
   156   error ("More than one " ^ name ^ " for " ^ quote c);
   157 
   158 fun remove_trtab trfuns = fold (Symtab.remove SynExt.eq_trfun) trfuns;
   159 
   160 fun update_trtab name trfuns tab = Symtab.extend (remove_trtab trfuns tab, trfuns)
   161   handle Symtab.DUP c => err_dup_trfun name c;
   162 
   163 fun merge_trtabs name tab1 tab2 = Symtab.merge SynExt.eq_trfun (tab1, tab2)
   164   handle Symtab.DUP c => err_dup_trfun name c;
   165 
   166 
   167 (* print (ast) translations *)
   168 
   169 fun lookup_tr' tab c = map fst (Symtab.lookup_list tab c);
   170 fun update_tr'tab trfuns = fold_rev (Symtab.update_list SynExt.eq_trfun) trfuns;
   171 fun remove_tr'tab trfuns = fold (Symtab.remove_list SynExt.eq_trfun) trfuns;
   172 fun merge_tr'tabs tab1 tab2 = Symtab.merge_list SynExt.eq_trfun (tab1, tab2);
   173 
   174 
   175 
   176 (** tables of token translation functions **)
   177 
   178 fun lookup_tokentr tabs modes =
   179   let val trs = distinct (eq_fst (op =)) (maps (these o AList.lookup (op =) tabs) (modes @ [""]))
   180   in fn c => Option.map fst (AList.lookup (op =) trs c) end;
   181 
   182 fun merge_tokentrtabs tabs1 tabs2 =
   183   let
   184     fun eq_tr ((c1, (_, s1)), (c2, (_, s2))) = c1 = c2 andalso s1 = s2;
   185 
   186     fun name (s, _) = implode (tl (Symbol.explode s));
   187 
   188     fun merge mode =
   189       let
   190         val trs1 = these (AList.lookup (op =) tabs1 mode);
   191         val trs2 = these (AList.lookup (op =) tabs2 mode);
   192         val trs = distinct eq_tr (trs1 @ trs2);
   193       in
   194         (case duplicates (eq_fst (op =)) trs of
   195           [] => (mode, trs)
   196         | dups => error ("More than one token translation function in mode " ^
   197             quote mode ^ " for " ^ commas_quote (map name dups)))
   198       end;
   199   in map merge (distinct (op =) (map fst (tabs1 @ tabs2))) end;
   200 
   201 fun extend_tokentrtab tokentrs tabs =
   202   let
   203     fun ins_tokentr (m, c, f) =
   204       AList.default (op =) (m, [])
   205       #> AList.map_entry (op =) m (cons ("_" ^ c, (f, stamp ())));
   206   in merge_tokentrtabs tabs (fold ins_tokentr tokentrs []) end;
   207 
   208 
   209 
   210 (** tables of translation rules **)
   211 
   212 type ruletab = (Ast.ast * Ast.ast) list Symtab.table;
   213 
   214 fun dest_ruletab tab = maps snd (Symtab.dest tab);
   215 
   216 
   217 (* empty, update, merge ruletabs *)
   218 
   219 val update_ruletab = fold_rev (fn r => Symtab.update_list (op =) (Ast.head_of_rule r, r));
   220 val remove_ruletab = fold (fn r => Symtab.remove_list (op =) (Ast.head_of_rule r, r));
   221 fun merge_ruletabs tab1 tab2 = Symtab.merge_list (op =) (tab1, tab2);
   222 
   223 
   224 
   225 (** datatype syntax **)
   226 
   227 datatype syntax =
   228   Syntax of {
   229     input: SynExt.xprod list,
   230     lexicon: Scan.lexicon,
   231     gram: Parser.gram,
   232     consts: string list,
   233     prmodes: string list,
   234     parse_ast_trtab: ((Proof.context -> Ast.ast list -> Ast.ast) * stamp) Symtab.table,
   235     parse_ruletab: ruletab,
   236     parse_trtab: ((Proof.context -> term list -> term) * stamp) Symtab.table,
   237     print_trtab: ((Proof.context -> bool -> typ -> term list -> term) * stamp) list Symtab.table,
   238     print_ruletab: ruletab,
   239     print_ast_trtab: ((Proof.context -> Ast.ast list -> Ast.ast) * stamp) list Symtab.table,
   240     tokentrtab: (string * (string * ((string -> output * int) * stamp)) list) list,
   241     prtabs: Printer.prtabs} * stamp;
   242 
   243 fun eq_syntax (Syntax (_, s1), Syntax (_, s2)) = s1 = s2;
   244 
   245 fun is_keyword (Syntax ({lexicon, ...}, _)) = Scan.is_literal lexicon o Symbol.explode;
   246 
   247 type mode = string * bool;
   248 val mode_default = ("", true);
   249 val mode_input = (PrintMode.input, true);
   250 
   251 
   252 (* empty_syntax *)
   253 
   254 val empty_syntax = Syntax
   255   ({input = [],
   256     lexicon = Scan.empty_lexicon,
   257     gram = Parser.empty_gram,
   258     consts = [],
   259     prmodes = [],
   260     parse_ast_trtab = Symtab.empty,
   261     parse_ruletab = Symtab.empty,
   262     parse_trtab = Symtab.empty,
   263     print_trtab = Symtab.empty,
   264     print_ruletab = Symtab.empty,
   265     print_ast_trtab = Symtab.empty,
   266     tokentrtab = [],
   267     prtabs = Printer.empty_prtabs}, stamp ());
   268 
   269 
   270 (* update_syntax *)
   271 
   272 fun update_syntax (mode, inout) syn_ext (Syntax (tabs, _)) =
   273   let
   274     val {input, lexicon, gram, consts = consts1, prmodes = prmodes1,
   275       parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab, print_ruletab,
   276       print_ast_trtab, tokentrtab, prtabs} = tabs;
   277     val SynExt.SynExt {xprods, consts = consts2, prmodes = prmodes2,
   278       parse_ast_translation, parse_rules, parse_translation, print_translation, print_rules,
   279       print_ast_translation, token_translation} = syn_ext;
   280     val input' = if inout then fold (insert (op =)) xprods input else input;
   281     val changed = length input <> length input';
   282     fun if_inout xs = if inout then xs else [];
   283   in
   284     Syntax
   285     ({input = input',
   286       lexicon = if changed then Scan.extend_lexicon (SynExt.delims_of xprods) lexicon else lexicon,
   287       gram = if changed then Parser.extend_gram gram xprods else gram,
   288       consts = Library.merge (op =) (consts1, filter_out NameSpace.is_qualified consts2),
   289       prmodes = insert (op =) mode (Library.merge (op =) (prmodes1, prmodes2)),
   290       parse_ast_trtab =
   291         update_trtab "parse ast translation" (if_inout parse_ast_translation) parse_ast_trtab,
   292       parse_ruletab = update_ruletab (if_inout parse_rules) parse_ruletab,
   293       parse_trtab = update_trtab "parse translation" (if_inout parse_translation) parse_trtab,
   294       print_trtab = update_tr'tab print_translation print_trtab,
   295       print_ruletab = update_ruletab print_rules print_ruletab,
   296       print_ast_trtab = update_tr'tab print_ast_translation print_ast_trtab,
   297       tokentrtab = extend_tokentrtab token_translation tokentrtab,
   298       prtabs = Printer.update_prtabs mode xprods prtabs}, stamp ())
   299   end;
   300 
   301 
   302 (* remove_syntax *)
   303 
   304 fun remove_syntax (mode, inout) syn_ext (Syntax (tabs, _)) =
   305   let
   306     val SynExt.SynExt {xprods, consts = _, prmodes = _,
   307       parse_ast_translation, parse_rules, parse_translation, print_translation, print_rules,
   308       print_ast_translation, token_translation = _} = syn_ext;
   309     val {input, lexicon, gram, consts, prmodes,
   310       parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab, print_ruletab,
   311       print_ast_trtab, tokentrtab, prtabs} = tabs;
   312     val input' = if inout then subtract (op =) xprods input else input;
   313     val changed = length input <> length input';
   314     fun if_inout xs = if inout then xs else [];
   315   in
   316     Syntax
   317     ({input = input',
   318       lexicon = if changed then Scan.make_lexicon (SynExt.delims_of input') else lexicon,
   319       gram = if changed then Parser.make_gram input' else gram,
   320       consts = consts,
   321       prmodes = prmodes,
   322       parse_ast_trtab = remove_trtab (if_inout parse_ast_translation) parse_ast_trtab,
   323       parse_ruletab = remove_ruletab (if_inout parse_rules) parse_ruletab,
   324       parse_trtab = remove_trtab (if_inout parse_translation) parse_trtab,
   325       print_trtab = remove_tr'tab print_translation print_trtab,
   326       print_ruletab = remove_ruletab print_rules print_ruletab,
   327       print_ast_trtab = remove_tr'tab print_ast_translation print_ast_trtab,
   328       tokentrtab = tokentrtab,
   329       prtabs = Printer.remove_prtabs mode xprods prtabs}, stamp ())
   330   end;
   331 
   332 
   333 (* merge_syntaxes *)
   334 
   335 fun merge_syntaxes (Syntax (tabs1, _)) (Syntax (tabs2, _)) =
   336   let
   337     val {input = input1, lexicon = lexicon1, gram = gram1, consts = consts1,
   338       prmodes = prmodes1, parse_ast_trtab = parse_ast_trtab1,
   339       parse_ruletab = parse_ruletab1, parse_trtab = parse_trtab1,
   340       print_trtab = print_trtab1, print_ruletab = print_ruletab1,
   341       print_ast_trtab = print_ast_trtab1, tokentrtab = tokentrtab1, prtabs = prtabs1} = tabs1;
   342 
   343     val {input = input2, lexicon = lexicon2, gram = gram2, consts = consts2,
   344       prmodes = prmodes2, parse_ast_trtab = parse_ast_trtab2,
   345       parse_ruletab = parse_ruletab2, parse_trtab = parse_trtab2,
   346       print_trtab = print_trtab2, print_ruletab = print_ruletab2,
   347       print_ast_trtab = print_ast_trtab2, tokentrtab = tokentrtab2, prtabs = prtabs2} = tabs2;
   348   in
   349     Syntax
   350     ({input = Library.merge (op =) (input1, input2),
   351       lexicon = Scan.merge_lexicons lexicon1 lexicon2,
   352       gram = Parser.merge_grams gram1 gram2,
   353       consts = sort_distinct string_ord (consts1 @ consts2),
   354       prmodes = Library.merge (op =) (prmodes1, prmodes2),
   355       parse_ast_trtab =
   356         merge_trtabs "parse ast translation" parse_ast_trtab1 parse_ast_trtab2,
   357       parse_ruletab = merge_ruletabs parse_ruletab1 parse_ruletab2,
   358       parse_trtab = merge_trtabs "parse translation" parse_trtab1 parse_trtab2,
   359       print_trtab = merge_tr'tabs print_trtab1 print_trtab2,
   360       print_ruletab = merge_ruletabs print_ruletab1 print_ruletab2,
   361       print_ast_trtab = merge_tr'tabs print_ast_trtab1 print_ast_trtab2,
   362       tokentrtab = merge_tokentrtabs tokentrtab1 tokentrtab2,
   363       prtabs = Printer.merge_prtabs prtabs1 prtabs2}, stamp ())
   364   end;
   365 
   366 
   367 (* basic syntax *)
   368 
   369 val basic_syn =
   370   empty_syntax
   371   |> update_syntax mode_default TypeExt.type_ext
   372   |> update_syntax mode_default SynExt.pure_ext;
   373 
   374 val basic_nonterms =
   375   (Lexicon.terminals @ [SynExt.logic, "type", "types", "sort", "classes",
   376     SynExt.args, SynExt.cargs, "pttrn", "pttrns", "idt", "idts", "aprop",
   377     "asms", SynExt.any_, SynExt.sprop, "num_const", "index", "struct"]);
   378 
   379 
   380 
   381 (** print syntax **)
   382 
   383 local
   384 
   385 fun pretty_strs_qs name strs =
   386   Pretty.strs (name :: map Library.quote (sort_strings strs));
   387 
   388 fun pretty_gram (Syntax (tabs, _)) =
   389   let
   390     val {lexicon, prmodes, gram, prtabs, ...} = tabs;
   391     val prmodes' = sort_strings (filter_out (equal "") prmodes);
   392   in
   393     [pretty_strs_qs "lexicon:" (Scan.dest_lexicon lexicon),
   394       Pretty.big_list "prods:" (Parser.pretty_gram gram),
   395       pretty_strs_qs "print modes:" prmodes']
   396   end;
   397 
   398 fun pretty_trans (Syntax (tabs, _)) =
   399   let
   400     fun pretty_trtab name tab =
   401       pretty_strs_qs name (Symtab.keys tab);
   402 
   403     fun pretty_ruletab name tab =
   404       Pretty.big_list name (map Ast.pretty_rule (dest_ruletab tab));
   405 
   406     fun pretty_tokentr (mode, trs) = Pretty.strs (Library.quote mode ^ ":" :: map fst trs);
   407 
   408     val {consts, parse_ast_trtab, parse_ruletab, parse_trtab, print_trtab,
   409       print_ruletab, print_ast_trtab, tokentrtab, ...} = tabs;
   410   in
   411     [pretty_strs_qs "consts:" consts,
   412       pretty_trtab "parse_ast_translation:" parse_ast_trtab,
   413       pretty_ruletab "parse_rules:" parse_ruletab,
   414       pretty_trtab "parse_translation:" parse_trtab,
   415       pretty_trtab "print_translation:" print_trtab,
   416       pretty_ruletab "print_rules:" print_ruletab,
   417       pretty_trtab "print_ast_translation:" print_ast_trtab,
   418       Pretty.big_list "token_translation:" (map pretty_tokentr tokentrtab)]
   419   end;
   420 
   421 in
   422 
   423 fun print_gram syn = Pretty.writeln (Pretty.chunks (pretty_gram syn));
   424 fun print_trans syn = Pretty.writeln (Pretty.chunks (pretty_trans syn));
   425 fun print_syntax syn = Pretty.writeln (Pretty.chunks (pretty_gram syn @ pretty_trans syn));
   426 
   427 end;
   428 
   429 
   430 
   431 (** read **)
   432 
   433 (* read_ast *)
   434 
   435 val ambiguity_level = ref 1;
   436 val ambiguity_is_error = ref false;
   437 
   438 fun read_asts ctxt is_logtype (Syntax (tabs, _)) xids root str =
   439   let
   440     val {lexicon, gram, parse_ast_trtab, ...} = tabs;
   441     val root' = if root <> "prop" andalso is_logtype root then SynExt.logic else root;
   442     val chars = Symbol.explode str;
   443     val pts = Parser.parse gram root' (Lexicon.tokenize lexicon xids chars);
   444 
   445     fun show_pt pt =
   446       Pretty.string_of (Ast.pretty_ast (hd (SynTrans.pts_to_asts ctxt (K NONE) [pt])));
   447   in
   448     if length pts > ! ambiguity_level then
   449       if ! ambiguity_is_error then error ("Ambiguous input " ^ quote str)
   450       else (warning ("Ambiguous input " ^ quote str ^ "\n" ^
   451           "produces " ^ string_of_int (length pts) ^ " parse trees.");
   452          List.app (warning o show_pt) pts)
   453     else ();
   454     SynTrans.pts_to_asts ctxt (lookup_tr parse_ast_trtab) pts
   455   end;
   456 
   457 
   458 (* read *)
   459 
   460 fun read ctxt is_logtype (syn as Syntax (tabs, _)) ty str =
   461   let
   462     val {parse_ruletab, parse_trtab, ...} = tabs;
   463     val asts = read_asts ctxt is_logtype syn false (SynExt.typ_to_nonterm ty) str;
   464   in
   465     SynTrans.asts_to_terms ctxt (lookup_tr parse_trtab)
   466       (map (Ast.normalize_ast (Symtab.lookup_list parse_ruletab)) asts)
   467   end;
   468 
   469 
   470 (* read terms *)
   471 
   472 (*brute-force disambiguation via type-inference*)
   473 fun disambig _ _ [t] = t
   474   | disambig pp check ts =
   475       let
   476         val errs = map check ts;
   477         val results = map_filter (fn (t, NONE) => SOME t | _ => NONE) (ts ~~ errs);
   478 
   479         val ambiguity = length ts;
   480         fun ambig_msg () =
   481           if ambiguity > 1 andalso ambiguity <= ! ambiguity_level then
   482             "Got more than one parse tree.\n\
   483             \Retry with smaller ambiguity_level for more information."
   484           else "";
   485       in
   486         if null results then cat_error (ambig_msg ()) (cat_lines (map_filter I errs))
   487         else if length results = 1 then
   488           (if ambiguity > ! ambiguity_level then
   489             warning "Fortunately, only one parse tree is type correct.\n\
   490               \You may still want to disambiguate your grammar or your input."
   491           else (); hd results)
   492         else cat_error (ambig_msg ()) ("More than one term is type correct:\n" ^
   493           cat_lines (map (Pretty.string_of_term pp) ts))
   494       end;
   495 
   496 fun standard_parse_term pp check get_sort map_const map_free map_type map_sort
   497     ctxt is_logtype syn ty str =
   498   read ctxt is_logtype syn ty str
   499   |> map (TypeExt.decode_term get_sort map_const map_free map_type map_sort)
   500   |> disambig (Printer.pp_show_brackets pp) check;
   501 
   502 
   503 (* read types *)
   504 
   505 fun standard_parse_typ ctxt syn get_sort map_sort str =
   506   (case read ctxt (K false) syn SynExt.typeT str of
   507     [t] => TypeExt.typ_of_term (get_sort (TypeExt.term_sorts map_sort t)) map_sort t
   508   | _ => error "read_typ: ambiguous syntax");
   509 
   510 
   511 (* read sorts *)
   512 
   513 fun standard_parse_sort ctxt syn map_sort str =
   514   (case read ctxt (K false) syn TypeExt.sortT str of
   515     [t] => TypeExt.sort_of_term map_sort t
   516   | _ => error "read_sort: ambiguous syntax");
   517 
   518 
   519 
   520 (** prepare translation rules **)
   521 
   522 datatype 'a trrule =
   523   ParseRule of 'a * 'a |
   524   PrintRule of 'a * 'a |
   525   ParsePrintRule of 'a * 'a;
   526 
   527 fun map_trrule f (ParseRule (x, y)) = ParseRule (f x, f y)
   528   | map_trrule f (PrintRule (x, y)) = PrintRule (f x, f y)
   529   | map_trrule f (ParsePrintRule (x, y)) = ParsePrintRule (f x, f y);
   530 
   531 fun parse_rule (ParseRule pats) = SOME pats
   532   | parse_rule (PrintRule _) = NONE
   533   | parse_rule (ParsePrintRule pats) = SOME pats;
   534 
   535 fun print_rule (ParseRule _) = NONE
   536   | print_rule (PrintRule pats) = SOME (swap pats)
   537   | print_rule (ParsePrintRule pats) = SOME (swap pats);
   538 
   539 
   540 local
   541 
   542 fun check_rule (rule as (lhs, rhs)) =
   543   (case Ast.rule_error rule of
   544     SOME msg =>
   545       error ("Error in syntax translation rule: " ^ msg ^ "\n" ^
   546         Ast.str_of_ast lhs ^ "  ->  " ^ Ast.str_of_ast rhs)
   547   | NONE => rule);
   548 
   549 fun read_pattern ctxt is_logtype syn (root, str) =
   550   let
   551     val Syntax ({consts, ...}, _) = syn;
   552 
   553     fun constify (ast as Ast.Constant _) = ast
   554       | constify (ast as Ast.Variable x) =
   555           if member (op =) consts x orelse NameSpace.is_qualified x then Ast.Constant x
   556           else ast
   557       | constify (Ast.Appl asts) = Ast.Appl (map constify asts);
   558   in
   559     (case read_asts ctxt is_logtype syn true root str of
   560       [ast] => constify ast
   561     | _ => error ("Syntactically ambiguous input: " ^ quote str))
   562   end handle ERROR msg =>
   563     cat_error msg ("The error(s) above occurred in translation pattern " ^
   564       quote str);
   565 
   566 fun prep_rules rd_pat raw_rules =
   567   let val rules = map (map_trrule rd_pat) raw_rules in
   568     (map check_rule (map_filter parse_rule rules),
   569       map check_rule (map_filter print_rule rules))
   570   end
   571 
   572 in
   573 
   574 val cert_rules = prep_rules I;
   575 
   576 fun read_rules ctxt is_logtype syn =
   577   prep_rules (read_pattern ctxt is_logtype syn);
   578 
   579 end;
   580 
   581 
   582 
   583 (** unparse terms, typs, sorts **)
   584 
   585 local
   586 
   587 fun unparse_t t_to_ast prt_t markup ctxt (syn as Syntax (tabs, _)) curried t =
   588   let
   589     val {print_trtab, print_ruletab, print_ast_trtab, tokentrtab, prtabs, ...} = tabs;
   590     val ast = t_to_ast ctxt (lookup_tr' print_trtab) t;
   591   in
   592     Pretty.markup markup (prt_t ctxt curried prtabs (lookup_tr' print_ast_trtab)
   593       (lookup_tokentr tokentrtab (print_mode_value ()))
   594       (Ast.normalize_ast (Symtab.lookup_list print_ruletab) ast))
   595   end;
   596 
   597 in
   598 
   599 fun standard_unparse_term extern =
   600   unparse_t Printer.term_to_ast (Printer.pretty_term_ast extern) Markup.term;
   601 
   602 fun standard_unparse_typ ctxt syn =
   603   unparse_t Printer.typ_to_ast Printer.pretty_typ_ast Markup.typ ctxt syn false;
   604 
   605 fun standard_unparse_sort ctxt syn =
   606   unparse_t Printer.sort_to_ast Printer.pretty_typ_ast Markup.sort ctxt syn false;
   607 
   608 end;
   609 
   610 
   611 
   612 (** modify syntax **)
   613 
   614 fun ext_syntax f decls = update_syntax mode_default (f decls);
   615 
   616 val update_type_gram       = ext_syntax Mixfix.syn_ext_types;
   617 val update_consts          = ext_syntax SynExt.syn_ext_const_names;
   618 val update_trfuns          = ext_syntax SynExt.syn_ext_trfuns;
   619 val update_advanced_trfuns = ext_syntax SynExt.syn_ext_advanced_trfuns;
   620 val extend_tokentrfuns     = ext_syntax SynExt.syn_ext_tokentrfuns;
   621 
   622 fun update_const_gram is_logtype prmode decls =
   623   update_syntax prmode (Mixfix.syn_ext_consts is_logtype decls);
   624 
   625 fun remove_const_gram is_logtype prmode decls =
   626   remove_syntax prmode (Mixfix.syn_ext_consts is_logtype decls);
   627 
   628 fun update_trrules ctxt is_logtype syn =
   629   ext_syntax SynExt.syn_ext_rules o read_rules ctxt is_logtype syn;
   630 
   631 fun remove_trrules ctxt is_logtype syn =
   632   remove_syntax mode_default o SynExt.syn_ext_rules o read_rules ctxt is_logtype syn;
   633 
   634 val update_trrules_i = ext_syntax SynExt.syn_ext_rules o cert_rules;
   635 val remove_trrules_i = remove_syntax mode_default o SynExt.syn_ext_rules o cert_rules;
   636 
   637 
   638 
   639 (** inner syntax operations **)
   640 
   641 (* (un)parsing *)
   642 
   643 local
   644   type operations =
   645    {parse_sort: Proof.context -> string -> sort,
   646     parse_typ: Proof.context -> string -> typ,
   647     parse_term: Proof.context -> string -> term,
   648     parse_prop: Proof.context -> string -> term,
   649     unparse_sort: Proof.context -> sort -> Pretty.T,
   650     unparse_typ: Proof.context -> typ -> Pretty.T,
   651     unparse_term: Proof.context -> term -> Pretty.T};
   652 
   653   val operations = ref (NONE: operations option);
   654 
   655   fun operation which ctxt x =
   656     (case ! operations of
   657       NONE => error "Inner syntax operations not yet installed"
   658     | SOME ops => which ops ctxt x);
   659 in
   660 
   661 val parse_sort = operation #parse_sort;
   662 val parse_typ = operation #parse_typ;
   663 val parse_term = operation #parse_term;
   664 val parse_prop = operation #parse_prop;
   665 val unparse_sort = operation #unparse_sort;
   666 val unparse_typ = operation #unparse_typ;
   667 val unparse_term = operation #unparse_term;
   668 
   669 fun install_operations ops = CRITICAL (fn () =>
   670   if is_some (! operations) then error "Inner syntax operations already installed"
   671   else operations := SOME ops);
   672 
   673 end;
   674 
   675 
   676 (* context-sensitive (un)checking *)
   677 
   678 local
   679 
   680 type key = int * bool;
   681 type 'a check = 'a list -> Proof.context -> ('a list * Proof.context) option;
   682 
   683 structure Checks = GenericDataFun
   684 (
   685   type T =
   686     ((key * ((string * typ check) * stamp) list) list *
   687      (key * ((string * term check) * stamp) list) list);
   688   val empty = ([], []);
   689   val extend = I;
   690   fun merge _ ((typ_checks1, term_checks1), (typ_checks2, term_checks2)) : T =
   691     (AList.join (op =) (K (Library.merge (eq_snd (op =)))) (typ_checks1, typ_checks2),
   692      AList.join (op =) (K (Library.merge (eq_snd (op =)))) (term_checks1, term_checks2));
   693 );
   694 
   695 fun gen_add which (key: key) name f =
   696   Checks.map (which (AList.map_default op = (key, []) (cons ((name, f), stamp ()))));
   697 
   698 fun check_stage fs = perhaps_loop (perhaps_apply (map uncurry fs));
   699 
   700 fun gen_check which uncheck ctxt0 xs0 =
   701   let
   702     val funs = which (Checks.get (Context.Proof ctxt0))
   703       |> map_filter (fn ((i, u), fs) => if uncheck = u then SOME (i, map (snd o fst) fs) else NONE)
   704       |> Library.sort (int_ord o pairself fst) |> map snd
   705       |> not uncheck ? map rev;
   706     val check_all = perhaps_apply (map check_stage funs);
   707   in #1 (perhaps check_all (xs0, ctxt0)) end;
   708 
   709 fun map_sort f S =
   710   (case f (TFree ("", S)) of
   711     TFree ("", S') => S'
   712   | _ => raise TYPE ("map_sort", [TFree ("", S)], []));
   713 
   714 in
   715 
   716 fun print_checks ctxt =
   717   let
   718     fun split_checks checks =
   719       List.partition (fn ((_, un), _) => not un) checks
   720       |> pairself (map (fn ((i, _), fs) => (i, map (fst o fst) fs))
   721           #> sort (int_ord o pairself fst));
   722     fun pretty_checks kind checks =
   723       checks |> map (fn (i, names) => Pretty.block
   724         [Pretty.str (kind ^ " (stage " ^ signed_string_of_int i ^ "):"),
   725           Pretty.brk 1, Pretty.strs names]);
   726 
   727     val (typs, terms) = Checks.get (Context.Proof ctxt);
   728     val (typ_checks, typ_unchecks) = split_checks typs;
   729     val (term_checks, term_unchecks) = split_checks terms;
   730   in
   731     pretty_checks "typ_checks" typ_checks @
   732     pretty_checks "term_checks" term_checks @
   733     pretty_checks "typ_unchecks" typ_unchecks @
   734     pretty_checks "term_unchecks" term_unchecks
   735   end |> Pretty.chunks |> Pretty.writeln;
   736 
   737 fun add_typ_check stage = gen_add apfst (stage, false);
   738 fun add_term_check stage = gen_add apsnd (stage, false);
   739 fun add_typ_uncheck stage = gen_add apfst (stage, true);
   740 fun add_term_uncheck stage = gen_add apsnd (stage, true);
   741 
   742 val check_typs = gen_check fst false;
   743 val check_terms = gen_check snd false;
   744 fun check_props ctxt = map (TypeInfer.constrain propT) #> check_terms ctxt;
   745 
   746 val check_typ = singleton o check_typs;
   747 val check_term = singleton o check_terms;
   748 val check_prop = singleton o check_props;
   749 val check_sort = map_sort o check_typ;
   750 
   751 val uncheck_typs = gen_check fst true;
   752 val uncheck_terms = gen_check snd true;
   753 val uncheck_sort = map_sort o singleton o uncheck_typs;
   754 
   755 end;
   756 
   757 
   758 (* derived operations for classrel and arity *)
   759 
   760 val uncheck_classrel = map o singleton o uncheck_sort;
   761 
   762 fun unparse_classrel ctxt cs = Pretty.block (flat
   763   (separate [Pretty.str " <", Pretty.brk 1] (map (single o unparse_sort ctxt o single) cs)));
   764 
   765 fun uncheck_arity ctxt (a, Ss, S) =
   766   let
   767     val T = Type (a, replicate (length Ss) dummyT);
   768     val a' =
   769       (case singleton (uncheck_typs ctxt) T of
   770         Type (a', _) => a'
   771       | T => raise TYPE ("uncheck_arity", [T], []));
   772     val Ss' = map (uncheck_sort ctxt) Ss;
   773     val S' = uncheck_sort ctxt S;
   774   in (a', Ss', S') end;
   775 
   776 fun unparse_arity ctxt (a, Ss, S) =
   777   let
   778     val prtT = unparse_typ ctxt (Type (a, []));
   779     val dom =
   780       if null Ss then []
   781       else [Pretty.list "(" ")" (map (unparse_sort ctxt) Ss), Pretty.brk 1];
   782   in Pretty.block ([prtT, Pretty.str " ::", Pretty.brk 1] @ dom @ [unparse_sort ctxt S]) end;
   783 
   784 
   785 (* read = parse + check *)
   786 
   787 fun read_sort ctxt = parse_sort ctxt #> check_sort ctxt;
   788 fun read_typ ctxt = parse_typ ctxt #> singleton (check_typs ctxt);
   789 
   790 fun read_terms ctxt = map (parse_term ctxt) #> check_terms ctxt;
   791 fun read_props ctxt = map (parse_prop ctxt) #> check_props ctxt;
   792 
   793 val read_term = singleton o read_terms;
   794 val read_prop = singleton o read_props;
   795 
   796 val read_sort_global = read_sort o ProofContext.init;
   797 val read_typ_global = read_typ o ProofContext.init;
   798 val read_term_global = read_term o ProofContext.init;
   799 val read_prop_global = read_prop o ProofContext.init;
   800 
   801 
   802 (* pretty = uncheck + unparse *)
   803 
   804 fun pretty_sort ctxt = uncheck_sort ctxt #> unparse_sort ctxt;
   805 fun pretty_classrel ctxt = uncheck_classrel ctxt #> unparse_classrel ctxt;
   806 fun pretty_arity ctxt = uncheck_arity ctxt #> unparse_arity ctxt;
   807 fun pretty_typ ctxt = singleton (uncheck_typs ctxt) #> unparse_typ ctxt;
   808 fun pretty_term ctxt = singleton (uncheck_terms ctxt) #> unparse_term ctxt;
   809 
   810 val string_of_sort = Pretty.string_of oo pretty_sort;
   811 val string_of_classrel = Pretty.string_of oo pretty_classrel;
   812 val string_of_arity = Pretty.string_of oo pretty_arity;
   813 val string_of_typ = Pretty.string_of oo pretty_typ;
   814 val string_of_term = Pretty.string_of oo pretty_term;
   815 
   816 (*pp operations -- deferred evaluation*)
   817 fun pp ctxt = Pretty.pp
   818  (fn x => pretty_term ctxt x,
   819   fn x => pretty_typ ctxt x,
   820   fn x => pretty_sort ctxt x,
   821   fn x => pretty_classrel ctxt x,
   822   fn x => pretty_arity ctxt x);
   823 
   824 (*export parts of internal Syntax structures*)
   825 open Lexicon SynExt Ast Parser TypeExt SynTrans Mixfix Printer;
   826 
   827 end;
   828 
   829 structure BasicSyntax: BASIC_SYNTAX = Syntax;
   830 open BasicSyntax;
   831 
   832 structure Hidden = struct end;
   833 structure Ast = Hidden;
   834 structure SynExt = Hidden;
   835 structure Parser = Hidden;
   836 structure TypeExt = Hidden;
   837 structure SynTrans = Hidden;
   838 structure Mixfix = Hidden;
   839 structure Printer = Hidden;