| author | ehmety | 
| Thu, 15 Nov 2001 16:46:38 +0100 | |
| changeset 12197 | d9320fb0a570 | 
| parent 12150 | f83dc4202b78 | 
| child 12785 | 27debaf2112d | 
| permissions | -rw-r--r-- | 
| 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 = | |
| 2698 | 9 | sig | 
| 548 | 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 | |
| 8577 | 14 | val antiquote_tr: string -> term -> term | 
| 15 | val quote_tr: string -> term -> term | |
| 5084 | 16 | val quote_antiquote_tr: string -> string -> string -> string * (term list -> term) | 
| 8577 | 17 | val antiquote_tr': string -> term -> term | 
| 18 | val quote_tr': string -> term -> term | |
| 5084 | 19 | val quote_antiquote_tr': string -> string -> string -> string * (term list -> term) | 
| 2698 | 20 | val mark_bound: string -> term | 
| 21 | val mark_boundT: string * typ -> term | |
| 22 | val variant_abs': string * typ * term -> string * term | |
| 23 | end; | |
| 548 | 24 | |
| 25 | signature SYN_TRANS1 = | |
| 2698 | 26 | sig | 
| 548 | 27 | include SYN_TRANS0 | 
| 1511 | 28 | val constrainAbsC: string | 
| 29 | val pure_trfuns: | |
| 30 | (string * (Ast.ast list -> Ast.ast)) list * | |
| 548 | 31 | (string * (term list -> term)) list * | 
| 32 | (string * (term list -> term)) list * | |
| 1511 | 33 | (string * (Ast.ast list -> Ast.ast)) list | 
| 4148 | 34 | val pure_trfunsT: (string * (bool -> typ -> term list -> term)) list | 
| 2698 | 35 | end; | 
| 548 | 36 | |
| 37 | signature SYN_TRANS = | |
| 2698 | 38 | sig | 
| 548 | 39 | include SYN_TRANS1 | 
| 1511 | 40 | val abs_tr': term -> term | 
| 4148 | 41 | val prop_tr': term -> term | 
| 1511 | 42 | val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast | 
| 43 | val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast | |
| 44 | val pt_to_ast: (string -> (Ast.ast list -> Ast.ast) option) -> Parser.parsetree -> Ast.ast | |
| 45 | val ast_to_term: (string -> (term list -> term) option) -> Ast.ast -> term | |
| 2698 | 46 | end; | 
| 548 | 47 | |
| 2698 | 48 | structure SynTrans: SYN_TRANS = | 
| 548 | 49 | struct | 
| 2698 | 50 | |
| 51 | ||
| 548 | 52 | (** parse (ast) translations **) | 
| 53 | ||
| 11491 | 54 | (* constify *) | 
| 55 | ||
| 56 | fun constify_ast_tr [Ast.Variable c] = Ast.Constant c | |
| 57 |   | constify_ast_tr asts = raise Ast.AST ("constify_ast_tr", asts);
 | |
| 58 | ||
| 59 | ||
| 548 | 60 | (* application *) | 
| 61 | ||
| 5690 | 62 | fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args) | 
| 63 |   | appl_ast_tr asts = raise Ast.AST ("appl_ast_tr", asts);
 | |
| 922 
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
 clasohm parents: 
639diff
changeset | 64 | |
| 5690 | 65 | fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args) | 
| 66 |   | applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts);
 | |
| 548 | 67 | |
| 68 | ||
| 69 | (* abstraction *) | |
| 70 | ||
| 5690 | 71 | fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Ast.Appl [Ast.Constant SynExt.constrainC, x, ty] | 
| 72 |   | idtyp_ast_tr (*"_idtyp"*) asts = raise Ast.AST ("idtyp_ast_tr", asts);
 | |
| 548 | 73 | |
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 74 | fun lambda_ast_tr (*"_lambda"*) [pats, body] = | 
| 5690 | 75 | Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body) | 
| 76 |   | lambda_ast_tr (*"_lambda"*) asts = raise Ast.AST ("lambda_ast_tr", asts);
 | |
