| author | bulwahn | 
| Sat, 24 Oct 2009 16:55:43 +0200 | |
| changeset 33145 | 1a22f7ca1dfc | 
| parent 32786 | f1ac4b515af9 | 
| child 35145 | f132a4fd8679 | 
| permissions | -rw-r--r-- | 
| 548 | 1 | (* Title: Pure/Syntax/syn_trans.ML | 
| 2 | Author: Tobias Nipkow and Markus Wenzel, TU Muenchen | |
| 3 | ||
| 4 | Syntax translation functions. | |
| 5 | *) | |
| 6 | ||
| 7 | signature SYN_TRANS0 = | |
| 2698 | 8 | sig | 
| 32738 | 9 | val eta_contract: bool Unsynchronized.ref | 
| 13762 | 10 | val atomic_abs_tr': string * typ * term -> term * term | 
| 32120 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 11 | val preserve_binder_abs_tr': string -> string -> string * (term list -> term) | 
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 12 | val preserve_binder_abs2_tr': string -> string -> string * (term list -> term) | 
| 548 | 13 | val mk_binder_tr: string * string -> string * (term list -> term) | 
| 14 | val mk_binder_tr': string * string -> string * (term list -> term) | |
| 15 | val dependent_tr': string * string -> term list -> term | |
| 8577 | 16 | val antiquote_tr: string -> term -> term | 
| 17 | val quote_tr: string -> term -> term | |
| 5084 | 18 | val quote_antiquote_tr: string -> string -> string -> string * (term list -> term) | 
| 8577 | 19 | val antiquote_tr': string -> term -> term | 
| 20 | val quote_tr': string -> term -> term | |
| 5084 | 21 | val quote_antiquote_tr': string -> string -> string -> string * (term list -> term) | 
| 2698 | 22 | val mark_bound: string -> term | 
| 23 | val mark_boundT: string * typ -> term | |
| 18958 | 24 | val bound_vars: (string * typ) list -> term -> term | 
| 20202 | 25 | val variant_abs: string * typ * term -> string * term | 
| 2698 | 26 | val variant_abs': string * typ * term -> string * term | 
| 27 | end; | |
| 548 | 28 | |
| 29 | signature SYN_TRANS1 = | |
| 2698 | 30 | sig | 
| 548 | 31 | include SYN_TRANS0 | 
| 14647 | 32 | val non_typed_tr': (term list -> term) -> bool -> typ -> term list -> term | 
| 33 |   val non_typed_tr'': ('a -> term list -> term) -> 'a -> bool -> typ -> term list -> term
 | |
| 1511 | 34 | val constrainAbsC: string | 
| 35 | val pure_trfuns: | |
| 36 | (string * (Ast.ast list -> Ast.ast)) list * | |
| 548 | 37 | (string * (term list -> term)) list * | 
| 38 | (string * (term list -> term)) list * | |
| 1511 | 39 | (string * (Ast.ast list -> Ast.ast)) list | 
| 4148 | 40 | val pure_trfunsT: (string * (bool -> typ -> term list -> term)) list | 
| 14697 | 41 | val struct_trfuns: string list -> | 
| 42 | (string * (Ast.ast list -> Ast.ast)) list * | |
| 43 | (string * (term list -> term)) list * | |
| 44 | (string * (bool -> typ -> term list -> term)) list * | |
| 45 | (string * (Ast.ast list -> Ast.ast)) list | |
| 2698 | 46 | end; | 
| 548 | 47 | |
| 48 | signature SYN_TRANS = | |
| 2698 | 49 | sig | 
| 548 | 50 | include SYN_TRANS1 | 
| 1511 | 51 | val abs_tr': term -> term | 
| 4148 | 52 | val prop_tr': term -> term | 
| 1511 | 53 | val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast | 
| 54 | val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast | |
| 21773 | 55 | val pts_to_asts: Proof.context -> | 
| 56 | (string -> (Proof.context -> Ast.ast list -> Ast.ast) option) -> | |
| 16612 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 57 | Parser.parsetree list -> Ast.ast list | 
| 21773 | 58 | val asts_to_terms: Proof.context -> | 
| 59 | (string -> (Proof.context -> term list -> term) option) -> Ast.ast list -> term list | |
| 2698 | 60 | end; | 
| 548 | 61 | |
| 2698 | 62 | structure SynTrans: SYN_TRANS = | 
| 548 | 63 | struct | 
| 2698 | 64 | |
| 65 | ||
| 548 | 66 | (** parse (ast) translations **) | 
| 67 | ||
| 11491 | 68 | (* constify *) | 
| 69 | ||
| 70 | fun constify_ast_tr [Ast.Variable c] = Ast.Constant c | |
| 71 |   | constify_ast_tr asts = raise Ast.AST ("constify_ast_tr", asts);
 | |
| 72 | ||
| 73 | ||
| 548 | 74 | (* application *) | 
| 75 | ||
| 5690 | 76 | fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args) | 
| 77 |   | 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 | 78 | |
| 5690 | 79 | fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args) | 
| 80 |   | applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts);
 | |
