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