| 548 | 77 | |
| 78 | val constrainAbsC = "_constrainAbs"; | |
| 79 | ||
| 5690 | 80 | fun abs_tr (*"_abs"*) [Free (x, T), body] = Term.absfree (x, T, body) | 
| 548 | 81 | | abs_tr (*"_abs"*) (ts as [Const (c, _) $ Free (x, T) $ tT, body]) = | 
| 5690 | 82 | if c = SynExt.constrainC | 
| 83 | then Lexicon.const constrainAbsC $ Term.absfree (x, T, body) $ tT | |
| 3777 | 84 |       else raise TERM ("abs_tr", ts)
 | 
| 85 |   | abs_tr (*"_abs"*) ts = raise TERM ("abs_tr", ts);
 | |
| 548 | 86 | |
| 87 | ||
| 88 | (* nondependent abstraction *) | |
| 89 | ||
| 5690 | 90 | fun k_tr (*"_K"*) [t] = Abs ("uu", dummyT, Term.incr_boundvars 1 t)
 | 
| 3777 | 91 |   | k_tr (*"_K"*) ts = raise TERM ("k_tr", ts);
 | 
| 548 | 92 | |
| 93 | ||
| 94 | (* binder *) | |
| 95 | ||
| 96 | fun mk_binder_tr (sy, name) = | |
| 97 | let | |
| 5690 | 98 | fun tr (Free (x, T), t) = Lexicon.const name $ Term.absfree (x, T, t) | 
| 548 | 99 |       | tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t))
 | 
| 100 | | tr (t1 as Const (c, _) $ Free (x, T) $ tT, t) = | |
| 5690 | 101 | if c = SynExt.constrainC then | 
| 102 | Lexicon.const name $ (Lexicon.const constrainAbsC $ Term.absfree (x, T, t) $ tT) | |
| 3777 | 103 |           else raise TERM ("binder_tr", [t1, t])
 | 
| 104 |       | tr (t1, t2) = raise TERM ("binder_tr", [t1, t2]);
 | |
| 548 | 105 | |
| 106 | fun binder_tr (*sy*) [idts, body] = tr (idts, body) | |
| 3777 | 107 |       | binder_tr (*sy*) ts = raise TERM ("binder_tr", ts);
 | 
| 548 | 108 | in | 
| 109 | (sy, binder_tr) | |
| 110 | end; | |
| 111 | ||
| 112 | ||
| 113 | (* meta propositions *) | |
| 114 | ||
| 5690 | 115 | fun aprop_tr (*"_aprop"*) [t] = Lexicon.const SynExt.constrainC $ t $ Lexicon.const "prop" | 
| 3777 | 116 |   | aprop_tr (*"_aprop"*) ts = raise TERM ("aprop_tr", ts);
 | 
| 548 | 117 | |
| 118 | fun ofclass_tr (*"_ofclass"*) [ty, cls] = | |
| 5690 | 119 | cls $ (Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $ | 
| 120 | (Lexicon.const "itself" $ ty)) | |
| 3777 | 121 |   | ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts);
 | 
| 548 | 122 | |
| 123 | ||
| 124 | (* meta implication *) | |
| 125 | ||
| 126 | fun bigimpl_ast_tr (*"_bigimpl"*) [asms, concl] = | |
| 5690 | 127 | Ast.fold_ast_p "==>" (Ast.unfold_ast "_asms" asms, concl) | 
| 128 |   | bigimpl_ast_tr (*"_bigimpl"*) asts = raise Ast.AST ("bigimpl_ast_tr", asts);
 | |
| 548 | 129 | |
| 130 | ||
| 4148 | 131 | (* type reflection *) | 
| 132 | ||
| 133 | fun type_tr (*"_TYPE"*) [ty] = | |
| 5690 | 134 | Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $ (Lexicon.const "itself" $ ty) | 
| 4148 | 135 |   | type_tr (*"_TYPE"*) ts = raise TERM ("type_tr", ts);
 | 