| 548 | 81 | |
| 82 | ||
| 83 | (* abstraction *) | |
| 84 | ||
| 17788 | 85 | fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Ast.Appl [Ast.Constant "_constrain", x, ty] | 
| 5690 | 86 |   | idtyp_ast_tr (*"_idtyp"*) asts = raise Ast.AST ("idtyp_ast_tr", asts);
 | 
| 548 | 87 | |
| 17788 | 88 | fun idtypdummy_ast_tr (*"_idtypdummy"*) [ty] = | 
| 89 | Ast.Appl [Ast.Constant "_constrain", Ast.Constant "_idtdummy", ty] | |
| 90 |   | idtypdummy_ast_tr (*"_idtypdummy"*) asts = raise Ast.AST ("idtyp_ast_tr", asts);
 | |
| 91 | ||
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 92 | fun lambda_ast_tr (*"_lambda"*) [pats, body] = | 
| 5690 | 93 | Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body) | 
| 94 |   | lambda_ast_tr (*"_lambda"*) asts = raise Ast.AST ("lambda_ast_tr", asts);
 | |
| 548 | 95 | |
| 96 | val constrainAbsC = "_constrainAbs"; | |
| 97 | ||
| 21773 | 98 | fun absfree_proper (x, T, t) = | 
| 99 |   if can Name.dest_internal x then error ("Illegal internal variable in abstraction: " ^ quote x)
 | |
| 100 | else Term.absfree (x, T, t); | |
| 101 | ||
| 102 | fun abs_tr (*"_abs"*) [Free (x, T), t] = absfree_proper (x, T, t) | |
| 17788 | 103 |   | abs_tr (*"_abs"*) [Const ("_idtdummy", T), t] = Term.absdummy (T, t)
 | 
| 104 |   | abs_tr (*"_abs"*) [Const ("_constrain", _) $ Free (x, T) $ tT, t] =
 | |
| 21773 | 105 | Lexicon.const constrainAbsC $ absfree_proper (x, T, t) $ tT | 
| 17788 | 106 |   | abs_tr (*"_abs"*) [Const ("_constrain", _) $ Const ("_idtdummy", T) $ tT, t] =
 | 
| 107 | Lexicon.const constrainAbsC $ Term.absdummy (T, t) $ tT | |
| 3777 | 108 |   | abs_tr (*"_abs"*) ts = raise TERM ("abs_tr", ts);
 | 
| 548 | 109 | |
| 110 | ||
| 111 | (* binder *) | |
| 112 | ||
| 21535 | 113 | fun mk_binder_tr (syn, name) = | 
| 548 | 114 | let | 
| 21773 | 115 | fun tr (Free (x, T), t) = Lexicon.const name $ absfree_proper (x, T, t) | 
| 17788 | 116 |       | tr (Const ("_idtdummy", T), t) = Lexicon.const name $ Term.absdummy (T, t)
 | 
| 117 |       | tr (Const ("_constrain", _) $ Free (x, T) $ tT, t) =
 | |
| 21773 | 118 | Lexicon.const name $ (Lexicon.const constrainAbsC $ absfree_proper (x, T, t) $ tT) | 
| 17788 | 119 |       | tr (Const ("_constrain", _) $ Const ("_idtdummy", T) $ tT, t) =
 | 
| 18342 
1b7109c10b7b
asts_to_terms: builtin free_fixed translation makes local binders work properly;
 wenzelm parents: 
18212diff
changeset | 120 | Lexicon.const name $ (Lexicon.const constrainAbsC $ Term.absdummy (T, t) $ tT) | 
| 548 | 121 |       | tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t))
 | 
| 3777 | 122 |       | tr (t1, t2) = raise TERM ("binder_tr", [t1, t2]);
 | 
| 548 | 123 | |
| 21535 | 124 | fun binder_tr [idts, body] = tr (idts, body) | 
| 125 |       | binder_tr ts = raise TERM ("binder_tr", ts);
 | |
| 126 | in (syn, binder_tr) end; | |
| 548 | 127 | |
| 128 | ||
| 28628 | 129 | (* type propositions *) | 
| 130 | ||
| 131 | fun mk_type ty = Lexicon.const "_constrain" $ Lexicon.const "TYPE" $ (Lexicon.const "itself" $ ty); | |
| 132 | ||
| 133 | fun ofclass_tr (*"_ofclass"*) [ty, cls] = cls $ mk_type ty | |
| 134 |   | ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts);
 | |
