| 548 |      1 | (*  Title:      Pure/Syntax/syn_trans.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | Syntax translation functions.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | signature SYN_TRANS0 =
 | 
|  |      9 | sig
 | 
|  |     10 |   val eta_contract: bool ref
 | 
|  |     11 |   val mk_binder_tr: string * string -> string * (term list -> term)
 | 
|  |     12 |   val mk_binder_tr': string * string -> string * (term list -> term)
 | 
|  |     13 |   val dependent_tr': string * string -> term list -> term
 | 
|  |     14 | end;
 | 
|  |     15 | 
 | 
|  |     16 | signature SYN_TRANS1 =
 | 
|  |     17 | sig
 | 
|  |     18 |   include SYN_TRANS0
 | 
|  |     19 |   structure Parser: PARSER
 | 
|  |     20 |   local open Parser.SynExt.Ast in
 | 
|  |     21 |     val constrainAbsC: string
 | 
|  |     22 |     val pure_trfuns:
 | 
|  |     23 |       (string * (ast list -> ast)) list *
 | 
|  |     24 |       (string * (term list -> term)) list *
 | 
|  |     25 |       (string * (term list -> term)) list *
 | 
|  |     26 |       (string * (ast list -> ast)) list
 | 
|  |     27 |   end
 | 
|  |     28 | end;
 | 
|  |     29 | 
 | 
|  |     30 | signature SYN_TRANS =
 | 
|  |     31 | sig
 | 
|  |     32 |   include SYN_TRANS1
 | 
|  |     33 |   local open Parser Parser.SynExt Parser.SynExt.Ast in
 | 
|  |     34 |     val abs_tr': term -> term
 | 
|  |     35 |     val prop_tr': bool -> term -> term
 | 
|  |     36 |     val appl_ast_tr': ast * ast list -> ast
 | 
|  |     37 |     val pt_to_ast: (string -> (ast list -> ast) option) -> parsetree -> ast
 | 
|  |     38 |     val ast_to_term: (string -> (term list -> term) option) -> ast -> term
 | 
|  |     39 |   end
 | 
|  |     40 | end;
 | 
|  |     41 | 
 | 
|  |     42 | functor SynTransFun(structure TypeExt: TYPE_EXT and Parser: PARSER
 | 
|  |     43 |   sharing TypeExt.SynExt = Parser.SynExt): SYN_TRANS =
 | 
|  |     44 | struct
 | 
|  |     45 | 
 | 
|  |     46 | structure Parser = Parser;
 | 
|  |     47 | open TypeExt Parser.Lexicon Parser.SynExt.Ast Parser.SynExt Parser;
 | 
|  |     48 | 
 | 
|  |     49 | 
 | 
|  |     50 | (** parse (ast) translations **)
 | 
|  |     51 | 
 | 
|  |     52 | (* application *)
 | 
|  |     53 | 
 | 
|  |     54 | fun appl_ast_tr (*"_appl"*) [f, args] = Appl (f :: unfold_ast "_args" args)
 | 
|  |     55 |   | appl_ast_tr (*"_appl"*) asts = raise_ast "appl_ast_tr" asts;
 | 
|  |     56 | 
 | 
|  |     57 | 
 | 
|  |     58 | (* abstraction *)
 | 
|  |     59 | 
 | 
|  |     60 | fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Appl [Constant constrainC, x, ty]
 | 
|  |     61 |   | idtyp_ast_tr (*"_idtyp"*) asts = raise_ast "idtyp_ast_tr" asts;
 | 
|  |     62 | 
 | 
|  |     63 | fun lambda_ast_tr (*"_lambda"*) [idts, body] =
 | 
|  |     64 |       fold_ast_p "_abs" (unfold_ast "_idts" idts, body)
 | 
|  |     65 |   | lambda_ast_tr (*"_lambda"*) asts = raise_ast "lambda_ast_tr" asts;
 | 
|  |     66 | 
 | 
|  |     67 | val constrainAbsC = "_constrainAbs";
 | 
|  |     68 | 
 | 
|  |     69 | fun abs_tr (*"_abs"*) [Free (x, T), body] = absfree (x, T, body)
 | 
|  |     70 |   | abs_tr (*"_abs"*) (ts as [Const (c, _) $ Free (x, T) $ tT, body]) =
 | 
|  |     71 |       if c = constrainC
 | 
|  |     72 |         then const constrainAbsC $ absfree (x, T, body) $ tT
 | 
|  |     73 |       else raise_term "abs_tr" ts
 | 
|  |     74 |   | abs_tr (*"_abs"*) ts = raise_term "abs_tr" ts;
 | 
|  |     75 | 
 | 
|  |     76 | 
 | 
|  |     77 | (* nondependent abstraction *)
 | 
|  |     78 | 
 | 
|  |     79 | fun k_tr (*"_K"*) [t] = Abs ("uu", dummyT, incr_boundvars 1 t)
 | 
|  |     80 |   | k_tr (*"_K"*) ts = raise_term "k_tr" ts;
 | 
|  |     81 | 
 | 
|  |     82 | 
 | 
|  |     83 | (* binder *)
 | 
|  |     84 | 
 | 
|  |     85 | fun mk_binder_tr (sy, name) =
 | 
|  |     86 |   let
 | 
|  |     87 |     fun tr (Free (x, T), t) = const name $ absfree (x, T, t)
 | 
|  |     88 |       | tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t))
 | 
