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