| 135 | ||
| 136 | fun sort_constraint_tr (*"_sort_constraint"*) [ty] = | |
| 137 | Lexicon.const "Pure.sort_constraint" $ mk_type ty | |
| 138 |   | sort_constraint_tr (*"_sort_constraint"*) ts = raise TERM ("sort_constraint_tr", ts);
 | |
| 139 | ||
| 140 | ||
| 548 | 141 | (* meta propositions *) | 
| 142 | ||
| 17788 | 143 | fun aprop_tr (*"_aprop"*) [t] = Lexicon.const "_constrain" $ t $ Lexicon.const "prop" | 
| 3777 | 144 |   | aprop_tr (*"_aprop"*) ts = raise TERM ("aprop_tr", ts);
 | 
| 548 | 145 | |
| 146 | ||
| 147 | (* meta implication *) | |
| 148 | ||
| 16612 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 149 | fun bigimpl_ast_tr (*"_bigimpl"*) (asts as [asms, concl]) = | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 150 | let val prems = | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 151 | (case Ast.unfold_ast_p "_asms" asms of | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 152 | (asms', Ast.Appl [Ast.Constant "_asm", asm']) => asms' @ [asm'] | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 153 |         | _ => raise Ast.AST ("bigimpl_ast_tr", asts))
 | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 154 | in Ast.fold_ast_p "==>" (prems, concl) end | 
| 15421 
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
 schirmer parents: 
14981diff
changeset | 155 |   | bigimpl_ast_tr (*"_bigimpl"*) asts = raise Ast.AST ("bigimpl_ast_tr", asts);
 | 
| 
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
 schirmer parents: 
14981diff
changeset | 156 | |
| 548 | 157 | |
| 23824 
8ad7131dbfcf
moved print_translations from Pure.thy to Syntax/syn_trans.ML;
 wenzelm parents: 
21773diff
changeset | 158 | (* type/term reflection *) | 
| 4148 | 159 | |
| 28628 | 160 | fun type_tr (*"_TYPE"*) [ty] = mk_type ty | 
| 4148 | 161 |   | type_tr (*"_TYPE"*) ts = raise TERM ("type_tr", ts);
 | 
| 162 | ||
| 548 | 163 | |
| 6761 | 164 | (* dddot *) | 
| 165 | ||
| 8595 | 166 | fun dddot_tr (*"_DDDOT"*) ts = Term.list_comb (Lexicon.var SynExt.dddot_indexname, ts); | 
| 6761 | 167 | |
| 168 | ||
| 5084 | 169 | (* quote / antiquote *) | 
| 170 | ||
| 8577 | 171 | fun antiquote_tr name = | 
| 172 | let | |
| 173 | fun tr i ((t as Const (c, _)) $ u) = | |
| 174 | if c = name then tr i u $ Bound i | |
| 175 | else tr i t $ tr i u | |
| 176 | | tr i (t $ u) = tr i t $ tr i u | |
| 177 | | tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t) | |
| 178 | | tr _ a = a; | |
| 179 | in tr 0 end; | |
| 180 | ||
| 181 | fun quote_tr name t = Abs ("s", dummyT, antiquote_tr name (Term.incr_boundvars 1 t));
 | |
| 182 | ||
| 5084 | 183 | fun quote_antiquote_tr quoteN antiquoteN name = | 
| 184 | let | |
| 8577 | 185 | fun tr [t] = Lexicon.const name $ quote_tr antiquoteN t | 
| 186 |       | tr ts = raise TERM ("quote_tr", ts);
 | |
| 187 | in (quoteN, tr) end; | |
| 5084 | 188 | |
| 189 | ||
| 14697 | 190 | (* indexed syntax *) | 
| 191 | ||
| 192 | fun struct_ast_tr (*"_struct"*) [Ast.Appl [Ast.Constant "_index", ast]] = ast | |
| 193 | | struct_ast_tr (*"_struct"*) asts = Ast.mk_appl (Ast.Constant "_struct") asts; | |
| 194 | ||
| 195 | fun index_ast_tr ast = | |
| 196 | Ast.mk_appl (Ast.Constant "_index") [Ast.mk_appl (Ast.Constant "_struct") [ast]]; | |
| 197 | ||
| 198 | fun indexdefault_ast_tr (*"_indexdefault"*) [] = | |
| 199 | index_ast_tr (Ast.Constant "_indexdefault") | |
| 200 | | indexdefault_ast_tr (*"_indexdefault"*) asts = | |
| 201 |       raise Ast.AST ("indexdefault_ast_tr", asts);
 | |