|  |     89 |       | tr (t1 as Const (c, _) $ Free (x, T) $ tT, t) =
 | 
|  |     90 |           if c = constrainC then
 | 
|  |     91 |             const name $ (const constrainAbsC $ absfree (x, T, t) $ tT)
 | 
|  |     92 |           else raise_term "binder_tr" [t1, t]
 | 
|  |     93 |       | tr (t1, t2) = raise_term "binder_tr" [t1, t2];
 | 
|  |     94 | 
 | 
|  |     95 |     fun binder_tr (*sy*) [idts, body] = tr (idts, body)
 | 
|  |     96 |       | binder_tr (*sy*) ts = raise_term "binder_tr" ts;
 | 
|  |     97 |   in
 | 
|  |     98 |     (sy, binder_tr)
 | 
|  |     99 |   end;
 | 
|  |    100 | 
 | 
|  |    101 | 
 | 
|  |    102 | (* meta propositions *)
 | 
|  |    103 | 
 | 
|  |    104 | fun aprop_tr (*"_aprop"*) [t] = const constrainC $ t $ const "prop"
 | 
|  |    105 |   | aprop_tr (*"_aprop"*) ts = raise_term "aprop_tr" ts;
 | 
|  |    106 | 
 | 
|  |    107 | fun ofclass_tr (*"_ofclass"*) [ty, cls] =
 | 
|  |    108 |       cls $ (const constrainC $ const "TYPE" $ (const "itself" $ ty))
 | 
|  |    109 |   | ofclass_tr (*"_ofclass"*) ts = raise_term "ofclass_tr" ts;
 | 
|  |    110 | 
 | 
|  |    111 | 
 | 
|  |    112 | (* meta implication *)
 | 
|  |    113 | 
 | 
|  |    114 | fun bigimpl_ast_tr (*"_bigimpl"*) [asms, concl] =
 | 
|  |    115 |       fold_ast_p "==>" (unfold_ast "_asms" asms, concl)
 | 
|  |    116 |   | bigimpl_ast_tr (*"_bigimpl"*) asts = raise_ast "bigimpl_ast_tr" asts;
 | 
|  |    117 | 
 | 
|  |    118 | 
 | 
|  |    119 | 
 | 
|  |    120 | (** print (ast) translations **)
 | 
|  |    121 | 
 | 
|  |    122 | (* application *)
 | 
|  |    123 | 
 | 
|  |    124 | fun appl_ast_tr' (f, []) = raise_ast "appl_ast_tr'" [f]
 | 
|  |    125 |   | appl_ast_tr' (f, args) = Appl [Constant "_appl", f, fold_ast "_args" args];
 | 
|  |    126 | 
 | 
|  |    127 | 
 | 
|  |    128 | (* abstraction *)
 | 
|  |    129 | 
 | 
|  |    130 | fun strip_abss vars_of body_of tm =
 | 
|  |    131 |   let
 | 
|  |    132 |     val vars = vars_of tm;
 | 
|  |    133 |     val body = body_of tm;
 | 
|  |    134 |     val rev_new_vars = rename_wrt_term body vars;
 | 
|  |    135 |   in
 | 