| 136 | ||
| 548 | 137 | |
| 6761 | 138 | (* dddot *) | 
| 139 | ||
| 8595 | 140 | fun dddot_tr (*"_DDDOT"*) ts = Term.list_comb (Lexicon.var SynExt.dddot_indexname, ts); | 
| 6761 | 141 | |
| 142 | ||
| 5084 | 143 | (* quote / antiquote *) | 
| 144 | ||
| 8577 | 145 | fun antiquote_tr name = | 
| 146 | let | |
| 147 | fun tr i ((t as Const (c, _)) $ u) = | |
| 148 | if c = name then tr i u $ Bound i | |
| 149 | else tr i t $ tr i u | |
| 150 | | tr i (t $ u) = tr i t $ tr i u | |
| 151 | | tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t) | |
| 152 | | tr _ a = a; | |
| 153 | in tr 0 end; | |
| 154 | ||
| 155 | fun quote_tr name t = Abs ("s", dummyT, antiquote_tr name (Term.incr_boundvars 1 t));
 | |
| 156 | ||
| 5084 | 157 | fun quote_antiquote_tr quoteN antiquoteN name = | 
| 158 | let | |
| 8577 | 159 | fun tr [t] = Lexicon.const name $ quote_tr antiquoteN t | 
| 160 |       | tr ts = raise TERM ("quote_tr", ts);
 | |
| 161 | in (quoteN, tr) end; | |
| 5084 | 162 | |
| 163 | ||
| 12122 | 164 | (* index variable *) | 
| 165 | ||
| 166 | fun indexvar_ast_tr [] = Ast.Variable "some_index" | |
| 167 |   | indexvar_ast_tr asts = raise Ast.AST ("indexvar_ast_tr", asts);
 | |
| 168 | ||
| 169 | ||
| 5084 | 170 | |
| 548 | 171 | (** print (ast) translations **) | 
| 172 | ||
| 173 | (* application *) | |
| 174 | ||
| 5690 | 175 | fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f])
 | 
| 176 | | appl_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_appl", f, Ast.fold_ast "_args" args]; | |
| 548 | 177 | |
| 5690 | 178 | fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f])
 | 
| 179 | | applC_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_applC", f, Ast.fold_ast "_cargs" args]; | |
| 922 
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
 clasohm parents: 
