src/Pure/Syntax/syntax_trans.ML
changeset 42284 326f57825e1a
parent 42278 088a2d69746f
child 42288 2074b31650e6
equal deleted inserted replaced
42283:25d9d836ed9c 42284:326f57825e1a
       
     1 (*  Title:      Pure/Syntax/syntax_trans.ML
       
     2     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
       
     3 
       
     4 Syntax translation functions.
       
     5 *)
       
     6 
       
     7 signature BASIC_SYNTAX_TRANS =
       
     8 sig
       
     9   val eta_contract: bool Config.T
       
    10 end
       
    11 
       
    12 signature SYNTAX_TRANS =
       
    13 sig
       
    14   include BASIC_SYNTAX_TRANS
       
    15   val no_brackets: unit -> bool
       
    16   val no_type_brackets: unit -> bool
       
    17   val abs_tr: term list -> term
       
    18   val mk_binder_tr: string * string -> string * (term list -> term)
       
    19   val antiquote_tr: string -> term -> term
       
    20   val quote_tr: string -> term -> term
       
    21   val quote_antiquote_tr: string -> string -> string -> string * (term list -> term)
       
    22   val non_typed_tr': (term list -> term) -> typ -> term list -> term
       
    23   val non_typed_tr'': ('a -> term list -> term) -> 'a -> typ -> term list -> term
       
    24   val tappl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
       
    25   val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
       
    26   val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
       
    27   val eta_contract_default: bool Unsynchronized.ref
       
    28   val eta_contract_raw: Config.raw
       
    29   val mark_bound: string -> term
       
    30   val mark_boundT: string * typ -> term
       
    31   val bound_vars: (string * typ) list -> term -> term
       
    32   val abs_tr': Proof.context -> term -> term
       
    33   val atomic_abs_tr': string * typ * term -> term * term
       
    34   val const_abs_tr': term -> term
       
    35   val mk_binder_tr': string * string -> string * (term list -> term)
       
    36   val preserve_binder_abs_tr': string -> string -> string * (term list -> term)
       
    37   val preserve_binder_abs2_tr': string -> string -> string * (term list -> term)
       
    38   val prop_tr': term -> term
       
    39   val variant_abs: string * typ * term -> string * term
       
    40   val variant_abs': string * typ * term -> string * term
       
    41   val dependent_tr': string * string -> term list -> term
       
    42   val antiquote_tr': string -> term -> term
       
    43   val quote_tr': string -> term -> term
       
    44   val quote_antiquote_tr': string -> string -> string -> string * (term list -> term)
       
    45   val update_name_tr': term -> term
       
    46   val pure_trfuns:
       
    47     (string * (Ast.ast list -> Ast.ast)) list *
       
    48     (string * (term list -> term)) list *
       
    49     (string * (term list -> term)) list *
       
    50     (string * (Ast.ast list -> Ast.ast)) list
       
    51   val struct_trfuns: string list ->
       
    52     (string * (Ast.ast list -> Ast.ast)) list *
       
    53     (string * (term list -> term)) list *
       
    54     (string * (typ -> term list -> term)) list *
       
    55     (string * (Ast.ast list -> Ast.ast)) list
       
    56 end;
       
    57 
       
    58 structure Syntax_Trans: SYNTAX_TRANS =
       
    59 struct
       
    60 
       
    61 (* print mode *)
       
    62 
       
    63 val bracketsN = "brackets";
       
    64 val no_bracketsN = "no_brackets";
       
    65 
       
    66 fun no_brackets () =
       
    67   find_first (fn mode => mode = bracketsN orelse mode = no_bracketsN)
       
    68     (print_mode_value ()) = SOME no_bracketsN;
       
    69 
       
    70 val type_bracketsN = "type_brackets";
       
    71 val no_type_bracketsN = "no_type_brackets";
       
    72 
       
    73 fun no_type_brackets () =
       
    74   find_first (fn mode => mode = type_bracketsN orelse mode = no_type_bracketsN)
       
    75     (print_mode_value ()) <> SOME type_bracketsN;
       
    76 
       
    77 
       
    78 
       
    79 (** parse (ast) translations **)
       
    80 
       
    81 (* strip_positions *)
       
    82 
       
    83 fun strip_positions_ast_tr [ast] = Ast.strip_positions ast
       
    84   | strip_positions_ast_tr asts = raise Ast.AST ("strip_positions_ast_tr", asts);
       
    85 
       
    86 
       
    87 (* constify *)
       
    88 
       
    89 fun constify_ast_tr [Ast.Variable c] = Ast.Constant c
       
    90   | constify_ast_tr asts = raise Ast.AST ("constify_ast_tr", asts);
       
    91 
       
    92 
       
    93 (* type syntax *)
       
    94 
       
    95 fun tapp_ast_tr [ty, c] = Ast.Appl [c, ty]
       
    96   | tapp_ast_tr asts = raise Ast.AST ("tapp_ast_tr", asts);
       
    97 
       
    98 fun tappl_ast_tr [ty, tys, c] = Ast.mk_appl c (ty :: Ast.unfold_ast "_types" tys)
       
    99   | tappl_ast_tr asts = raise Ast.AST ("tappl_ast_tr", asts);
       
   100 
       
   101 fun bracket_ast_tr [dom, cod] = Ast.fold_ast_p "\\<^type>fun" (Ast.unfold_ast "_types" dom, cod)
       
   102   | bracket_ast_tr asts = raise Ast.AST ("bracket_ast_tr", asts);
       
   103 
       
   104 
       
   105 (* application *)
       
   106 
       
   107 fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args)
       
   108   | appl_ast_tr asts = raise Ast.AST ("appl_ast_tr", asts);
       
   109 
       
   110 fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args)
       
   111   | applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts);
       
   112 
       
   113 
       
   114 (* abstraction *)
       
   115 
       
   116 fun idtyp_ast_tr [x, ty] = Ast.Appl [Ast.Constant "_constrain", x, ty]
       
   117   | idtyp_ast_tr asts = raise Ast.AST ("idtyp_ast_tr", asts);
       
   118 
       
   119 fun idtypdummy_ast_tr [ty] = Ast.Appl [Ast.Constant "_constrain", Ast.Constant "_idtdummy", ty]
       
   120   | idtypdummy_ast_tr asts = raise Ast.AST ("idtyp_ast_tr", asts);
       
   121 
       
   122 fun lambda_ast_tr [pats, body] = Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body)
       
   123   | lambda_ast_tr asts = raise Ast.AST ("lambda_ast_tr", asts);
       
   124 
       
   125 fun absfree_proper (x, T, t) =
       
   126   if can Name.dest_internal x
       
   127   then error ("Illegal internal variable in abstraction: " ^ quote x)
       
   128   else Term.absfree (x, T, t);
       
   129 
       
   130 fun abs_tr [Free (x, T), t] = absfree_proper (x, T, t)
       
   131   | abs_tr [Const ("_idtdummy", T), t] = Term.absdummy (T, t)
       
   132   | abs_tr [Const ("_constrain", _) $ x $ tT, t] =
       
   133       Lexicon.const "_constrainAbs" $ abs_tr [x, t] $ tT
       
   134   | abs_tr ts = raise TERM ("abs_tr", ts);
       
   135 
       
   136 
       
   137 (* binder *)
       
   138 
       
   139 fun mk_binder_tr (syn, name) =
       
   140   let
       
   141     fun err ts = raise TERM ("binder_tr: " ^ syn, ts)
       
   142     fun binder_tr [Const ("_idts", _) $ idt $ idts, t] = binder_tr [idt, binder_tr [idts, t]]
       
   143       | binder_tr [x, t] =
       
   144           let val abs = abs_tr [x, t] handle TERM _ => err [x, t]
       
   145           in Lexicon.const name $ abs end
       
   146       | binder_tr ts = err ts;
       
   147   in (syn, binder_tr) end;
       
   148 
       
   149 
       
   150 (* type propositions *)
       
   151 
       
   152 fun mk_type ty =
       
   153   Lexicon.const "_constrain" $
       
   154     Lexicon.const "\\<^const>TYPE" $ (Lexicon.const "\\<^type>itself" $ ty);
       
   155 
       
   156 fun ofclass_tr [ty, cls] = cls $ mk_type ty
       
   157   | ofclass_tr ts = raise TERM ("ofclass_tr", ts);
       
   158 
       
   159 fun sort_constraint_tr [ty] = Lexicon.const "\\<^const>Pure.sort_constraint" $ mk_type ty
       
   160   | sort_constraint_tr ts = raise TERM ("sort_constraint_tr", ts);
       
   161 
       
   162 
       
   163 (* meta propositions *)
       
   164 
       
   165 fun aprop_tr [t] = Lexicon.const "_constrain" $ t $ Lexicon.const "\\<^type>prop"
       
   166   | aprop_tr ts = raise TERM ("aprop_tr", ts);
       
   167 
       
   168 
       
   169 (* meta implication *)
       
   170 
       
   171 fun bigimpl_ast_tr (asts as [asms, concl]) =
       
   172       let val prems =
       
   173         (case Ast.unfold_ast_p "_asms" asms of
       
   174           (asms', Ast.Appl [Ast.Constant "_asm", asm']) => asms' @ [asm']
       
   175         | _ => raise Ast.AST ("bigimpl_ast_tr", asts))
       
   176       in Ast.fold_ast_p "\\<^const>==>" (prems, concl) end
       
   177   | bigimpl_ast_tr asts = raise Ast.AST ("bigimpl_ast_tr", asts);
       
   178 
       
   179 
       
   180 (* type/term reflection *)
       
   181 
       
   182 fun type_tr [ty] = mk_type ty
       
   183   | type_tr ts = raise TERM ("type_tr", ts);
       
   184 
       
   185 
       
   186 (* dddot *)
       
   187 
       
   188 fun dddot_tr ts = Term.list_comb (Lexicon.var Syn_Ext.dddot_indexname, ts);
       
   189 
       
   190 
       
   191 (* quote / antiquote *)
       
   192 
       
   193 fun antiquote_tr name =
       
   194   let
       
   195     fun tr i ((t as Const (c, _)) $ u) =
       
   196           if c = name then tr i u $ Bound i
       
   197           else tr i t $ tr i u
       
   198       | tr i (t $ u) = tr i t $ tr i u
       
   199       | tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t)
       
   200       | tr _ a = a;
       
   201   in tr 0 end;
       
   202 
       
   203 fun quote_tr name t = Abs ("s", dummyT, antiquote_tr name (Term.incr_boundvars 1 t));
       
   204 
       
   205 fun quote_antiquote_tr quoteN antiquoteN name =
       
   206   let
       
   207     fun tr [t] = Lexicon.const name $ quote_tr antiquoteN t
       
   208       | tr ts = raise TERM ("quote_tr", ts);
       
   209   in (quoteN, tr) end;
       
   210 
       
   211 
       
   212 (* corresponding updates *)
       
   213 
       
   214 fun update_name_tr (Free (x, T) :: ts) = list_comb (Free (suffix "_update" x, T), ts)
       
   215   | update_name_tr (Const (x, T) :: ts) = list_comb (Const (suffix "_update" x, T), ts)
       
   216   | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
       
   217       if Term_Position.is_position ty then list_comb (c $ update_name_tr [t] $ ty, ts)
       
   218       else
       
   219         list_comb (c $ update_name_tr [t] $
       
   220           (Lexicon.fun_type $
       
   221             (Lexicon.fun_type $ Lexicon.dummy_type $ ty) $ Lexicon.dummy_type), ts)
       
   222   | update_name_tr ts = raise TERM ("update_name_tr", ts);
       
   223 
       
   224 
       
   225 (* indexed syntax *)
       
   226 
       
   227 fun struct_ast_tr [Ast.Appl [Ast.Constant "_index", ast]] = ast
       
   228   | struct_ast_tr asts = Ast.mk_appl (Ast.Constant "_struct") asts;
       
   229 
       
   230 fun index_ast_tr ast =
       
   231   Ast.mk_appl (Ast.Constant "_index") [Ast.mk_appl (Ast.Constant "_struct") [ast]];
       
   232 
       
   233 fun indexdefault_ast_tr [] = index_ast_tr (Ast.Constant "_indexdefault")
       
   234   | indexdefault_ast_tr asts = raise Ast.AST ("indexdefault_ast_tr", asts);
       
   235 
       
   236 fun indexnum_ast_tr [ast] = index_ast_tr (Ast.mk_appl (Ast.Constant "_indexnum") [ast])
       
   237   | indexnum_ast_tr asts = raise Ast.AST ("indexnum_ast_tr", asts);
       
   238 
       
   239 fun indexvar_ast_tr [] = Ast.mk_appl (Ast.Constant "_index") [Ast.Variable "some_index"]
       
   240   | indexvar_ast_tr asts = raise Ast.AST ("indexvar_ast_tr", asts);
       
   241 
       
   242 fun index_tr [t] = t
       
   243   | index_tr ts = raise TERM ("index_tr", ts);
       
   244 
       
   245 
       
   246 (* implicit structures *)
       
   247 
       
   248 fun the_struct structs i =
       
   249   if 1 <= i andalso i <= length structs then nth structs (i - 1)
       
   250   else error ("Illegal reference to implicit structure #" ^ string_of_int i);
       
   251 
       
   252 fun struct_tr structs [Const ("_indexdefault", _)] =
       
   253       Lexicon.free (the_struct structs 1)
       
   254   | struct_tr structs [t as (Const ("_indexnum", _) $ Const (s, _))] =
       
   255       Lexicon.free (the_struct structs
       
   256         (case Lexicon.read_nat s of SOME n => n | NONE => raise TERM ("struct_tr", [t])))
       
   257   | struct_tr _ ts = raise TERM ("struct_tr", ts);
       
   258 
       
   259 
       
   260 
       
   261 (** print (ast) translations **)
       
   262 
       
   263 (* types *)
       
   264 
       
   265 fun non_typed_tr' f _ ts = f ts;
       
   266 fun non_typed_tr'' f x _ ts = f x ts;
       
   267 
       
   268 
       
   269 (* type syntax *)
       
   270 
       
   271 fun tappl_ast_tr' (f, []) = raise Ast.AST ("tappl_ast_tr'", [f])
       
   272   | tappl_ast_tr' (f, [ty]) = Ast.Appl [Ast.Constant "_tapp", ty, f]
       
   273   | tappl_ast_tr' (f, ty :: tys) =
       
   274       Ast.Appl [Ast.Constant "_tappl", ty, Ast.fold_ast "_types" tys, f];
       
   275 
       
   276 fun fun_ast_tr' asts =
       
   277   if no_brackets () orelse no_type_brackets () then raise Match
       
   278   else
       
   279     (case Ast.unfold_ast_p "\\<^type>fun" (Ast.Appl (Ast.Constant "\\<^type>fun" :: asts)) of
       
   280       (dom as _ :: _ :: _, cod)
       
   281         => Ast.Appl [Ast.Constant "_bracket", Ast.fold_ast "_types" dom, cod]
       
   282     | _ => raise Match);
       
   283 
       
   284 
       
   285 (* application *)
       
   286 
       
   287 fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f])
       
   288   | appl_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_appl", f, Ast.fold_ast "_args" args];
       
   289 
       
   290 fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f])
       
   291   | applC_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_applC", f, Ast.fold_ast "_cargs" args];
       
   292 
       
   293 
       
   294 (* partial eta-contraction before printing *)
       
   295 
       
   296 fun eta_abs (Abs (a, T, t)) =
       
   297       (case eta_abs t of
       
   298         t' as Const ("_aprop", _) $ _ => Abs (a, T, t')
       
   299       | t' as f $ u =>
       
   300           (case eta_abs u of
       
   301             Bound 0 =>
       
   302               if Term.is_dependent f then Abs (a, T, t')
       
   303               else incr_boundvars ~1 f
       
   304           | _ => Abs (a, T, t'))
       
   305       | t' => Abs (a, T, t'))
       
   306   | eta_abs t = t;
       
   307 
       
   308 val eta_contract_default = Unsynchronized.ref true;
       
   309 val eta_contract_raw = Config.declare "eta_contract" (fn _ => Config.Bool (! eta_contract_default));
       
   310 val eta_contract = Config.bool eta_contract_raw;
       
   311 
       
   312 fun eta_contr ctxt tm =
       
   313   if Config.get ctxt eta_contract then eta_abs tm else tm;
       
   314 
       
   315 
       
   316 (* abstraction *)
       
   317 
       
   318 fun mark_boundT (x, T) = Const ("_bound", T --> T) $ Free (x, T);
       
   319 fun mark_bound x = mark_boundT (x, dummyT);
       
   320 
       
   321 fun bound_vars vars body =
       
   322   subst_bounds (map mark_boundT (Term.rename_wrt_term body vars), body);
       
   323 
       
   324 fun strip_abss vars_of body_of tm =
       
   325   let
       
   326     val vars = vars_of tm;
       
   327     val body = body_of tm;
       
   328     val rev_new_vars = Term.rename_wrt_term body vars;
       
   329     fun subst (x, T) b =
       
   330       if can Name.dest_internal x andalso not (Term.is_dependent b)
       
   331       then (Const ("_idtdummy", T), incr_boundvars ~1 b)
       
   332       else (mark_boundT (x, T), Term.subst_bound (mark_bound x, b));
       
   333     val (rev_vars', body') = fold_map subst rev_new_vars body;
       
   334   in (rev rev_vars', body') end;
       
   335 
       
   336 
       
   337 fun abs_tr' ctxt tm =
       
   338   uncurry (fold_rev (fn x => fn t => Lexicon.const "_abs" $ x $ t))
       
   339     (strip_abss strip_abs_vars strip_abs_body (eta_contr ctxt tm));
       
   340 
       
   341 fun atomic_abs_tr' (x, T, t) =
       
   342   let val [xT] = Term.rename_wrt_term t [(x, T)]
       
   343   in (mark_boundT xT, subst_bound (mark_bound (fst xT), t)) end;
       
   344 
       
   345 fun abs_ast_tr' asts =
       
   346   (case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of
       
   347     ([], _) => raise Ast.AST ("abs_ast_tr'", asts)
       
   348   | (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]);
       
   349 
       
   350 fun const_abs_tr' t =
       
   351   (case eta_abs t of
       
   352     Abs (_, _, t') =>
       
   353       if Term.is_dependent t' then raise Match
       
   354       else incr_boundvars ~1 t'
       
   355   | _ => raise Match);
       
   356 
       
   357 
       
   358 (* binders *)
       
   359 
       
   360 fun mk_binder_tr' (name, syn) =
       
   361   let
       
   362     fun mk_idts [] = raise Match    (*abort translation*)
       
   363       | mk_idts [idt] = idt
       
   364       | mk_idts (idt :: idts) = Lexicon.const "_idts" $ idt $ mk_idts idts;
       
   365 
       
   366     fun tr' t =
       
   367       let
       
   368         val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t;
       
   369       in Lexicon.const syn $ mk_idts xs $ bd end;
       
   370 
       
   371     fun binder_tr' (t :: ts) = Term.list_comb (tr' (Lexicon.const name $ t), ts)
       
   372       | binder_tr' [] = raise Match;
       
   373   in (name, binder_tr') end;
       
   374 
       
   375 fun preserve_binder_abs_tr' name syn = (name, fn Abs abs :: ts =>
       
   376   let val (x, t) = atomic_abs_tr' abs
       
   377   in list_comb (Lexicon.const syn $ x $ t, ts) end);
       
   378 
       
   379 fun preserve_binder_abs2_tr' name syn = (name, fn A :: Abs abs :: ts =>
       
   380   let val (x, t) = atomic_abs_tr' abs
       
   381   in list_comb (Lexicon.const syn $ x $ A $ t, ts) end);
       
   382 
       
   383 
       
   384 (* idtyp constraints *)
       
   385 
       
   386 fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant "_constrain", x, ty], xs] =
       
   387       Ast.Appl [Ast.Constant a, Ast.Appl [Ast.Constant "_idtyp", x, ty], xs]
       
   388   | idtyp_ast_tr' _ _ = raise Match;
       
   389 
       
   390 
       
   391 (* meta propositions *)
       
   392 
       
   393 fun prop_tr' tm =
       
   394   let
       
   395     fun aprop t = Lexicon.const "_aprop" $ t;
       
   396 
       
   397     fun is_prop Ts t =
       
   398       fastype_of1 (Ts, t) = propT handle TERM _ => false;
       
   399 
       
   400     fun is_term (Const ("Pure.term", _) $ _) = true
       
   401       | is_term _ = false;
       
   402 
       
   403     fun tr' _ (t as Const _) = t
       
   404       | tr' Ts (t as Const ("_bound", _) $ u) =
       
   405           if is_prop Ts u then aprop t else t
       
   406       | tr' _ (t as Free (x, T)) =
       
   407           if T = propT then aprop (Lexicon.free x) else t
       
   408       | tr' _ (t as Var (xi, T)) =
       
   409           if T = propT then aprop (Lexicon.var xi) else t
       
   410       | tr' Ts (t as Bound _) =
       
   411           if is_prop Ts t then aprop t else t
       
   412       | tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t)
       
   413       | tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
       
   414           if is_prop Ts t andalso not (is_term t) then Const ("_type_prop", T) $ tr' Ts t1
       
   415           else tr' Ts t1 $ tr' Ts t2
       
   416       | tr' Ts (t as t1 $ t2) =
       
   417           (if is_Const (Term.head_of t) orelse not (is_prop Ts t)
       
   418             then I else aprop) (tr' Ts t1 $ tr' Ts t2);
       
   419   in tr' [] tm end;
       
   420 
       
   421 
       
   422 (* meta implication *)
       
   423 
       
   424 fun impl_ast_tr' asts =
       
   425   if no_brackets () then raise Match
       
   426   else
       
   427     (case Ast.unfold_ast_p "\\<^const>==>" (Ast.Appl (Ast.Constant "\\<^const>==>" :: asts)) of
       
   428       (prems as _ :: _ :: _, concl) =>
       
   429         let
       
   430           val (asms, asm) = split_last prems;
       
   431           val asms' = Ast.fold_ast_p "_asms" (asms, Ast.Appl [Ast.Constant "_asm", asm]);
       
   432         in Ast.Appl [Ast.Constant "_bigimpl", asms', concl] end
       
   433     | _ => raise Match);
       
   434 
       
   435 
       
   436 (* dependent / nondependent quantifiers *)
       
   437 
       
   438 fun var_abs mark (x, T, b) =
       
   439   let val ([x'], _) = Name.variants [x] (Term.declare_term_names b Name.context)
       
   440   in (x', subst_bound (mark (x', T), b)) end;
       
   441 
       
   442 val variant_abs = var_abs Free;
       
   443 val variant_abs' = var_abs mark_boundT;
       
   444 
       
   445 fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) =
       
   446       if Term.is_dependent B then
       
   447         let val (x', B') = variant_abs' (x, dummyT, B);
       
   448         in Term.list_comb (Lexicon.const q $ mark_boundT (x', T) $ A $ B', ts) end
       
   449       else Term.list_comb (Lexicon.const r $ A $ incr_boundvars ~1 B, ts)
       
   450   | dependent_tr' _ _ = raise Match;
       
   451 
       
   452 
       
   453 (* quote / antiquote *)
       
   454 
       
   455 fun antiquote_tr' name =
       
   456   let
       
   457     fun tr' i (t $ u) =
       
   458           if u aconv Bound i then Lexicon.const name $ tr' i t
       
   459           else tr' i t $ tr' i u
       
   460       | tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t)
       
   461       | tr' i a = if a aconv Bound i then raise Match else a;
       
   462   in tr' 0 end;
       
   463 
       
   464 fun quote_tr' name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' name t)
       
   465   | quote_tr' _ _ = raise Match;
       
   466 
       
   467 fun quote_antiquote_tr' quoteN antiquoteN name =
       
   468   let
       
   469     fun tr' (t :: ts) = Term.list_comb (Lexicon.const quoteN $ quote_tr' antiquoteN t, ts)
       
   470       | tr' _ = raise Match;
       
   471   in (name, tr') end;
       
   472 
       
   473 
       
   474 (* corresponding updates *)
       
   475 
       
   476 local
       
   477 
       
   478 fun upd_type (Type ("fun", [Type ("fun", [_, T]), _])) = T
       
   479   | upd_type _ = dummyT;
       
   480 
       
   481 fun upd_tr' (x_upd, T) =
       
   482   (case try (unsuffix "_update") x_upd of
       
   483     SOME x => (x, upd_type T)
       
   484   | NONE => raise Match);
       
   485 
       
   486 in
       
   487 
       
   488 fun update_name_tr' (Free x) = Free (upd_tr' x)
       
   489   | update_name_tr' ((c as Const ("_free", _)) $ Free x) = c $ Free (upd_tr' x)
       
   490   | update_name_tr' (Const x) = Const (upd_tr' x)
       
   491   | update_name_tr' _ = raise Match;
       
   492 
       
   493 end;
       
   494 
       
   495 
       
   496 (* indexed syntax *)
       
   497 
       
   498 fun index_ast_tr' [Ast.Appl [Ast.Constant "_struct", ast]] = ast
       
   499   | index_ast_tr' _ = raise Match;
       
   500 
       
   501 
       
   502 (* implicit structures *)
       
   503 
       
   504 fun the_struct' structs s =
       
   505   [(case Lexicon.read_nat s of
       
   506     SOME i => Ast.Variable (the_struct structs i handle ERROR _ => raise Match)
       
   507   | NONE => raise Match)] |> Ast.mk_appl (Ast.Constant "_free");
       
   508 
       
   509 fun struct_ast_tr' structs [Ast.Constant "_indexdefault"] = the_struct' structs "1"
       
   510   | struct_ast_tr' structs [Ast.Appl [Ast.Constant "_indexnum", Ast.Constant s]] =
       
   511       the_struct' structs s
       
   512   | struct_ast_tr' _ _ = raise Match;
       
   513 
       
   514 
       
   515 
       
   516 (** Pure translations **)
       
   517 
       
   518 val pure_trfuns =
       
   519   ([("_strip_positions", strip_positions_ast_tr),
       
   520     ("_constify", constify_ast_tr),
       
   521     ("_tapp", tapp_ast_tr),
       
   522     ("_tappl", tappl_ast_tr),
       
   523     ("_bracket", bracket_ast_tr),
       
   524     ("_appl", appl_ast_tr),
       
   525     ("_applC", applC_ast_tr),
       
   526     ("_lambda", lambda_ast_tr),
       
   527     ("_idtyp", idtyp_ast_tr),
       
   528     ("_idtypdummy", idtypdummy_ast_tr),
       
   529     ("_bigimpl", bigimpl_ast_tr),
       
   530     ("_indexdefault", indexdefault_ast_tr),
       
   531     ("_indexnum", indexnum_ast_tr),
       
   532     ("_indexvar", indexvar_ast_tr),
       
   533     ("_struct", struct_ast_tr)],
       
   534    [("_abs", abs_tr),
       
   535     ("_aprop", aprop_tr),
       
   536     ("_ofclass", ofclass_tr),
       
   537     ("_sort_constraint", sort_constraint_tr),
       
   538     ("_TYPE", type_tr),
       
   539     ("_DDDOT", dddot_tr),
       
   540     ("_update_name", update_name_tr),
       
   541     ("_index", index_tr)],
       
   542    ([]: (string * (term list -> term)) list),
       
   543    [("\\<^type>fun", fun_ast_tr'),
       
   544     ("_abs", abs_ast_tr'),
       
   545     ("_idts", idtyp_ast_tr' "_idts"),
       
   546     ("_pttrns", idtyp_ast_tr' "_pttrns"),
       
   547     ("\\<^const>==>", impl_ast_tr'),
       
   548     ("_index", index_ast_tr')]);
       
   549 
       
   550 fun struct_trfuns structs =
       
   551   ([], [("_struct", struct_tr structs)], [], [("_struct", struct_ast_tr' structs)]);
       
   552 
       
   553 end;
       
   554 
       
   555 structure Basic_Syntax_Trans: BASIC_SYNTAX_TRANS = Syntax_Trans;
       
   556 open Basic_Syntax_Trans;