author | wenzelm |
Tue, 05 Apr 2011 14:25:18 +0200 | |
changeset 42224 | 578a51fae383 |
parent 42223 | 098c86e53153 |
child 42242 | 39261908e12f |
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 |
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
37216
diff
changeset
|
9 |
val eta_contract_default: bool Unsynchronized.ref |
39163
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
wenzelm
parents:
39128
diff
changeset
|
10 |
val eta_contract_raw: Config.raw |
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
37216
diff
changeset
|
11 |
val eta_contract: bool Config.T |
13762 | 12 |
val atomic_abs_tr': string * typ * term -> term * term |
42086
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
13 |
val const_abs_tr': term -> term |
42085 | 14 |
val mk_binder_tr: string * string -> string * (term list -> term) |
15 |
val mk_binder_tr': string * string -> string * (term list -> term) |
|
32120
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
haftmann
parents:
31542
diff
changeset
|
16 |
val preserve_binder_abs_tr': string -> string -> string * (term list -> term) |
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
haftmann
parents:
31542
diff
changeset
|
17 |
val preserve_binder_abs2_tr': string -> string -> string * (term list -> term) |
548 | 18 |
val dependent_tr': 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) |
8577 | 22 |
val antiquote_tr': string -> term -> term |
23 |
val quote_tr': string -> term -> term |
|
5084 | 24 |
val quote_antiquote_tr': string -> string -> string -> string * (term list -> term) |
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
25 |
val update_name_tr': term -> term |
2698 | 26 |
val mark_bound: string -> term |
27 |
val mark_boundT: string * typ -> term |
|
18958 | 28 |
val bound_vars: (string * typ) list -> term -> term |
20202 | 29 |
val variant_abs: string * typ * term -> string * term |
2698 | 30 |
val variant_abs': string * typ * term -> string * term |
31 |
end; |
|
548 | 32 |
|
33 |
signature SYN_TRANS1 = |
|
2698 | 34 |
sig |
548 | 35 |
include SYN_TRANS0 |
14647 | 36 |
val non_typed_tr': (term list -> term) -> bool -> typ -> term list -> term |
37 |
val non_typed_tr'': ('a -> term list -> term) -> 'a -> bool -> typ -> term list -> term |
|
1511 | 38 |
val constrainAbsC: string |
42152
6c17259724b2
Hoare syntax: standard abstraction syntax admits source positions;
wenzelm
parents:
42086
diff
changeset
|
39 |
val abs_tr: term list -> term |
1511 | 40 |
val pure_trfuns: |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
41 |
(string * (Ast.ast list -> Ast.ast)) list * |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
42 |
(string * (term list -> term)) list * |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
43 |
(string * (term list -> term)) list * |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
44 |
(string * (Ast.ast list -> Ast.ast)) list |
4148 | 45 |
val pure_trfunsT: (string * (bool -> typ -> term list -> term)) list |
14697 | 46 |
val struct_trfuns: string list -> |
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
47 |
(string * (Ast.ast list -> Ast.ast)) list * |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
48 |
(string * (term list -> term)) list * |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
49 |
(string * (bool -> typ -> term list -> term)) list * |
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
50 |
(string * (Ast.ast list -> Ast.ast)) list |
42204 | 51 |
type type_context |
2698 | 52 |
end; |
548 | 53 |
|
54 |
signature SYN_TRANS = |
|
2698 | 55 |
sig |
548 | 56 |
include SYN_TRANS1 |
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
37216
diff
changeset
|
57 |
val abs_tr': Proof.context -> term -> term |
4148 | 58 |
val prop_tr': term -> term |
1511 | 59 |
val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
60 |
val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
|
42204 | 61 |
val parsetree_to_ast: Proof.context -> type_context -> bool -> |
62 |
(string -> (Proof.context -> Ast.ast list -> Ast.ast) option) -> |
|
42223
098c86e53153
more precise propagation of reports/results through some inner syntax layers;
wenzelm
parents:
42204
diff
changeset
|
63 |
Parser.parsetree -> Position.reports * Ast.ast Exn.result |
42043
2055515c9d3f
clarified Syn_Trans.parsetree_to_ast and Syn_Trans.ast_to_term;
wenzelm
parents:
40235
diff
changeset
|
64 |
val ast_to_term: Proof.context -> |
2055515c9d3f
clarified Syn_Trans.parsetree_to_ast and Syn_Trans.ast_to_term;
wenzelm
parents:
40235
diff
changeset
|
65 |
(string -> (Proof.context -> term list -> term) option) -> Ast.ast -> term |
2698 | 66 |
end; |
548 | 67 |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
36502
diff
changeset
|
68 |
structure Syn_Trans: SYN_TRANS = |
548 | 69 |
struct |
2698 | 70 |
|
71 |
||
548 | 72 |
(** parse (ast) translations **) |
73 |
||
42057
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
74 |
(* strip_positions *) |
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
75 |
|
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
76 |
fun strip_positions_ast_tr [ast] = Type_Ext.strip_positions_ast ast |
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
77 |
| strip_positions_ast_tr asts = raise Ast.AST ("strip_positions_ast_tr", asts); |
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
78 |
|
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
79 |
|
11491 | 80 |
(* constify *) |
81 |
||
82 |
fun constify_ast_tr [Ast.Variable c] = Ast.Constant c |
|
83 |
| constify_ast_tr asts = raise Ast.AST ("constify_ast_tr", asts); |
|
84 |
||
85 |
||
548 | 86 |
(* application *) |
87 |
||
5690 | 88 |
fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args) |
89 |
| 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
|
90 |
|
5690 | 91 |
fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args) |
92 |
| applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts); |
|
548 | 93 |
|
94 |
||
95 |
(* abstraction *) |
|
96 |
||
17788 | 97 |
fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Ast.Appl [Ast.Constant "_constrain", x, ty] |
5690 | 98 |
| idtyp_ast_tr (*"_idtyp"*) asts = raise Ast.AST ("idtyp_ast_tr", asts); |
548 | 99 |
|
17788 | 100 |
fun idtypdummy_ast_tr (*"_idtypdummy"*) [ty] = |
101 |
Ast.Appl [Ast.Constant "_constrain", Ast.Constant "_idtdummy", ty] |
|
102 |
| idtypdummy_ast_tr (*"_idtypdummy"*) asts = raise Ast.AST ("idtyp_ast_tr", asts); |
|
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
103 |
fun lambda_ast_tr (*"_lambda"*) [pats, body] = |
5690 | 104 |
Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body) |
105 |
| lambda_ast_tr (*"_lambda"*) asts = raise Ast.AST ("lambda_ast_tr", asts); |
|
548 | 106 |
|
107 |
val constrainAbsC = "_constrainAbs"; |
|
108 |
||
21773 | 109 |
fun absfree_proper (x, T, t) = |
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
110 |
if can Name.dest_internal x |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
111 |
then error ("Illegal internal variable in abstraction: " ^ quote x) |
21773 | 112 |
else Term.absfree (x, T, t); |
113 |
||
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
114 |
fun abs_tr [Free (x, T), t] = absfree_proper (x, T, t) |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
115 |
| abs_tr [Const ("_idtdummy", T), t] = Term.absdummy (T, t) |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
116 |
| abs_tr [Const ("_constrain", _) $ x $ tT, t] = Lexicon.const constrainAbsC $ abs_tr [x, t] $ tT |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
117 |
| abs_tr ts = raise TERM ("abs_tr", ts); |
548 | 118 |
|
119 |
||
120 |
(* binder *) |
|
121 |
||
21535 | 122 |
fun mk_binder_tr (syn, name) = |
548 | 123 |
let |
42055 | 124 |
fun err ts = raise TERM ("binder_tr: " ^ syn, ts) |
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
125 |
fun binder_tr [Const ("_idts", _) $ idt $ idts, t] = binder_tr [idt, binder_tr [idts, t]] |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
126 |
| binder_tr [x, t] = |
42055 | 127 |
let val abs = abs_tr [x, t] handle TERM _ => err [x, t] |
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
128 |
in Lexicon.const name $ abs end |
42055 | 129 |
| binder_tr ts = err ts; |
21535 | 130 |
in (syn, binder_tr) end; |
548 | 131 |
|
132 |
||
28628 | 133 |
(* type propositions *) |
134 |
||
35255 | 135 |
fun mk_type ty = |
136 |
Lexicon.const "_constrain" $ |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
137 |
Lexicon.const "\\<^const>TYPE" $ (Lexicon.const "\\<^type>itself" $ ty); |
28628 | 138 |
|
139 |
fun ofclass_tr (*"_ofclass"*) [ty, cls] = cls $ mk_type ty |
|
140 |
| ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts); |
|
141 |
||
142 |
fun sort_constraint_tr (*"_sort_constraint"*) [ty] = |
|
35255 | 143 |
Lexicon.const "\\<^const>Pure.sort_constraint" $ mk_type ty |
28628 | 144 |
| sort_constraint_tr (*"_sort_constraint"*) ts = raise TERM ("sort_constraint_tr", ts); |
145 |
||
146 |
||
548 | 147 |
(* meta propositions *) |
148 |
||
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
149 |
fun aprop_tr (*"_aprop"*) [t] = Lexicon.const "_constrain" $ t $ Lexicon.const "\\<^type>prop" |
3777 | 150 |
| aprop_tr (*"_aprop"*) ts = raise TERM ("aprop_tr", ts); |
548 | 151 |
|
152 |
||
153 |
(* meta implication *) |
|
154 |
||
16612
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
155 |
fun bigimpl_ast_tr (*"_bigimpl"*) (asts as [asms, concl]) = |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
156 |
let val prems = |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
157 |
(case Ast.unfold_ast_p "_asms" asms of |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
158 |
(asms', Ast.Appl [Ast.Constant "_asm", asm']) => asms' @ [asm'] |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
159 |
| _ => raise Ast.AST ("bigimpl_ast_tr", asts)) |
35255 | 160 |
in Ast.fold_ast_p "\\<^const>==>" (prems, concl) end |
15421
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
schirmer
parents:
14981
diff
changeset
|
161 |
| 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
|
162 |
|
548 | 163 |
|
23824
8ad7131dbfcf
moved print_translations from Pure.thy to Syntax/syn_trans.ML;
wenzelm
parents:
21773
diff
changeset
|
164 |
(* type/term reflection *) |
4148 | 165 |
|
28628 | 166 |
fun type_tr (*"_TYPE"*) [ty] = mk_type ty |
4148 | 167 |
| type_tr (*"_TYPE"*) ts = raise TERM ("type_tr", ts); |
168 |
||
548 | 169 |
|
6761 | 170 |
(* dddot *) |
171 |
||
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
36502
diff
changeset
|
172 |
fun dddot_tr (*"_DDDOT"*) ts = Term.list_comb (Lexicon.var Syn_Ext.dddot_indexname, ts); |
6761 | 173 |
|
174 |
||
5084 | 175 |
(* quote / antiquote *) |
176 |
||
8577 | 177 |
fun antiquote_tr name = |
178 |
let |
|
179 |
fun tr i ((t as Const (c, _)) $ u) = |
|
180 |
if c = name then tr i u $ Bound i |
|
181 |
else tr i t $ tr i u |
|
182 |
| tr i (t $ u) = tr i t $ tr i u |
|
183 |
| tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t) |
|
184 |
| tr _ a = a; |
|
185 |
in tr 0 end; |
|
186 |
||
187 |
fun quote_tr name t = Abs ("s", dummyT, antiquote_tr name (Term.incr_boundvars 1 t)); |
|
188 |
||
5084 | 189 |
fun quote_antiquote_tr quoteN antiquoteN name = |
190 |
let |
|
8577 | 191 |
fun tr [t] = Lexicon.const name $ quote_tr antiquoteN t |
192 |
| tr ts = raise TERM ("quote_tr", ts); |
|
193 |
in (quoteN, tr) end; |
|
5084 | 194 |
|
195 |
||
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
196 |
(* corresponding updates *) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
197 |
|
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
198 |
fun update_name_tr (Free (x, T) :: ts) = list_comb (Free (suffix "_update" x, T), ts) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
199 |
| update_name_tr (Const (x, T) :: ts) = list_comb (Const (suffix "_update" x, T), ts) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
200 |
| update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) = |
42053
006095137a81
update_name_tr: more precise handling of explicit constraints, including positions;
wenzelm
parents:
42048
diff
changeset
|
201 |
if Type_Ext.is_position ty then list_comb (c $ update_name_tr [t] $ ty, ts) |
006095137a81
update_name_tr: more precise handling of explicit constraints, including positions;
wenzelm
parents:
42048
diff
changeset
|
202 |
else |
006095137a81
update_name_tr: more precise handling of explicit constraints, including positions;
wenzelm
parents:
42048
diff
changeset
|
203 |
list_comb (c $ update_name_tr [t] $ |
42080
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
204 |
(Lexicon.fun_type $ |
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
205 |
(Lexicon.fun_type $ Lexicon.dummy_type $ ty) $ Lexicon.dummy_type), ts) |
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
206 |
| update_name_tr ts = raise TERM ("update_name_tr", ts); |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
207 |
|
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
208 |
|
14697 | 209 |
(* indexed syntax *) |
210 |
||
211 |
fun struct_ast_tr (*"_struct"*) [Ast.Appl [Ast.Constant "_index", ast]] = ast |
|
212 |
| struct_ast_tr (*"_struct"*) asts = Ast.mk_appl (Ast.Constant "_struct") asts; |
|
213 |
||
214 |
fun index_ast_tr ast = |
|
215 |
Ast.mk_appl (Ast.Constant "_index") [Ast.mk_appl (Ast.Constant "_struct") [ast]]; |
|
216 |
||
217 |
fun indexdefault_ast_tr (*"_indexdefault"*) [] = |
|
218 |
index_ast_tr (Ast.Constant "_indexdefault") |
|
219 |
| indexdefault_ast_tr (*"_indexdefault"*) asts = |
|
220 |
raise Ast.AST ("indexdefault_ast_tr", asts); |
|
221 |
||
222 |
fun indexnum_ast_tr (*"_indexnum"*) [ast] = |
|
223 |
index_ast_tr (Ast.mk_appl (Ast.Constant "_indexnum") [ast]) |
|
224 |
| indexnum_ast_tr (*"_indexnum"*) asts = raise Ast.AST ("indexnum_ast_tr", asts); |
|
12122 | 225 |
|
14697 | 226 |
fun indexvar_ast_tr (*"_indexvar"*) [] = |
227 |
Ast.mk_appl (Ast.Constant "_index") [Ast.Variable "some_index"] |
|
228 |
| indexvar_ast_tr (*"_indexvar"*) asts = raise Ast.AST ("indexvar_ast_tr", asts); |
|
229 |
||
230 |
fun index_tr (*"_index"*) [t] = t |
|
231 |
| index_tr (*"_index"*) ts = raise TERM ("index_tr", ts); |
|
232 |
||
233 |
||
234 |
(* implicit structures *) |
|
235 |
||
236 |
fun the_struct structs i = |
|
30146 | 237 |
if 1 <= i andalso i <= length structs then nth structs (i - 1) |
31542
3371a3c19bb1
reraise exceptions to preserve position information;
wenzelm
parents:
30146
diff
changeset
|
238 |
else error ("Illegal reference to implicit structure #" ^ string_of_int i); |
14697 | 239 |
|
240 |
fun struct_tr structs (*"_struct"*) [Const ("_indexdefault", _)] = |
|
241 |
Lexicon.free (the_struct structs 1) |
|
242 |
| struct_tr structs (*"_struct"*) [t as (Const ("_indexnum", _) $ Const (s, _))] = |
|
243 |
Lexicon.free (the_struct structs |
|
15531 | 244 |
(case Lexicon.read_nat s of SOME n => n | NONE => raise TERM ("struct_tr", [t]))) |
14697 | 245 |
| struct_tr _ (*"_struct"*) ts = raise TERM ("struct_tr", ts); |
12122 | 246 |
|
247 |
||
5084 | 248 |
|
548 | 249 |
(** print (ast) translations **) |
250 |
||
14647 | 251 |
(* types *) |
252 |
||
253 |
fun non_typed_tr' f _ _ ts = f ts; |
|
254 |
fun non_typed_tr'' f x _ _ ts = f x ts; |
|
255 |
||
256 |
||
548 | 257 |
(* application *) |
258 |
||
5690 | 259 |
fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f]) |
260 |
| appl_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_appl", f, Ast.fold_ast "_args" args]; |
|
548 | 261 |
|
5690 | 262 |
fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f]) |
263 |
| 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
|
264 |
|
548 | 265 |
|
42085 | 266 |
(* partial eta-contraction before printing *) |
267 |
||
268 |
fun eta_abs (Abs (a, T, t)) = |
|
269 |
(case eta_abs t of |
|
270 |
t' as Const ("_aprop", _) $ _ => Abs (a, T, t') |
|
271 |
| t' as f $ u => |
|
272 |
(case eta_abs u of |
|
273 |
Bound 0 => |
|
274 |
if Term.is_dependent f then Abs (a, T, t') |
|
275 |
else incr_boundvars ~1 f |
|
276 |
| _ => Abs (a, T, t')) |
|
277 |
| t' => Abs (a, T, t')) |
|
278 |
| eta_abs t = t; |
|
279 |
||
280 |
val eta_contract_default = Unsynchronized.ref true; |
|
281 |
val eta_contract_raw = Config.declare "eta_contract" (fn _ => Config.Bool (! eta_contract_default)); |
|
282 |
val eta_contract = Config.bool eta_contract_raw; |
|
283 |
||
284 |
fun eta_contr ctxt tm = |
|
285 |
if Config.get ctxt eta_contract then eta_abs tm else tm; |
|
286 |
||
287 |
||
548 | 288 |
(* abstraction *) |
289 |
||
19311 | 290 |
fun mark_boundT (x, T) = Const ("_bound", T --> T) $ Free (x, T); |
2698 | 291 |
fun mark_bound x = mark_boundT (x, dummyT); |
292 |
||
18958 | 293 |
fun bound_vars vars body = |
294 |
subst_bounds (map mark_boundT (Term.rename_wrt_term body vars), body); |
|
295 |
||
548 | 296 |
fun strip_abss vars_of body_of tm = |
297 |
let |
|
298 |
val vars = vars_of tm; |
|
299 |
val body = body_of tm; |
|
29276 | 300 |
val rev_new_vars = Term.rename_wrt_term body vars; |
21750 | 301 |
fun subst (x, T) b = |
42083
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents:
42080
diff
changeset
|
302 |
if can Name.dest_internal x andalso not (Term.is_dependent b) |
21750 | 303 |
then (Const ("_idtdummy", T), incr_boundvars ~1 b) |
304 |
else (mark_boundT (x, T), Term.subst_bound (mark_bound x, b)); |
|
305 |
val (rev_vars', body') = fold_map subst rev_new_vars body; |
|
306 |
in (rev rev_vars', body') end; |
|
548 | 307 |
|
3928
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
wenzelm
parents:
3777
diff
changeset
|
308 |
|
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
37216
diff
changeset
|
309 |
fun abs_tr' ctxt tm = |
19473 | 310 |
uncurry (fold_rev (fn x => fn t => Lexicon.const "_abs" $ x $ t)) |
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
37216
diff
changeset
|
311 |
(strip_abss strip_abs_vars strip_abs_body (eta_contr ctxt tm)); |
548 | 312 |
|
14697 | 313 |
fun atomic_abs_tr' (x, T, t) = |
29276 | 314 |
let val [xT] = Term.rename_wrt_term t [(x, T)] |
14697 | 315 |
in (mark_boundT xT, subst_bound (mark_bound (fst xT), t)) end; |
13762 | 316 |
|
42085 | 317 |
fun abs_ast_tr' asts = |
5690 | 318 |
(case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of |
319 |
([], _) => raise Ast.AST ("abs_ast_tr'", asts) |
|
320 |
| (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]); |
|
548 | 321 |
|
42086
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
322 |
fun const_abs_tr' t = |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
323 |
(case eta_abs t of |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
324 |
Abs (_, _, t') => |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
325 |
if Term.is_dependent t' then raise Match |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
326 |
else incr_boundvars ~1 t' |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
327 |
| _ => raise Match); |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
328 |
|
32120
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
haftmann
parents:
31542
diff
changeset
|
329 |
|
42085 | 330 |
(* binders *) |
548 | 331 |
|
21535 | 332 |
fun mk_binder_tr' (name, syn) = |
548 | 333 |
let |
334 |
fun mk_idts [] = raise Match (*abort translation*) |
|
335 |
| mk_idts [idt] = idt |
|
5690 | 336 |
| mk_idts (idt :: idts) = Lexicon.const "_idts" $ idt $ mk_idts idts; |
548 | 337 |
|
338 |
fun tr' t = |
|
339 |
let |
|
340 |
val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t; |
|
21535 | 341 |
in Lexicon.const syn $ mk_idts xs $ bd end; |
548 | 342 |
|
21535 | 343 |
fun binder_tr' (t :: ts) = Term.list_comb (tr' (Lexicon.const name $ t), ts) |
344 |
| binder_tr' [] = raise Match; |
|
345 |
in (name, binder_tr') end; |
|
548 | 346 |
|
42085 | 347 |
fun preserve_binder_abs_tr' name syn = (name, fn Abs abs :: ts => |
348 |
let val (x, t) = atomic_abs_tr' abs |
|
349 |
in list_comb (Lexicon.const syn $ x $ t, ts) end); |
|
350 |
||
351 |
fun preserve_binder_abs2_tr' name syn = (name, fn A :: Abs abs :: ts => |
|
352 |
let val (x, t) = atomic_abs_tr' abs |
|
353 |
in list_comb (Lexicon.const syn $ x $ A $ t, ts) end); |
|
354 |
||
548 | 355 |
|
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
356 |
(* idtyp constraints *) |
548 | 357 |
|
42045 | 358 |
fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant "_constrain", x, ty], xs] = |
359 |
Ast.Appl [Ast.Constant a, Ast.Appl [Ast.Constant "_idtyp", x, ty], xs] |
|
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
360 |
| idtyp_ast_tr' _ _ = raise Match; |
548 | 361 |
|
362 |
||
28628 | 363 |
(* type propositions *) |
364 |
||
35255 | 365 |
fun type_prop_tr' _ (*"_type_prop"*) T [Const ("\\<^const>Pure.sort_constraint", _)] = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
36502
diff
changeset
|
366 |
Lexicon.const "_sort_constraint" $ Type_Ext.term_of_typ true T |
28628 | 367 |
| type_prop_tr' show_sorts (*"_type_prop"*) T [t] = |
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
36502
diff
changeset
|
368 |
Lexicon.const "_ofclass" $ Type_Ext.term_of_typ show_sorts T $ t |
28628 | 369 |
| type_prop_tr' _ (*"_type_prop"*) T ts = raise TYPE ("type_prop_tr'", [T], ts); |
370 |
||
371 |
||
548 | 372 |
(* meta propositions *) |
373 |
||
4148 | 374 |
fun prop_tr' tm = |
548 | 375 |
let |
5690 | 376 |
fun aprop t = Lexicon.const "_aprop" $ t; |
548 | 377 |
|
2698 | 378 |
fun is_prop Ts t = |
379 |
fastype_of1 (Ts, t) = propT handle TERM _ => false; |
|
548 | 380 |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
381 |
fun is_term (Const ("Pure.term", _) $ _) = true |
19848 | 382 |
| is_term _ = false; |
383 |
||
548 | 384 |
fun tr' _ (t as Const _) = t |
18478
29a5070b517c
prop_tr': proper handling of aprop marked as bound;
wenzelm
parents:
18342
diff
changeset
|
385 |
| tr' Ts (t as Const ("_bound", _) $ u) = |
29a5070b517c
prop_tr': proper handling of aprop marked as bound;
wenzelm
parents:
18342
diff
changeset
|
386 |
if is_prop Ts u then aprop t else t |
2698 | 387 |
| tr' _ (t as Free (x, T)) = |
5690 | 388 |
if T = propT then aprop (Lexicon.free x) else t |
2698 | 389 |
| tr' _ (t as Var (xi, T)) = |
5690 | 390 |
if T = propT then aprop (Lexicon.var xi) else t |
2698 | 391 |
| tr' Ts (t as Bound _) = |
392 |
if is_prop Ts t then aprop t else t |
|
393 |
| tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t) |
|
35429
afa8cf9e63d8
authentic syntax for classes and type constructors;
wenzelm
parents:
35255
diff
changeset
|
394 |
| tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) = |
28628 | 395 |
if is_prop Ts t andalso not (is_term t) then Const ("_type_prop", T) $ tr' Ts t1 |
2698 | 396 |
else tr' Ts t1 $ tr' Ts t2 |
397 |
| tr' Ts (t as t1 $ t2) = |
|
5690 | 398 |
(if is_Const (Term.head_of t) orelse not (is_prop Ts t) |
2698 | 399 |
then I else aprop) (tr' Ts t1 $ tr' Ts t2); |
16612
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
400 |
in tr' [] tm end; |
548 | 401 |
|
2698 | 402 |
|
548 | 403 |
(* meta implication *) |
404 |
||
405 |
fun impl_ast_tr' (*"==>"*) asts = |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
36502
diff
changeset
|
406 |
if Type_Ext.no_brackets () then raise Match |
10572 | 407 |
else |
35255 | 408 |
(case Ast.unfold_ast_p "\\<^const>==>" (Ast.Appl (Ast.Constant "\\<^const>==>" :: asts)) of |
16612
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
409 |
(prems as _ :: _ :: _, concl) => |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
410 |
let |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
411 |
val (asms, asm) = split_last prems; |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
412 |
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
|
413 |
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
|
414 |
| _ => raise Match); |
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
schirmer
parents:
14981
diff
changeset
|
415 |
|
548 | 416 |
|
4148 | 417 |
(* type reflection *) |
418 |
||
419 |
fun type_tr' show_sorts (*"TYPE"*) (Type ("itself", [T])) ts = |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
36502
diff
changeset
|
420 |
Term.list_comb (Lexicon.const "_TYPE" $ Type_Ext.term_of_typ show_sorts T, ts) |
4148 | 421 |
| type_tr' _ _ _ = raise Match; |
422 |
||
423 |
||
19577 | 424 |
(* type constraints *) |
425 |
||
426 |
fun type_constraint_tr' show_sorts (*"_type_constraint_"*) (Type ("fun", [T, _])) (t :: ts) = |
|
37216
3165bc303f66
modernized some structure names, keeping a few legacy aliases;
wenzelm
parents:
36502
diff
changeset
|
427 |
Term.list_comb (Lexicon.const Syn_Ext.constrainC $ t $ Type_Ext.term_of_typ show_sorts T, ts) |
19577 | 428 |
| type_constraint_tr' _ _ _ = raise Match; |
429 |
||
430 |
||
548 | 431 |
(* dependent / nondependent quantifiers *) |
432 |
||
20202 | 433 |
fun var_abs mark (x, T, b) = |
434 |
let val ([x'], _) = Name.variants [x] (Term.declare_term_names b Name.context) |
|
435 |
in (x', subst_bound (mark (x', T), b)) end; |
|
436 |
||
437 |
val variant_abs = var_abs Free; |
|
438 |
val variant_abs' = var_abs mark_boundT; |
|
2698 | 439 |
|
548 | 440 |
fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) = |
42083
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents:
42080
diff
changeset
|
441 |
if Term.is_dependent B then |
2698 | 442 |
let val (x', B') = variant_abs' (x, dummyT, B); |
5690 | 443 |
in Term.list_comb (Lexicon.const q $ mark_boundT (x', T) $ A $ B', ts) end |
42084
532b3a76103f
dependent_tr': formal treatment of bounds after stripping Abs, although it should only happen for malformed terms, since print_translations work top-down;
wenzelm
parents:
42083
diff
changeset
|
444 |
else Term.list_comb (Lexicon.const r $ A $ incr_boundvars ~1 B, ts) |
548 | 445 |
| dependent_tr' _ _ = raise Match; |
446 |
||
447 |
||
5084 | 448 |
(* quote / antiquote *) |
449 |
||
8577 | 450 |
fun antiquote_tr' name = |
451 |
let |
|
452 |
fun tr' i (t $ u) = |
|
42084
532b3a76103f
dependent_tr': formal treatment of bounds after stripping Abs, although it should only happen for malformed terms, since print_translations work top-down;
wenzelm
parents:
42083
diff
changeset
|
453 |
if u aconv Bound i then Lexicon.const name $ tr' i t |
532b3a76103f
dependent_tr': formal treatment of bounds after stripping Abs, although it should only happen for malformed terms, since print_translations work top-down;
wenzelm
parents:
42083
diff
changeset
|
454 |
else tr' i t $ tr' i u |
8577 | 455 |
| tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t) |
18139 | 456 |
| tr' i a = if a aconv Bound i then raise Match else a; |
8577 | 457 |
in tr' 0 end; |
458 |
||
459 |
fun quote_tr' name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' name t) |
|
460 |
| quote_tr' _ _ = raise Match; |
|
461 |
||
5084 | 462 |
fun quote_antiquote_tr' quoteN antiquoteN name = |
463 |
let |
|
8577 | 464 |
fun tr' (t :: ts) = Term.list_comb (Lexicon.const quoteN $ quote_tr' antiquoteN t, ts) |
465 |
| tr' _ = raise Match; |
|
466 |
in (name, tr') end; |
|
5084 | 467 |
|
468 |
||
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
469 |
(* corresponding updates *) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
470 |
|
42080
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
471 |
local |
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
472 |
|
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
473 |
fun upd_type (Type ("fun", [Type ("fun", [_, T]), _])) = T |
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
474 |
| upd_type _ = dummyT; |
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
475 |
|
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
476 |
fun upd_tr' (x_upd, T) = |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
477 |
(case try (unsuffix "_update") x_upd of |
42080
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
478 |
SOME x => (x, upd_type T) |
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
479 |
| NONE => raise Match); |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
480 |
|
42080
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
481 |
in |
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
482 |
|
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
483 |
fun update_name_tr' (Free x) = Free (upd_tr' x) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
484 |
| update_name_tr' ((c as Const ("_free", _)) $ Free x) = c $ Free (upd_tr' x) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
485 |
| update_name_tr' (Const x) = Const (upd_tr' x) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
486 |
| update_name_tr' _ = raise Match; |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
487 |
|
42080
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
488 |
end; |
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
489 |
|
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
490 |
|
14697 | 491 |
(* indexed syntax *) |
548 | 492 |
|
14697 | 493 |
fun index_ast_tr' (*"_index"*) [Ast.Appl [Ast.Constant "_struct", ast]] = ast |
494 |
| index_ast_tr' _ = raise Match; |
|
495 |
||
496 |
||
497 |
(* implicit structures *) |
|
498 |
||
499 |
fun the_struct' structs s = |
|
500 |
[(case Lexicon.read_nat s of |
|
18678 | 501 |
SOME i => Ast.Variable (the_struct structs i handle ERROR _ => raise Match) |
15531 | 502 |
| NONE => raise Match)] |> Ast.mk_appl (Ast.Constant "_free"); |
14697 | 503 |
|
504 |
fun struct_ast_tr' structs (*"_struct"*) [Ast.Constant "_indexdefault"] = |
|
505 |
the_struct' structs "1" |
|
506 |
| struct_ast_tr' structs (*"_struct"*) [Ast.Appl [Ast.Constant "_indexnum", Ast.Constant s]] = |
|
507 |
the_struct' structs s |
|
508 |
| struct_ast_tr' _ _ = raise Match; |
|
509 |
||
510 |
||
511 |
||
512 |
(** Pure translations **) |
|
548 | 513 |
|
514 |
val pure_trfuns = |
|
42057
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
515 |
([("_strip_positions", strip_positions_ast_tr), |
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
wenzelm
parents:
42055
diff
changeset
|
516 |
("_constify", constify_ast_tr), |
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
517 |
("_appl", appl_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
518 |
("_applC", applC_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
519 |
("_lambda", lambda_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
520 |
("_idtyp", idtyp_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
521 |
("_idtypdummy", idtypdummy_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
522 |
("_bigimpl", bigimpl_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
523 |
("_indexdefault", indexdefault_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
524 |
("_indexnum", indexnum_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
525 |
("_indexvar", indexvar_ast_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
526 |
("_struct", struct_ast_tr)], |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
527 |
[("_abs", abs_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
528 |
("_aprop", aprop_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
529 |
("_ofclass", ofclass_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
530 |
("_sort_constraint", sort_constraint_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
531 |
("_TYPE", type_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
532 |
("_DDDOT", dddot_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
533 |
("_update_name", update_name_tr), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
534 |
("_index", index_tr)], |
35198 | 535 |
([]: (string * (term list -> term)) list), |
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
536 |
[("_abs", abs_ast_tr'), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
537 |
("_idts", idtyp_ast_tr' "_idts"), |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
538 |
("_pttrns", idtyp_ast_tr' "_pttrns"), |
35255 | 539 |
("\\<^const>==>", impl_ast_tr'), |
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
540 |
("_index", index_ast_tr')]); |
548 | 541 |
|
2698 | 542 |
val pure_trfunsT = |
35255 | 543 |
[("_type_prop", type_prop_tr'), |
544 |
("\\<^const>TYPE", type_tr'), |
|
545 |
("_type_constraint_", type_constraint_tr')]; |
|
2698 | 546 |
|
14697 | 547 |
fun struct_trfuns structs = |
548 |
([], [("_struct", struct_tr structs)], [], [("_struct", struct_ast_tr' structs)]); |
|
549 |
||
548 | 550 |
|
551 |
||
42043
2055515c9d3f
clarified Syn_Trans.parsetree_to_ast and Syn_Trans.ast_to_term;
wenzelm
parents:
40235
diff
changeset
|
552 |
(** parsetree_to_ast **) |
548 | 553 |
|
42204 | 554 |
type type_context = |
555 |
{get_class: string -> string, |
|
556 |
get_type: string -> string, |
|
557 |
markup_class: string -> Markup.T list, |
|
558 |
markup_type: string -> Markup.T list}; |
|
559 |
||
560 |
fun parsetree_to_ast ctxt (type_context: type_context) constrain_pos trf parsetree = |
|
548 | 561 |
let |
42204 | 562 |
val {get_class, get_type, markup_class, markup_type} = type_context; |
563 |
||
564 |
val reports = Unsynchronized.ref ([]: Position.reports); |
|
565 |
fun report pos = Position.reports reports [pos]; |
|
566 |
||
548 | 567 |
fun trans a args = |
568 |
(case trf a of |
|
15531 | 569 |
NONE => Ast.mk_appl (Ast.Constant a) args |
23937
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents:
23824
diff
changeset
|
570 |
| SOME f => f ctxt args); |
548 | 571 |
|
42204 | 572 |
fun ast_of (Parser.Node ("_class_name", [Parser.Tip tok])) = |
573 |
let |
|
574 |
val c = get_class (Lexicon.str_of_token tok); |
|
575 |
val _ = report (Lexicon.pos_of_token tok) markup_class c; |
|
576 |
in Ast.Constant (Lexicon.mark_class c) end |
|
577 |
| ast_of (Parser.Node ("_type_name", [Parser.Tip tok])) = |
|
578 |
let |
|
579 |
val c = get_type (Lexicon.str_of_token tok); |
|
580 |
val _ = report (Lexicon.pos_of_token tok) markup_type c; |
|
581 |
in Ast.Constant (Lexicon.mark_type c) end |
|
582 |
| ast_of (Parser.Node ("_constrain_position", [pt as Parser.Tip tok])) = |
|
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
583 |
if constrain_pos then |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
584 |
Ast.Appl [Ast.Constant "_constrain", ast_of pt, |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
585 |
Ast.Variable (Lexicon.encode_position (Lexicon.pos_of_token tok))] |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
586 |
else ast_of pt |
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
587 |
| ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts) |
5690 | 588 |
| ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok); |
42204 | 589 |
|
42223
098c86e53153
more precise propagation of reports/results through some inner syntax layers;
wenzelm
parents:
42204
diff
changeset
|
590 |
val ast = Exn.interruptible_capture ast_of parsetree; |
42204 | 591 |
in (! reports, ast) end; |
548 | 592 |
|
593 |
||
594 |
||
42043
2055515c9d3f
clarified Syn_Trans.parsetree_to_ast and Syn_Trans.ast_to_term;
wenzelm
parents:
40235
diff
changeset
|
595 |
(** ast_to_term **) |
548 | 596 |
|
42043
2055515c9d3f
clarified Syn_Trans.parsetree_to_ast and Syn_Trans.ast_to_term;
wenzelm
parents:
40235
diff
changeset
|
597 |
fun ast_to_term ctxt trf = |
548 | 598 |
let |
599 |
fun trans a args = |
|
600 |
(case trf a of |
|
15531 | 601 |
NONE => Term.list_comb (Lexicon.const a, args) |
23937
66e1f24d655d
eliminated transform_failure (to avoid critical section for main transactions);
wenzelm
parents:
23824
diff
changeset
|
602 |
| SOME f => f ctxt args); |
548 | 603 |
|
5690 | 604 |
fun term_of (Ast.Constant a) = trans a [] |
605 |
| term_of (Ast.Variable x) = Lexicon.read_var x |
|
606 |
| term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) = |
|
548 | 607 |
trans a (map term_of asts) |
5690 | 608 |
| term_of (Ast.Appl (ast :: (asts as _ :: _))) = |
609 |
Term.list_comb (term_of ast, map term_of asts) |
|
610 |
| term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]); |
|
42043
2055515c9d3f
clarified Syn_Trans.parsetree_to_ast and Syn_Trans.ast_to_term;
wenzelm
parents:
40235
diff
changeset
|
611 |
in term_of end; |
548 | 612 |
|
613 |
end; |