|  |    136 |     (map Free (rev rev_new_vars),
 | 
|  |    137 |       subst_bounds (map (free o #1) rev_new_vars, body))
 | 
|  |    138 |   end;
 | 
|  |    139 | 
 | 
|  |    140 | (*do (partial) eta-contraction before printing*)
 | 
|  |    141 | 
 | 
|  |    142 | val eta_contract = ref false;
 | 
|  |    143 | 
 | 
|  |    144 | fun eta_contr tm =
 | 
|  |    145 |   let
 | 
|  |    146 |     fun eta_abs (Abs (a, T, t)) =
 | 
|  |    147 |           (case eta_abs t of
 | 
|  |    148 |             t' as f $ u =>
 | 
|  |    149 |               (case eta_abs u of
 | 
|  |    150 |                 Bound 0 =>
 | 
|  |    151 |                   if not (0 mem loose_bnos f) then incr_boundvars ~1 f
 | 
|  |    152 |                   else Abs (a, T, t')
 | 
|  |    153 |               | _ => Abs (a, T, t'))
 | 
|  |    154 |           | t' => Abs (a, T, t'))
 | 
|  |    155 |       | eta_abs t = t;
 | 
|  |    156 |   in
 | 
|  |    157 |     if ! eta_contract then eta_abs tm else tm
 | 
|  |    158 |   end;
 | 
|  |    159 | 
 | 
|  |    160 | 
 | 
|  |    161 | fun abs_tr' tm =
 | 
|  |    162 |   foldr (fn (x, t) => const "_abs" $ x $ t)
 | 
|  |    163 |     (strip_abss strip_abs_vars strip_abs_body (eta_contr tm));
 | 
|  |    164 | 
 | 
|  |    165 | 
 | 
|  |    166 | fun abs_ast_tr' (*"_abs"*) asts =
 | 
|  |    167 |   (case unfold_ast_p "_abs" (Appl (Constant "_abs" :: asts)) of
 | 
|  |    168 |     ([], _) => raise_ast "abs_ast_tr'" asts
 | 
|  |    169 |   | (xs, body) => Appl [Constant "_lambda", fold_ast "_idts" xs, body]);
 | 
|  |    170 | 
 | 
|  |    171 | 
 | 
|  |    172 | (* binder *)
 | 
|  |    173 | 
 | 
|  |    174 | fun mk_binder_tr' (name, sy) =
 | 
|  |    175 |   let
 | 
|  |    176 |     fun mk_idts [] = raise Match    (*abort translation*)
 | 
|  |    177 |       | mk_idts [idt] = idt
 | 
|  |    178 |       | mk_idts (idt :: idts) = const "_idts" $ idt $ mk_idts idts;
 | 
|  |    179 | 
 | 
|  |    180 |     fun tr' t =
 | 
|  |    181 |       let
 | 
|  |    182 |         val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t;
 | 
|  |    183 |       in
 | 
|  |    184 |         const sy $ mk_idts xs $ bd
 | 
|  |    185 |       end;
 | 
|  |    186 | 
 | 
|  |    187 |     fun binder_tr' (*name*) (t :: ts) =
 | 
|  |    188 |           list_comb (tr' (const name $ t), ts)
 | 
|  |    189 |       | binder_tr' (*name*) [] = raise Match;
 | 
|  |    190 |   in
 | 
|  |    191 |     (name, binder_tr')
 | 
|  |    192 |   end;
 | 
|  |    193 | 
 | 
|  |    194 | 
 | 
|  |    195 | (* idts *)
 | 
|  |    196 | 
 | 
|  |    197 | fun idts_ast_tr' (*"_idts"*) [Appl [Constant c, x, ty], xs] =
 | 
|  |    198 |       if c = constrainC then
 | 
|  |    199 |         Appl [Constant "_idts", Appl [Constant "_idtyp", x, ty], xs]
 | 
|  |    200 |       else raise Match
 | 
|  |    201 |   | idts_ast_tr' (*"_idts"*) _ = raise Match;
 | 
|  |    202 | 
 | 
|  |    203 | 
 | 
|  |    204 | (* meta propositions *)
 | 
|  |    205 | 
 | 
|  |    206 | fun prop_tr' show_sorts tm =
 | 
|  |    207 |   let
 | 
|  |    208 |     fun aprop t = const "_aprop" $ t;
 | 
|  |    209 | 
 | 
|  |    210 |     fun is_prop tys t =
 | 
|  |    211 |       fastype_of1 (tys, t) = propT handle TERM _ => false;
 | 
|  |    212 | 
 | 
|  |    213 |     fun tr' _ (t as Const _) = t
 | 
|  |    214 |       | tr' _ (t as Free (x, ty)) =
 | 
|  |    215 |           if ty = propT then aprop (free x) else t
 | 
|  |    216 |       | tr' _ (t as Var (xi, ty)) =
 | 
|  |    217 |           if ty = propT then aprop (var xi) else t
 | 
|  |    218 |       | tr' tys (t as Bound _) =
 | 
|  |    219 |           if is_prop tys t then aprop t else t
 | 
|  |    220 |       | tr' tys (Abs (x, ty, t)) = Abs (x, ty, tr' (ty :: tys) t)
 | 
|  |    221 |       | tr' tys (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [ty])))) =
 | 
|  |    222 |           if is_prop tys t then
 | 
|  |    223 |             const "_ofclass" $ term_of_typ show_sorts ty $ tr' tys t1
 | 
|  |    224 |           else tr' tys t1 $ tr' tys t2
 | 
|  |    225 |       | tr' tys (t as t1 $ t2) =
 | 
|  |    226 |           (if is_Const (head_of t) orelse not (is_prop tys t)
 | 
|  |    227 |             then I else aprop) (tr' tys t1 $ tr' tys t2);
 | 
|  |    228 |   in
 | 
|  |    229 |     tr' [] tm
 | 
|  |    230 |   end;
 | 
|  |    231 | 
 | 
|  |    232 | 
 | 
|  |    233 | (* meta implication *)
 | 
|  |    234 | 
 | 
|  |    235 | fun impl_ast_tr' (*"==>"*) asts =
 | 
|  |    236 |   (case unfold_ast_p "==>" (Appl (Constant "==>" :: asts)) of
 | 
|  |    237 |     (asms as _ :: _ :: _, concl)
 | 
|  |    238 |       => Appl [Constant "_bigimpl", fold_ast "_asms" asms, concl]
 | 
|  |    239 |   | _ => raise Match);
 | 
|  |    240 | 
 | 
|  |    241 | 
 | 
|  |    242 | (* dependent / nondependent quantifiers *)
 | 
|  |    243 | 
 | 
|  |    244 | fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) =
 | 
|  |    245 |       if 0 mem (loose_bnos B) then
 | 
|  |    246 |         let val (x', B') = variant_abs (x, dummyT, B);
 | 
|  |    247 |         in list_comb (const q $ Free (x', T) $ A $ B', ts) end
 | 
|  |    248 |       else list_comb (const r $ A $ B, ts)
 | 
|  |    249 |   | dependent_tr' _ _ = raise Match;
 | 
|  |    250 | 
 | 
|  |    251 | 
 | 
|  |    252 | 
 | 
|  |    253 | (** pure_trfuns **)
 | 
|  |    254 | 
 | 
|  |    255 | val pure_trfuns =
 | 
|  |    256 |  ([(applC, appl_ast_tr), ("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr),
 | 
|  |    257 |     ("_bigimpl", bigimpl_ast_tr)],
 | 
| 639 |    258 |   [("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr), ("_K", k_tr)],
 | 
| 548 |    259 |   [],
 | 
| 639 |    260 |   [("_abs", abs_ast_tr'), ("_idts", idts_ast_tr'), ("==>", impl_ast_tr')]);
 | 
| 548 |    261 | 
 | 
|  |    262 | 
 | 
|  |    263 | 
 | 
|  |    264 | (** pt_to_ast **)
 | 
|  |    265 | 
 | 
|  |    266 | fun pt_to_ast trf pt =
 | 
|  |    267 |   let
 | 
|  |    268 |     fun trans a args =
 | 
|  |    269 |       (case trf a of
 | 
|  |    270 |         None => mk_appl (Constant a) args
 | 
|  |    271 |       | Some f => f args handle exn
 | 
|  |    272 |           => (writeln ("Error in parse ast translation for " ^ quote a); raise exn));
 | 
|  |    273 | 
 | 
|  |    274 |     fun ast_of (Node (a, pts)) = trans a (map ast_of pts)
 | 
|  |    275 |       | ast_of (Tip tok) = Variable (str_of_token tok);
 | 
|  |    276 |   in
 | 
|  |    277 |     ast_of pt
 | 
|  |    278 |   end;
 | 
|  |    279 | 
 | 
|  |    280 | 
 | 
|  |    281 | 
 | 
|  |    282 | (** ast_to_term **)
 | 
|  |    283 | 
 | 
|  |    284 | fun ast_to_term trf ast =
 | 
|  |    285 |   let
 | 
|  |    286 |     fun trans a args =
 | 
|  |    287 |       (case trf a of
 | 
|  |    288 |         None => list_comb (const a, args)
 | 
|  |    289 |       | Some f => f args handle exn
 | 
|  |    290 |           => (writeln ("Error in parse translation for " ^ quote a); raise exn));
 | 
|  |    291 | 
 | 
|  |    292 |     fun term_of (Constant a) = trans a []
 | 
|  |    293 |       | term_of (Variable x) = scan_var x
 | 
|  |    294 |       | term_of (Appl (Constant a :: (asts as _ :: _))) =
 | 
|  |    295 |           trans a (map term_of asts)
 | 
|  |    296 |       | term_of (Appl (ast :: (asts as _ :: _))) =
 | 
|  |    297 |           list_comb (term_of ast, map term_of asts)
 | 
|  |    298 |       | term_of (ast as Appl _) = raise_ast "ast_to_term: malformed ast" [ast];
 | 
|  |    299 |   in
 | 
|  |    300 |     term_of ast
 | 
|  |    301 |   end;
 | 
|  |    302 | 
 | 
|  |    303 | 
 | 
|  |    304 | end;
 | 
|  |    305 | 
 |