| 202 | ||
| 203 | fun indexnum_ast_tr (*"_indexnum"*) [ast] = | |
| 204 | index_ast_tr (Ast.mk_appl (Ast.Constant "_indexnum") [ast]) | |
| 205 |   | indexnum_ast_tr (*"_indexnum"*) asts = raise Ast.AST ("indexnum_ast_tr", asts);
 | |
| 12122 | 206 | |
| 14697 | 207 | fun indexvar_ast_tr (*"_indexvar"*) [] = | 
| 208 | Ast.mk_appl (Ast.Constant "_index") [Ast.Variable "some_index"] | |
| 209 |   | indexvar_ast_tr (*"_indexvar"*) asts = raise Ast.AST ("indexvar_ast_tr", asts);
 | |
| 210 | ||
| 211 | fun index_tr (*"_index"*) [t] = t | |
| 212 |   | index_tr (*"_index"*) ts = raise TERM ("index_tr", ts);
 | |
| 213 | ||
| 214 | ||
| 215 | (* implicit structures *) | |
| 216 | ||
| 217 | fun the_struct structs i = | |
| 30146 | 218 | if 1 <= i andalso i <= length structs then nth structs (i - 1) | 
| 31542 
3371a3c19bb1
reraise exceptions to preserve position information;
 wenzelm parents: 
30146diff
changeset | 219 |   else error ("Illegal reference to implicit structure #" ^ string_of_int i);
 | 
| 14697 | 220 | |
| 221 | fun struct_tr structs (*"_struct"*) [Const ("_indexdefault", _)] =
 | |
| 222 | Lexicon.free (the_struct structs 1) | |
| 223 |   | struct_tr structs (*"_struct"*) [t as (Const ("_indexnum", _) $ Const (s, _))] =
 | |
| 224 | Lexicon.free (the_struct structs | |
| 15531 | 225 |         (case Lexicon.read_nat s of SOME n => n | NONE => raise TERM ("struct_tr", [t])))
 | 
| 14697 | 226 |   | struct_tr _ (*"_struct"*) ts = raise TERM ("struct_tr", ts);
 | 
| 12122 | 227 | |
| 228 | ||
| 5084 | 229 | |
| 548 | 230 | (** print (ast) translations **) | 
| 231 | ||
| 14647 | 232 | (* types *) | 
| 233 | ||
| 234 | fun non_typed_tr' f _ _ ts = f ts; | |
| 235 | fun non_typed_tr'' f x _ _ ts = f x ts; | |
| 236 | ||
| 237 | ||
| 548 | 238 | (* application *) | 
| 239 | ||
| 5690 | 240 | fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f])
 | 
| 241 | | appl_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_appl", f, Ast.fold_ast "_args" args]; | |
| 548 | 242 | |
| 5690 | 243 | fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f])
 | 
| 244 | | 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 | 245 | |
| 548 | 246 | |
| 247 | (* abstraction *) | |
| 248 | ||
| 19311 | 249 | fun mark_boundT (x, T) = Const ("_bound", T --> T) $ Free (x, T);
 | 
| 2698 | 250 | fun mark_bound x = mark_boundT (x, dummyT); | 
| 251 | ||
| 18958 | 252 | fun bound_vars vars body = | 
| 253 | subst_bounds (map mark_boundT (Term.rename_wrt_term body vars), body); | |
| 254 | ||
| 548 | 255 | fun strip_abss vars_of body_of tm = | 
| 256 | let | |
| 257 | val vars = vars_of tm; | |
| 258 | val body = body_of tm; | |
| 29276 | 259 | val rev_new_vars = Term.rename_wrt_term body vars; | 
| 21750 | 260 | fun subst (x, T) b = | 
| 261 | if can Name.dest_internal x andalso not (Term.loose_bvar1 (b, 0)) | |
| 262 |       then (Const ("_idtdummy", T), incr_boundvars ~1 b)
 | |
| 263 | else (mark_boundT (x, T), Term.subst_bound (mark_bound x, b)); | |
| 264 | val (rev_vars', body') = fold_map subst rev_new_vars body; | |
| 265 | in (rev rev_vars', body') end; | |
| 548 | 266 | |
| 3928 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 267 | |
| 548 | 268 | (*do (partial) eta-contraction before printing*) | 
| 269 | ||
| 32738 | 270 | val eta_contract = Unsynchronized.ref true; | 
| 548 | 271 | |
| 272 | fun eta_contr tm = | |
| 273 | let | |
| 3928 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 274 |     fun is_aprop (Const ("_aprop", _)) = true
 | 
| 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 275 | | is_aprop _ = false; | 
| 
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
 wenzelm parents: 