639diff
changeset | 180 | |
| 548 | 181 | |
| 182 | (* abstraction *) | |
| 183 | ||
| 5690 | 184 | fun mark_boundT x_T = Lexicon.const "_bound" $ Free x_T; | 
| 2698 | 185 | fun mark_bound x = mark_boundT (x, dummyT); | 
| 186 | ||
| 548 | 187 | fun strip_abss vars_of body_of tm = | 
| 188 | let | |
| 189 | val vars = vars_of tm; | |
| 190 | val body = body_of tm; | |
| 191 | val rev_new_vars = rename_wrt_term body vars; | |
| 192 | in | |
| 2698 | 193 | (map mark_boundT (rev rev_new_vars), | 
| 194 | subst_bounds (map (mark_bound o #1) rev_new_vars, body)) | |
| 548 | 195 | end; | 
| 196 | ||
| 3928 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 197 | |
| 548 | 198 | (*do (partial) eta-contraction before printing*) | 
| 199 | ||
| 1326 | 200 | val eta_contract = ref true; | 
| 548 | 201 | |
| 202 | fun eta_contr tm = | |
| 203 | let | |
| 3928 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 204 |     fun is_aprop (Const ("_aprop", _)) = true
 | 
| 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 205 | | is_aprop _ = false; | 
| 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 206 | |
| 548 | 207 | fun eta_abs (Abs (a, T, t)) = | 
| 208 | (case eta_abs t of | |
| 209 | t' as f $ u => | |
| 210 | (case eta_abs u of | |
| 211 | Bound 0 => | |
| 5084 | 212 | if Term.loose_bvar1 (f, 0) orelse is_aprop f then Abs (a, T, t') | 
| 3928 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 213 | else incr_boundvars ~1 f | 
| 548 | 214 | | _ => Abs (a, T, t')) | 
| 215 | | t' => Abs (a, T, t')) | |
| 216 | | eta_abs t = t; | |
| 217 | in | |
| 218 | if ! eta_contract then eta_abs tm else tm | |
| 219 | end; | |
| 220 | ||
| 221 | ||
| 222 | fun abs_tr' tm = | |
| 5690 | 223 | foldr (fn (x, t) => Lexicon.const "_abs" $ x $ t) | 
| 548 | 224 | (strip_abss strip_abs_vars strip_abs_body (eta_contr tm)); | 
| 225 | ||
| 226 | ||
| 227 | fun abs_ast_tr' (*"_abs"*) asts = | |
| 5690 | 228 | (case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of | 
| 229 |     ([], _) => raise Ast.AST ("abs_ast_tr'", asts)
 | |
| 230 | | (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]); | |
| 548 | 231 | |
| 232 | ||
| 233 | (* binder *) | |
| 234 | ||
| 235 | fun mk_binder_tr' (name, sy) = | |
| 236 | let | |
| 237 | fun mk_idts [] = raise Match (*abort translation*) | |
| 238 | | mk_idts [idt] = idt | |
| 5690 | 239 | | mk_idts (idt :: idts) = Lexicon.const "_idts" $ idt $ mk_idts idts; | 
| 548 | 240 | |
| 241 | fun tr' t = | |
| 242 | let | |
| 243 | val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t; | |
| 5690 | 244 | in Lexicon.const sy $ mk_idts xs $ bd end; | 
| 548 | 245 | |
| 5690 | 246 | fun binder_tr' (*name*) (t :: ts) = Term.list_comb (tr' (Lexicon.const name $ t), ts) | 
| 548 | 247 | | binder_tr' (*name*) [] = raise Match; | 
| 248 | in | |
| 249 | (name, binder_tr') | |
| 250 | end; | |
| 251 | ||
| 252 | ||
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 253 | (* idtyp constraints *) | 
| 548 | 254 | |
| 5690 | 255 | fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant c, x, ty], xs] = | 
| 256 | if c = SynExt.constrainC then | |
| 257 | Ast.Appl [ Ast.Constant a, Ast.Appl [Ast.Constant "_idtyp", x, ty], xs] | |
| 548 | 258 | else raise Match | 
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 259 | | idtyp_ast_tr' _ _ = raise Match; | 
| 548 | 260 | |
| 261 | ||
| 262 | (* meta propositions *) | |
| 263 | ||
| 4148 | 264 | fun prop_tr' tm = | 
| 548 | 265 | let | 
| 5690 | 266 | fun aprop t = Lexicon.const "_aprop" $ t; | 
| 548 | 267 | |
| 2698 | 268 | fun is_prop Ts t = | 
| 269 | fastype_of1 (Ts, t) = propT handle TERM _ => false; | |
| 548 | 270 | |
| 271 | fun tr' _ (t as Const _) = t | |
| 2698 | 272 | | tr' _ (t as Free (x, T)) = | 
| 5690 | 273 | if T = propT then aprop (Lexicon.free x) else t | 
| 2698 | 274 | | tr' _ (t as Var (xi, T)) = | 
| 5690 | 275 | if T = propT then aprop (Lexicon.var xi) else t | 
| 2698 | 276 | | tr' Ts (t as Bound _) = | 
| 277 | if is_prop Ts t then aprop t else t | |
| 278 | | tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t) | |
| 279 |       | tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
 | |
| 4148 | 280 |           if is_prop Ts t then Const ("_mk_ofclass", T) $ tr' Ts t1
 | 
| 2698 | 281 | else tr' Ts t1 $ tr' Ts t2 | 
| 282 | | tr' Ts (t as t1 $ t2) = | |
| 5690 | 283 | (if is_Const (Term.head_of t) orelse not (is_prop Ts t) | 
| 2698 | 284 | then I else aprop) (tr' Ts t1 $ tr' Ts t2); | 
| 548 | 285 | in | 
| 286 | tr' [] tm | |
| 287 | end; | |
| 288 | ||
| 4148 | 289 | fun mk_ofclass_tr' show_sorts (*"_mk_ofclass"*) T [t] = | 
| 5690 | 290 | Lexicon.const "_ofclass" $ TypeExt.term_of_typ show_sorts T $ t | 
| 4148 | 291 |   | mk_ofclass_tr' _ (*"_mk_ofclass"*) T ts = raise TYPE ("mk_ofclass_tr'", [T], ts);
 | 
| 2698 | 292 | |
| 293 | ||
| 548 | 294 | (* meta implication *) | 
| 295 | ||
| 296 | fun impl_ast_tr' (*"==>"*) asts = | |
| 10572 | 297 | if TypeExt.no_brackets () then raise Match | 
| 298 | else | |
| 299 | (case Ast.unfold_ast_p "==>" (Ast.Appl (Ast.Constant "==>" :: asts)) of | |
| 300 | (asms as _ :: _ :: _, concl) | |
| 301 | => Ast.Appl [Ast.Constant "_bigimpl", Ast.fold_ast "_asms" asms, concl] | |
| 302 | | _ => raise Match); | |
| 548 | 303 | |
| 304 | ||
| 12150 | 305 | (* meta conjunction *) | 
| 306 | ||
| 307 | fun meta_conjunction_tr' (*"all"*) | |
| 308 |       [Abs (_, _, Const ("==>", _) $
 | |
| 309 |         (Const ("==>", _) $ A $ (Const ("==>", _) $ B $ (Const ("_aprop", _) $ Bound 0))) $
 | |
| 310 |         (Const ("_aprop", _) $ Bound 0))] =
 | |
| 311 | if 0 mem_int Term.loose_bnos A orelse 0 mem_int Term.loose_bnos B then raise Match | |
| 312 | else Lexicon.const "_meta_conjunction" $ A $ B | |
| 313 | | meta_conjunction_tr' (*"all"*) ts = raise Match; | |
| 314 | ||
| 315 | ||
| 4148 | 316 | (* type reflection *) | 
| 317 | ||
| 318 | fun type_tr' show_sorts (*"TYPE"*) (Type ("itself", [T])) ts =
 | |
| 5690 | 319 | Term.list_comb (Lexicon.const "_TYPE" $ TypeExt.term_of_typ show_sorts T, ts) | 
| 4148 | 320 | | type_tr' _ _ _ = raise Match; | 
| 321 | ||
| 322 | ||
| 548 | 323 | (* dependent / nondependent quantifiers *) | 
| 324 | ||
| 2698 | 325 | fun variant_abs' (x, T, B) = | 
| 326 | let val x' = variant (add_term_names (B, [])) x in | |
| 327 | (x', subst_bound (mark_boundT (x', T), B)) | |
| 328 | end; | |
| 329 | ||
| 548 | 330 | fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) = | 
| 5084 | 331 | if Term.loose_bvar1 (B, 0) then | 
| 2698 | 332 | let val (x', B') = variant_abs' (x, dummyT, B); | 
| 5690 | 333 | in Term.list_comb (Lexicon.const q $ mark_boundT (x', T) $ A $ B', ts) end | 
| 334 | else Term.list_comb (Lexicon.const r $ A $ B, ts) | |
| 548 | 335 | | dependent_tr' _ _ = raise Match; | 
| 336 | ||
| 337 | ||
| 5084 | 338 | (* quote / antiquote *) | 
| 339 | ||
| 8577 | 340 | fun antiquote_tr' name = | 
| 341 | let | |
| 342 | fun tr' i (t $ u) = | |
| 343 | if u = Bound i then Lexicon.const name $ tr' i t | |
| 344 | else tr' i t $ tr' i u | |
| 345 | | tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t) | |
| 346 | | tr' i a = if a = Bound i then raise Match else a; | |
| 347 | in tr' 0 end; | |
| 348 | ||
| 349 | fun quote_tr' name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' name t) | |
| 350 | | quote_tr' _ _ = raise Match; | |
| 351 | ||
| 5084 | 352 | fun quote_antiquote_tr' quoteN antiquoteN name = | 
| 353 | let | |
| 8577 | 354 | fun tr' (t :: ts) = Term.list_comb (Lexicon.const quoteN $ quote_tr' antiquoteN t, ts) | 
| 355 | | tr' _ = raise Match; | |
| 356 | in (name, tr') end; | |
| 5084 | 357 | |
| 358 | ||
| 548 | 359 | |
| 360 | (** pure_trfuns **) | |
| 361 | ||
| 362 | val pure_trfuns = | |
| 11491 | 363 |  ([("_constify", constify_ast_tr), ("_appl", appl_ast_tr), ("_applC", applC_ast_tr),
 | 
| 12122 | 364 |    ("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr), ("_bigimpl", bigimpl_ast_tr),
 | 
| 365 |    ("_indexvar", indexvar_ast_tr)],
 | |
| 922 
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
 clasohm parents: 
639diff
changeset | 366 |   [("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr),
 | 
| 6761 | 367 |    ("_TYPE", type_tr), ("_DDDOT", dddot_tr), ("_K", k_tr)],
 | 
| 12150 | 368 |   [("all", meta_conjunction_tr')],
 | 
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 369 |   [("_abs", abs_ast_tr'), ("_idts", idtyp_ast_tr' "_idts"),
 | 
| 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 370 |    ("_pttrns", idtyp_ast_tr' "_pttrns"), ("==>", impl_ast_tr')]);
 | 
| 548 | 371 | |
| 2698 | 372 | val pure_trfunsT = | 
| 4148 | 373 |   [("_mk_ofclass", mk_ofclass_tr'), ("TYPE", type_tr')];
 | 
| 2698 | 374 | |
| 548 | 375 | |
| 376 | ||
| 377 | (** pt_to_ast **) | |
| 378 | ||
| 379 | fun pt_to_ast trf pt = | |
| 380 | let | |
| 381 | fun trans a args = | |
| 382 | (case trf a of | |
| 5690 | 383 | None => Ast.mk_appl (Ast.Constant a) args | 
| 548 | 384 | | Some f => f args handle exn | 
| 987 | 385 |           => (writeln ("Error in parse ast translation for " ^ quote a);
 | 
| 386 | raise exn)); | |
| 548 | 387 | |
| 987 | 388 | (*translate pt bottom-up*) | 
| 5690 | 389 | fun ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts) | 
| 390 | | ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok); | |
| 548 | 391 | in | 
| 1178 | 392 | ast_of pt | 
| 548 | 393 | end; | 
| 394 | ||
| 395 | ||
| 396 | ||
| 397 | (** ast_to_term **) | |
| 398 | ||
| 399 | fun ast_to_term trf ast = | |
| 400 | let | |
| 401 | fun trans a args = | |
| 402 | (case trf a of | |
| 5690 | 403 | None => Term.list_comb (Lexicon.const a, args) | 
| 548 | 404 | | Some f => f args handle exn | 
| 987 | 405 |           => (writeln ("Error in parse translation for " ^ quote a);
 | 
| 406 | raise exn)); | |
| 548 | 407 | |
| 5690 | 408 | fun term_of (Ast.Constant a) = trans a [] | 
| 409 | | term_of (Ast.Variable x) = Lexicon.read_var x | |
| 410 | | term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) = | |
| 548 | 411 | trans a (map term_of asts) | 
| 5690 | 412 | | term_of (Ast.Appl (ast :: (asts as _ :: _))) = | 
| 413 | Term.list_comb (term_of ast, map term_of asts) | |
| 414 |       | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]);
 | |
| 548 | 415 | in | 
| 416 | term_of ast | |
| 417 | end; | |
| 418 | ||
| 419 | end; |