3777diff
changeset | 276 | |
| 548 | 277 | fun eta_abs (Abs (a, T, t)) = | 
| 278 | (case eta_abs t of | |
| 279 | t' as f $ u => | |
| 280 | (case eta_abs u of | |
| 281 | Bound 0 => | |
| 5084 | 282 | 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 | 283 | else incr_boundvars ~1 f | 
| 548 | 284 | | _ => Abs (a, T, t')) | 
| 285 | | t' => Abs (a, T, t')) | |
| 286 | | eta_abs t = t; | |
| 287 | in | |
| 288 | if ! eta_contract then eta_abs tm else tm | |
| 289 | end; | |
| 290 | ||
| 291 | ||
| 292 | fun abs_tr' tm = | |
| 19473 | 293 | uncurry (fold_rev (fn x => fn t => Lexicon.const "_abs" $ x $ t)) | 
| 548 | 294 | (strip_abss strip_abs_vars strip_abs_body (eta_contr tm)); | 
| 295 | ||
| 14697 | 296 | fun atomic_abs_tr' (x, T, t) = | 
| 29276 | 297 | let val [xT] = Term.rename_wrt_term t [(x, T)] | 
| 14697 | 298 | in (mark_boundT xT, subst_bound (mark_bound (fst xT), t)) end; | 
| 13762 | 299 | |
| 548 | 300 | fun abs_ast_tr' (*"_abs"*) asts = | 
| 5690 | 301 | (case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of | 
| 302 |     ([], _) => raise Ast.AST ("abs_ast_tr'", asts)
 | |
| 303 | | (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]); | |
| 548 | 304 | |
| 32120 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 305 | fun preserve_binder_abs_tr' name syn = (name, fn (Abs abs :: ts) => | 
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 306 | let val (x, t) = atomic_abs_tr' abs | 
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 307 | in list_comb (Lexicon.const syn $ x $ t, ts) end); | 
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 308 | |
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 309 | fun preserve_binder_abs2_tr' name syn = (name, fn (A :: Abs abs :: ts) => | 
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 310 | let val (x, t) = atomic_abs_tr' abs | 
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 311 | in list_comb (Lexicon.const syn $ x $ A $ t, ts) end); | 
| 
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
 haftmann parents: 
31542diff
changeset | 312 | |
| 548 | 313 | |
| 314 | (* binder *) | |
| 315 | ||
| 21535 | 316 | fun mk_binder_tr' (name, syn) = | 
| 548 | 317 | let | 
| 318 | fun mk_idts [] = raise Match (*abort translation*) | |
| 319 | | mk_idts [idt] = idt | |
| 5690 | 320 | | mk_idts (idt :: idts) = Lexicon.const "_idts" $ idt $ mk_idts idts; | 
| 548 | 321 | |
| 322 | fun tr' t = | |
| 323 | let | |
| 324 | val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t; | |
| 21535 | 325 | in Lexicon.const syn $ mk_idts xs $ bd end; | 
| 548 | 326 | |
| 21535 | 327 | fun binder_tr' (t :: ts) = Term.list_comb (tr' (Lexicon.const name $ t), ts) | 
| 328 | | binder_tr' [] = raise Match; | |
| 329 | in (name, binder_tr') end; | |
| 548 | 330 | |
| 331 | ||
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 332 | (* idtyp constraints *) | 
| 548 | 333 | |
| 5690 | 334 | fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant c, x, ty], xs] = | 
| 17788 | 335 | if c = "_constrain" then | 
| 28628 | 336 | Ast.Appl [Ast.Constant a, Ast.Appl [Ast.Constant "_idtyp", x, ty], xs] | 
| 548 | 337 | else raise Match | 
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 338 | | idtyp_ast_tr' _ _ = raise Match; | 
| 548 | 339 | |
| 340 | ||
| 28628 | 341 | (* type propositions *) | 
| 342 | ||
| 343 | fun type_prop_tr' _ (*"_type_prop"*) T [Const ("Pure.sort_constraint", _)] =
 | |
| 344 | Lexicon.const "_sort_constraint" $ TypeExt.term_of_typ true T | |
| 345 | | type_prop_tr' show_sorts (*"_type_prop"*) T [t] = | |
| 346 | Lexicon.const "_ofclass" $ TypeExt.term_of_typ show_sorts T $ t | |
| 347 |   | type_prop_tr' _ (*"_type_prop"*) T ts = raise TYPE ("type_prop_tr'", [T], ts);
 | |
| 348 | ||
| 349 | ||
| 548 | 350 | (* meta propositions *) | 
| 351 | ||
| 4148 | 352 | fun prop_tr' tm = | 
| 548 | 353 | let | 
| 5690 | 354 | fun aprop t = Lexicon.const "_aprop" $ t; | 
| 548 | 355 | |
| 2698 | 356 | fun is_prop Ts t = | 
| 357 | fastype_of1 (Ts, t) = propT handle TERM _ => false; | |
| 548 | 358 | |
| 26424 | 359 |     fun is_term (Const ("Pure.term", _) $ _) = true
 | 
| 19848 | 360 | | is_term _ = false; | 
| 361 | ||
| 548 | 362 | fun tr' _ (t as Const _) = t | 
| 18478 
29a5070b517c
prop_tr': proper handling of aprop marked as bound;
 wenzelm parents: 
18342diff
changeset | 363 |       | tr' Ts (t as Const ("_bound", _) $ u) =
 | 
| 
29a5070b517c
prop_tr': proper handling of aprop marked as bound;
 wenzelm parents: 
18342diff
changeset | 364 | if is_prop Ts u then aprop t else t | 
| 2698 | 365 | | tr' _ (t as Free (x, T)) = | 
| 5690 | 366 | if T = propT then aprop (Lexicon.free x) else t | 
| 2698 | 367 | | tr' _ (t as Var (xi, T)) = | 
| 5690 | 368 | if T = propT then aprop (Lexicon.var xi) else t | 
| 2698 | 369 | | tr' Ts (t as Bound _) = | 
| 370 | if is_prop Ts t then aprop t else t | |
| 371 | | tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t) | |
| 372 |       | tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
 | |
| 28628 | 373 |           if is_prop Ts t andalso not (is_term t) then Const ("_type_prop", T) $ tr' Ts t1
 | 
| 2698 | 374 | else tr' Ts t1 $ tr' Ts t2 | 
| 375 | | tr' Ts (t as t1 $ t2) = | |
| 5690 | 376 | (if is_Const (Term.head_of t) orelse not (is_prop Ts t) | 
| 2698 | 377 | then I else aprop) (tr' Ts t1 $ tr' Ts t2); | 
| 16612 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 378 | in tr' [] tm end; | 
| 548 | 379 | |
| 2698 | 380 | |
| 548 | 381 | (* meta implication *) | 
| 382 | ||
| 383 | fun impl_ast_tr' (*"==>"*) asts = | |
| 10572 | 384 | if TypeExt.no_brackets () then raise Match | 
| 385 | else | |
| 386 | (case Ast.unfold_ast_p "==>" (Ast.Appl (Ast.Constant "==>" :: asts)) of | |
| 16612 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 387 | (prems as _ :: _ :: _, concl) => | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 388 | let | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 389 | val (asms, asm) = split_last prems; | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 390 | val asms' = Ast.fold_ast_p "_asms" (asms, Ast.Appl [Ast.Constant "_asm", asm]); | 
| 
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
 wenzelm parents: 
15570diff
changeset | 391 | in Ast.Appl [Ast.Constant "_bigimpl", asms', concl] end | 
| 15421 
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
 schirmer parents: 
14981diff
changeset | 392 | | _ => raise Match); | 
| 
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
 schirmer parents: 
14981diff
changeset | 393 | |
| 548 | 394 | |
| 4148 | 395 | (* type reflection *) | 
| 396 | ||
| 397 | fun type_tr' show_sorts (*"TYPE"*) (Type ("itself", [T])) ts =
 | |
| 5690 | 398 | Term.list_comb (Lexicon.const "_TYPE" $ TypeExt.term_of_typ show_sorts T, ts) | 
| 4148 | 399 | | type_tr' _ _ _ = raise Match; | 
| 400 | ||
| 401 | ||
| 19577 | 402 | (* type constraints *) | 
| 403 | ||
| 404 | fun type_constraint_tr' show_sorts (*"_type_constraint_"*) (Type ("fun", [T, _])) (t :: ts) =
 | |
| 405 | Term.list_comb (Lexicon.const SynExt.constrainC $ t $ TypeExt.term_of_typ show_sorts T, ts) | |
| 406 | | type_constraint_tr' _ _ _ = raise Match; | |
| 407 | ||
| 408 | ||
| 548 | 409 | (* dependent / nondependent quantifiers *) | 
| 410 | ||
| 20202 | 411 | fun var_abs mark (x, T, b) = | 
| 412 | let val ([x'], _) = Name.variants [x] (Term.declare_term_names b Name.context) | |
| 413 | in (x', subst_bound (mark (x', T), b)) end; | |
| 414 | ||
| 415 | val variant_abs = var_abs Free; | |
| 416 | val variant_abs' = var_abs mark_boundT; | |
| 2698 | 417 | |
| 548 | 418 | fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) = | 
| 5084 | 419 | if Term.loose_bvar1 (B, 0) then | 
| 2698 | 420 | let val (x', B') = variant_abs' (x, dummyT, B); | 
| 5690 | 421 | in Term.list_comb (Lexicon.const q $ mark_boundT (x', T) $ A $ B', ts) end | 
| 422 | else Term.list_comb (Lexicon.const r $ A $ B, ts) | |
| 548 | 423 | | dependent_tr' _ _ = raise Match; | 
| 424 | ||
| 425 | ||
| 5084 | 426 | (* quote / antiquote *) | 
| 427 | ||
| 8577 | 428 | fun antiquote_tr' name = | 
| 429 | let | |
| 430 | fun tr' i (t $ u) = | |
| 18139 | 431 | if u aconv Bound i then Lexicon.const name $ tr' i t | 
| 8577 | 432 | else tr' i t $ tr' i u | 
| 433 | | tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t) | |
| 18139 | 434 | | tr' i a = if a aconv Bound i then raise Match else a; | 
| 8577 | 435 | in tr' 0 end; | 
| 436 | ||
| 437 | fun quote_tr' name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' name t) | |
| 438 | | quote_tr' _ _ = raise Match; | |
| 439 | ||
| 5084 | 440 | fun quote_antiquote_tr' quoteN antiquoteN name = | 
| 441 | let | |
| 8577 | 442 | fun tr' (t :: ts) = Term.list_comb (Lexicon.const quoteN $ quote_tr' antiquoteN t, ts) | 
| 443 | | tr' _ = raise Match; | |
| 444 | in (name, tr') end; | |
| 5084 | 445 | |
| 446 | ||
| 14697 | 447 | (* indexed syntax *) | 
| 548 | 448 | |
| 14697 | 449 | fun index_ast_tr' (*"_index"*) [Ast.Appl [Ast.Constant "_struct", ast]] = ast | 
| 450 | | index_ast_tr' _ = raise Match; | |
| 451 | ||
| 452 | ||
| 453 | (* implicit structures *) | |
| 454 | ||
| 455 | fun the_struct' structs s = | |
| 456 | [(case Lexicon.read_nat s of | |
| 18678 | 457 | SOME i => Ast.Variable (the_struct structs i handle ERROR _ => raise Match) | 
| 15531 | 458 | | NONE => raise Match)] |> Ast.mk_appl (Ast.Constant "_free"); | 
| 14697 | 459 | |
| 460 | fun struct_ast_tr' structs (*"_struct"*) [Ast.Constant "_indexdefault"] = | |
| 461 | the_struct' structs "1" | |
| 462 | | struct_ast_tr' structs (*"_struct"*) [Ast.Appl [Ast.Constant "_indexnum", Ast.Constant s]] = | |
| 463 | the_struct' structs s | |
| 464 | | struct_ast_tr' _ _ = raise Match; | |
| 465 | ||
| 466 | ||
| 467 | ||
| 468 | (** Pure translations **) | |
| 548 | 469 | |
| 470 | val pure_trfuns = | |
| 11491 | 471 |  ([("_constify", constify_ast_tr), ("_appl", appl_ast_tr), ("_applC", applC_ast_tr),
 | 
| 17788 | 472 |    ("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr), ("_idtypdummy", idtypdummy_ast_tr),
 | 
| 473 |    ("_bigimpl", bigimpl_ast_tr), ("_indexdefault", indexdefault_ast_tr),
 | |
| 474 |    ("_indexnum", indexnum_ast_tr), ("_indexvar", indexvar_ast_tr), ("_struct", struct_ast_tr)],
 | |
| 922 
196ca0973a6d
added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
 clasohm parents: 
639diff
changeset | 475 |   [("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr),
 | 
| 28856 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28628diff
changeset | 476 |    ("_sort_constraint", sort_constraint_tr), ("_TYPE", type_tr),
 | 
| 
5e009a80fe6d
Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
 wenzelm parents: 
28628diff
changeset | 477 |    ("_DDDOT", dddot_tr), ("_index", index_tr)],
 | 
| 19131 | 478 | ([]: (string * (term list -> term)) list), | 
| 3691 
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
 wenzelm parents: 
2698diff
changeset | 479 |   [("_abs", abs_ast_tr'), ("_idts", idtyp_ast_tr' "_idts"),
 | 
| 14697 | 480 |    ("_pttrns", idtyp_ast_tr' "_pttrns"), ("==>", impl_ast_tr'),
 | 
| 481 |    ("_index", index_ast_tr')]);
 | |
| 548 | 482 | |
| 2698 | 483 | val pure_trfunsT = | 
| 28628 | 484 |   [("_type_prop", type_prop_tr'), ("TYPE", type_tr'), ("_type_constraint_", type_constraint_tr')];
 | 
| 2698 | 485 | |
| 14697 | 486 | fun struct_trfuns structs = | 
| 487 |   ([], [("_struct", struct_tr structs)], [], [("_struct", struct_ast_tr' structs)]);
 | |
| 488 | ||
| 548 | 489 | |
| 490 | ||
| 14868 | 491 | (** pts_to_asts **) | 
| 548 | 492 | |
| 21773 | 493 | fun pts_to_asts ctxt trf pts = | 
| 548 | 494 | let | 
| 495 | fun trans a args = | |
| 496 | (case trf a of | |
| 15531 | 497 | NONE => Ast.mk_appl (Ast.Constant a) args | 
| 23937 
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
 wenzelm parents: 
23824diff
changeset | 498 | | SOME f => f ctxt args); | 
| 548 | 499 | |
| 987 | 500 | (*translate pt bottom-up*) | 
| 5690 | 501 | fun ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts) | 
| 502 | | ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok); | |
| 14798 
702cb4859cab
Modified functions pt_to_ast and ast_to_term to improve handling
 berghofe parents: 
14697diff
changeset | 503 | |
| 23963 
c2ee97a963db
moved exception capture/release to structure Exn;
 wenzelm parents: 
23937diff
changeset | 504 | val exn_results = map (Exn.capture ast_of) pts; | 
| 
c2ee97a963db
moved exception capture/release to structure Exn;
 wenzelm parents: 
23937diff
changeset | 505 | val exns = map_filter Exn.get_exn exn_results; | 
| 
c2ee97a963db
moved exception capture/release to structure Exn;
 wenzelm parents: 
23937diff
changeset | 506 | val results = map_filter Exn.get_result exn_results | 
| 31542 
3371a3c19bb1
reraise exceptions to preserve position information;
 wenzelm parents: 
30146diff
changeset | 507 | in (case (results, exns) of ([], exn :: _) => reraise exn | _ => results) end; | 
| 548 | 508 | |
| 509 | ||
| 510 | ||
| 14798 
702cb4859cab
Modified functions pt_to_ast and ast_to_term to improve handling
 berghofe parents: 
14697diff
changeset | 511 | (** asts_to_terms **) | 
| 548 | 512 | |
| 21773 | 513 | fun asts_to_terms ctxt trf asts = | 
| 548 | 514 | let | 
| 515 | fun trans a args = | |
| 516 | (case trf a of | |
| 15531 | 517 | NONE => Term.list_comb (Lexicon.const a, args) | 
| 23937 
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
 wenzelm parents: 
23824diff
changeset | 518 | | SOME f => f ctxt args); | 
| 548 | 519 | |
| 5690 | 520 | fun term_of (Ast.Constant a) = trans a [] | 
| 521 | | term_of (Ast.Variable x) = Lexicon.read_var x | |
| 522 | | term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) = | |
| 548 | 523 | trans a (map term_of asts) | 
| 5690 | 524 | | term_of (Ast.Appl (ast :: (asts as _ :: _))) = | 
| 525 | Term.list_comb (term_of ast, map term_of asts) | |
| 526 |       | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]);
 | |
| 14798 
702cb4859cab
Modified functions pt_to_ast and ast_to_term to improve handling
 berghofe parents: 
14697diff
changeset | 527 | |
| 19005 | 528 | val free_fixed = Term.map_aterms | 
| 529 | (fn t as Const (c, T) => | |
| 530 | (case try (unprefix Lexicon.fixedN) c of | |
| 531 | NONE => t | |
| 18342 
1b7109c10b7b
asts_to_terms: builtin free_fixed translation makes local binders work properly;
 wenzelm parents: 
18212diff
changeset | 532 | | SOME x => Free (x, T)) | 
| 19005 | 533 | | t => t); | 
| 18342 
1b7109c10b7b
asts_to_terms: builtin free_fixed translation makes local binders work properly;
 wenzelm parents: 
18212diff
changeset | 534 | |
| 23963 
c2ee97a963db
moved exception capture/release to structure Exn;
 wenzelm parents: 
23937diff
changeset | 535 | val exn_results = map (Exn.capture (term_of #> free_fixed)) asts; | 
| 
c2ee97a963db
moved exception capture/release to structure Exn;
 wenzelm parents: 
23937diff
changeset | 536 | val exns = map_filter Exn.get_exn exn_results; | 
| 
c2ee97a963db
moved exception capture/release to structure Exn;
 wenzelm parents: 
23937diff
changeset | 537 | val results = map_filter Exn.get_result exn_results | 
| 31542 
3371a3c19bb1
reraise exceptions to preserve position information;
 wenzelm parents: 
30146diff
changeset | 538 | in (case (results, exns) of ([], exn :: _) => reraise exn | _ => results) end; | 
| 548 | 539 | |
| 540 | end; |