author | wenzelm |
Thu, 03 Jan 2019 15:55:48 +0100 | |
changeset 69577 | 015f43ee4bb7 |
parent 69575 | f77cc54f6d47 |
child 74442 | f5c5006d142e |
permissions | -rw-r--r-- |
42284 | 1 |
(* Title: Pure/Syntax/syntax_trans.ML |
548 | 2 |
Author: Tobias Nipkow and Markus Wenzel, TU Muenchen |
3 |
||
4 |
Syntax translation functions. |
|
5 |
*) |
|
6 |
||
42284 | 7 |
signature BASIC_SYNTAX_TRANS = |
8 |
sig |
|
9 |
val eta_contract: bool Config.T |
|
10 |
end |
|
11 |
||
12 |
signature SYNTAX_TRANS = |
|
2698 | 13 |
sig |
42284 | 14 |
include BASIC_SYNTAX_TRANS |
45057
86c9b73158a8
default print mode for Isabelle/Scala, not just Isabelle/jEdit;
wenzelm
parents:
44433
diff
changeset
|
15 |
val bracketsN: string |
86c9b73158a8
default print mode for Isabelle/Scala, not just Isabelle/jEdit;
wenzelm
parents:
44433
diff
changeset
|
16 |
val no_bracketsN: string |
42284 | 17 |
val no_brackets: unit -> bool |
45057
86c9b73158a8
default print mode for Isabelle/Scala, not just Isabelle/jEdit;
wenzelm
parents:
44433
diff
changeset
|
18 |
val type_bracketsN: string |
86c9b73158a8
default print mode for Isabelle/Scala, not just Isabelle/jEdit;
wenzelm
parents:
44433
diff
changeset
|
19 |
val no_type_bracketsN: string |
42284 | 20 |
val no_type_brackets: unit -> bool |
21 |
val abs_tr: term list -> term |
|
52143 | 22 |
val mk_binder_tr: string * string -> string * (Proof.context -> term list -> term) |
42284 | 23 |
val antiquote_tr: string -> term -> term |
24 |
val quote_tr: string -> term -> term |
|
52143 | 25 |
val quote_antiquote_tr: string -> string -> string -> |
26 |
string * (Proof.context -> term list -> term) |
|
27 |
val non_typed_tr': (Proof.context -> term list -> term) -> |
|
28 |
Proof.context -> typ -> term list -> term |
|
42284 | 29 |
val tappl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
30 |
val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
|
31 |
val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast |
|
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
32 |
val mark_bound_abs: string * typ -> term |
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
33 |
val mark_bound_body: string * typ -> term |
42284 | 34 |
val bound_vars: (string * typ) list -> term -> term |
35 |
val abs_tr': Proof.context -> term -> term |
|
13762 | 36 |
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
|
37 |
val const_abs_tr': term -> term |
52143 | 38 |
val mk_binder_tr': string * string -> string * (Proof.context -> term list -> term) |
39 |
val preserve_binder_abs_tr': string -> string -> string * (Proof.context -> term list -> term) |
|
40 |
val preserve_binder_abs2_tr': string -> string -> string * (Proof.context -> term list -> term) |
|
42284 | 41 |
val variant_abs: string * typ * term -> string * term |
42 |
val variant_abs': string * typ * term -> string * term |
|
548 | 43 |
val dependent_tr': string * string -> term list -> term |
8577 | 44 |
val antiquote_tr': string -> term -> term |
45 |
val quote_tr': string -> term -> term |
|
52143 | 46 |
val quote_antiquote_tr': string -> string -> string -> |
47 |
string * (Proof.context -> term list -> term) |
|
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
48 |
val update_name_tr': term -> term |
59152
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
49 |
val get_idents: Proof.context -> {structs: string list, fixes: string list} |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
50 |
val put_idents: {structs: string list, fixes: string list} -> Proof.context -> Proof.context |
52143 | 51 |
val pure_parse_ast_translation: (string * (Proof.context -> Ast.ast list -> Ast.ast)) list |
52 |
val pure_parse_translation: (string * (Proof.context -> term list -> term)) list |
|
53 |
val pure_print_ast_translation: (string * (Proof.context -> Ast.ast list -> Ast.ast)) list |
|
2698 | 54 |
end; |
548 | 55 |
|
42284 | 56 |
structure Syntax_Trans: SYNTAX_TRANS = |
548 | 57 |
struct |
2698 | 58 |
|
42476 | 59 |
structure Syntax = Lexicon.Syntax; |
60 |
||
61 |
||
42262 | 62 |
(* print mode *) |
63 |
||
64 |
val bracketsN = "brackets"; |
|
65 |
val no_bracketsN = "no_brackets"; |
|
66 |
||
67 |
fun no_brackets () = |
|
68 |
find_first (fn mode => mode = bracketsN orelse mode = no_bracketsN) |
|
69 |
(print_mode_value ()) = SOME no_bracketsN; |
|
70 |
||
71 |
val type_bracketsN = "type_brackets"; |
|
72 |
val no_type_bracketsN = "no_type_brackets"; |
|
73 |
||
74 |
fun no_type_brackets () = |
|
75 |
find_first (fn mode => mode = type_bracketsN orelse mode = no_type_bracketsN) |
|
76 |
(print_mode_value ()) <> SOME type_bracketsN; |
|
77 |
||
78 |
||
2698 | 79 |
|
548 | 80 |
(** parse (ast) translations **) |
81 |
||
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
|
82 |
(* 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
|
83 |
|
42264 | 84 |
fun strip_positions_ast_tr [ast] = Ast.strip_positions ast |
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
|
85 |
| 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
|
86 |
|
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
|
87 |
|
11491 | 88 |
(* constify *) |
89 |
||
46236
ae79f2978a67
position constraints for numerals enable PIDE markup;
wenzelm
parents:
45389
diff
changeset
|
90 |
fun constify_ast_tr [Ast.Appl [c as Ast.Constant "_constrain", ast1, ast2]] = |
ae79f2978a67
position constraints for numerals enable PIDE markup;
wenzelm
parents:
45389
diff
changeset
|
91 |
Ast.Appl [c, constify_ast_tr [ast1], ast2] |
ae79f2978a67
position constraints for numerals enable PIDE markup;
wenzelm
parents:
45389
diff
changeset
|
92 |
| constify_ast_tr [Ast.Variable c] = Ast.Constant c |
11491 | 93 |
| constify_ast_tr asts = raise Ast.AST ("constify_ast_tr", asts); |
94 |
||
95 |
||
42262 | 96 |
(* type syntax *) |
97 |
||
98 |
fun tapp_ast_tr [ty, c] = Ast.Appl [c, ty] |
|
99 |
| tapp_ast_tr asts = raise Ast.AST ("tapp_ast_tr", asts); |
|
100 |
||
101 |
fun tappl_ast_tr [ty, tys, c] = Ast.mk_appl c (ty :: Ast.unfold_ast "_types" tys) |
|
102 |
| tappl_ast_tr asts = raise Ast.AST ("tappl_ast_tr", asts); |
|
103 |
||
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
104 |
fun bracket_ast_tr [dom, cod] = Ast.fold_ast_p "\<^type>fun" (Ast.unfold_ast "_types" dom, cod) |
42262 | 105 |
| bracket_ast_tr asts = raise Ast.AST ("bracket_ast_tr", asts); |
106 |
||
107 |
||
548 | 108 |
(* application *) |
109 |
||
5690 | 110 |
fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args) |
111 |
| 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
|
112 |
|
5690 | 113 |
fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args) |
114 |
| applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts); |
|
548 | 115 |
|
116 |
||
117 |
(* abstraction *) |
|
118 |
||
42278 | 119 |
fun idtyp_ast_tr [x, ty] = Ast.Appl [Ast.Constant "_constrain", x, ty] |
120 |
| idtyp_ast_tr asts = raise Ast.AST ("idtyp_ast_tr", asts); |
|
548 | 121 |
|
42278 | 122 |
fun lambda_ast_tr [pats, body] = Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body) |
123 |
| lambda_ast_tr asts = raise Ast.AST ("lambda_ast_tr", asts); |
|
548 | 124 |
|
44241 | 125 |
fun absfree_proper (x, T) t = |
55948 | 126 |
if Name.is_internal x |
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
127 |
then error ("Illegal internal variable in abstraction: " ^ quote x) |
44241 | 128 |
else absfree (x, T) t; |
21773 | 129 |
|
44241 | 130 |
fun abs_tr [Free x, t] = absfree_proper x t |
131 |
| abs_tr [Const ("_idtdummy", T), t] = absdummy T t |
|
42284 | 132 |
| abs_tr [Const ("_constrain", _) $ x $ tT, t] = |
42476 | 133 |
Syntax.const "_constrainAbs" $ abs_tr [x, t] $ tT |
42048
afd11ca8e018
support for encoded positions (for id_position, longid_position) as pseudo type-constraints -- still inactive;
wenzelm
parents:
42045
diff
changeset
|
134 |
| abs_tr ts = raise TERM ("abs_tr", ts); |
548 | 135 |
|
136 |
||
137 |
(* binder *) |
|
138 |
||
21535 | 139 |
fun mk_binder_tr (syn, name) = |
548 | 140 |
let |
42055 | 141 |
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
|
142 |
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
|
143 |
| binder_tr [x, t] = |
42055 | 144 |
let val abs = abs_tr [x, t] handle TERM _ => err [x, t] |
42476 | 145 |
in Syntax.const name $ abs end |
42055 | 146 |
| binder_tr ts = err ts; |
52143 | 147 |
in (syn, fn _ => binder_tr) end; |
548 | 148 |
|
149 |
||
28628 | 150 |
(* type propositions *) |
151 |
||
35255 | 152 |
fun mk_type ty = |
42476 | 153 |
Syntax.const "_constrain" $ |
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
154 |
Syntax.const "\<^const>Pure.type" $ (Syntax.const "\<^type>itself" $ ty); |
28628 | 155 |
|
42278 | 156 |
fun ofclass_tr [ty, cls] = cls $ mk_type ty |
157 |
| ofclass_tr ts = raise TERM ("ofclass_tr", ts); |
|
28628 | 158 |
|
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
159 |
fun sort_constraint_tr [ty] = Syntax.const "\<^const>Pure.sort_constraint" $ mk_type ty |
42278 | 160 |
| sort_constraint_tr ts = raise TERM ("sort_constraint_tr", ts); |
28628 | 161 |
|
162 |
||
548 | 163 |
(* meta propositions *) |
164 |
||
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
165 |
fun aprop_tr [t] = Syntax.const "_constrain" $ t $ Syntax.const "\<^type>prop" |
42278 | 166 |
| aprop_tr ts = raise TERM ("aprop_tr", ts); |
548 | 167 |
|
168 |
||
169 |
(* meta implication *) |
|
170 |
||
42278 | 171 |
fun bigimpl_ast_tr (asts as [asms, concl]) = |
16612
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
172 |
let val prems = |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
173 |
(case Ast.unfold_ast_p "_asms" asms of |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
174 |
(asms', Ast.Appl [Ast.Constant "_asm", asm']) => asms' @ [asm'] |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
175 |
| _ => raise Ast.AST ("bigimpl_ast_tr", asts)) |
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
176 |
in Ast.fold_ast_p "\<^const>Pure.imp" (prems, concl) end |
42278 | 177 |
| bigimpl_ast_tr asts = raise Ast.AST ("bigimpl_ast_tr", asts); |
15421
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
schirmer
parents:
14981
diff
changeset
|
178 |
|
548 | 179 |
|
23824
8ad7131dbfcf
moved print_translations from Pure.thy to Syntax/syn_trans.ML;
wenzelm
parents:
21773
diff
changeset
|
180 |
(* type/term reflection *) |
4148 | 181 |
|
42278 | 182 |
fun type_tr [ty] = mk_type ty |
183 |
| type_tr ts = raise TERM ("type_tr", ts); |
|
4148 | 184 |
|
548 | 185 |
|
5084 | 186 |
(* quote / antiquote *) |
187 |
||
8577 | 188 |
fun antiquote_tr name = |
189 |
let |
|
190 |
fun tr i ((t as Const (c, _)) $ u) = |
|
191 |
if c = name then tr i u $ Bound i |
|
192 |
else tr i t $ tr i u |
|
193 |
| tr i (t $ u) = tr i t $ tr i u |
|
194 |
| tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t) |
|
195 |
| tr _ a = a; |
|
196 |
in tr 0 end; |
|
197 |
||
198 |
fun quote_tr name t = Abs ("s", dummyT, antiquote_tr name (Term.incr_boundvars 1 t)); |
|
199 |
||
5084 | 200 |
fun quote_antiquote_tr quoteN antiquoteN name = |
201 |
let |
|
42476 | 202 |
fun tr [t] = Syntax.const name $ quote_tr antiquoteN t |
8577 | 203 |
| tr ts = raise TERM ("quote_tr", ts); |
52143 | 204 |
in (quoteN, fn _ => tr) end; |
5084 | 205 |
|
206 |
||
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
207 |
(* corresponding updates *) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
208 |
|
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
209 |
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
|
210 |
| 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
|
211 |
| update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) = |
42264 | 212 |
if Term_Position.is_position ty then list_comb (c $ update_name_tr [t] $ ty, ts) |
42053
006095137a81
update_name_tr: more precise handling of explicit constraints, including positions;
wenzelm
parents:
42048
diff
changeset
|
213 |
else |
006095137a81
update_name_tr: more precise handling of explicit constraints, including positions;
wenzelm
parents:
42048
diff
changeset
|
214 |
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
|
215 |
(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
|
216 |
(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
|
217 |
| 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
|
218 |
|
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
219 |
|
14697 | 220 |
(* indexed syntax *) |
221 |
||
59152
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
222 |
structure Idents = Proof_Data |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
223 |
( |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
224 |
type T = {structs: string list, fixes: string list}; |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
225 |
fun init _ : T = {structs = [], fixes = []}; |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
226 |
); |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
227 |
|
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
228 |
val get_idents = Idents.get; |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
229 |
val put_idents = Idents.put; |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
230 |
|
45389
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
231 |
fun indexdefault_ast_tr [] = |
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
232 |
Ast.Appl [Ast.Constant "_index", |
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
233 |
Ast.Appl [Ast.Constant "_struct", Ast.Constant "_indexdefault"]] |
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
234 |
| indexdefault_ast_tr asts = raise Ast.AST ("indexdefault_ast_tr", asts); |
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
235 |
|
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
236 |
fun indexvar_ast_tr [] = Ast.Appl [Ast.Constant "_index", Ast.Variable "some_index"] |
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
237 |
| indexvar_ast_tr asts = raise Ast.AST ("indexvar_ast_tr", asts); |
bc0d50f8ae19
discontinued numbered structure indexes (legacy feature);
wenzelm
parents:
45057
diff
changeset
|
238 |
|
42278 | 239 |
fun struct_ast_tr [Ast.Appl [Ast.Constant "_index", ast]] = ast |
240 |
| struct_ast_tr asts = Ast.mk_appl (Ast.Constant "_struct") asts; |
|
14697 | 241 |
|
42278 | 242 |
fun index_tr [t] = t |
243 |
| index_tr ts = raise TERM ("index_tr", ts); |
|
14697 | 244 |
|
59152
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
245 |
fun struct_tr ctxt [Const ("_indexdefault", _)] = |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
246 |
(case #structs (get_idents ctxt) of |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
247 |
x :: _ => Syntax.const (Lexicon.mark_fixed x) |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
248 |
| _ => error "Illegal reference to implicit structure") |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
249 |
| struct_tr _ ts = raise TERM ("struct_tr", ts); |
12122 | 250 |
|
251 |
||
5084 | 252 |
|
548 | 253 |
(** print (ast) translations **) |
254 |
||
14647 | 255 |
(* types *) |
256 |
||
52143 | 257 |
fun non_typed_tr' f ctxt _ ts = f ctxt ts; |
14647 | 258 |
|
259 |
||
42262 | 260 |
(* type syntax *) |
261 |
||
262 |
fun tappl_ast_tr' (f, []) = raise Ast.AST ("tappl_ast_tr'", [f]) |
|
263 |
| tappl_ast_tr' (f, [ty]) = Ast.Appl [Ast.Constant "_tapp", ty, f] |
|
264 |
| tappl_ast_tr' (f, ty :: tys) = |
|
265 |
Ast.Appl [Ast.Constant "_tappl", ty, Ast.fold_ast "_types" tys, f]; |
|
266 |
||
267 |
fun fun_ast_tr' asts = |
|
268 |
if no_brackets () orelse no_type_brackets () then raise Match |
|
269 |
else |
|
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
270 |
(case Ast.unfold_ast_p "\<^type>fun" (Ast.Appl (Ast.Constant "\<^type>fun" :: asts)) of |
42262 | 271 |
(dom as _ :: _ :: _, cod) |
272 |
=> Ast.Appl [Ast.Constant "_bracket", Ast.fold_ast "_types" dom, cod] |
|
273 |
| _ => raise Match); |
|
274 |
||
275 |
||
548 | 276 |
(* application *) |
277 |
||
5690 | 278 |
fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f]) |
279 |
| appl_ast_tr' (f, args) = Ast.Appl [Ast.Constant "_appl", f, Ast.fold_ast "_args" args]; |
|
548 | 280 |
|
5690 | 281 |
fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f]) |
282 |
| 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
|
283 |
|
548 | 284 |
|
42085 | 285 |
(* partial eta-contraction before printing *) |
286 |
||
287 |
fun eta_abs (Abs (a, T, t)) = |
|
288 |
(case eta_abs t of |
|
289 |
t' as Const ("_aprop", _) $ _ => Abs (a, T, t') |
|
290 |
| t' as f $ u => |
|
291 |
(case eta_abs u of |
|
292 |
Bound 0 => |
|
293 |
if Term.is_dependent f then Abs (a, T, t') |
|
294 |
else incr_boundvars ~1 f |
|
295 |
| _ => Abs (a, T, t')) |
|
296 |
| t' => Abs (a, T, t')) |
|
297 |
| eta_abs t = t; |
|
298 |
||
69575 | 299 |
val eta_contract = Config.declare_option_bool ("eta_contract", \<^here>); |
69577 | 300 |
fun eta_contr ctxt = Config.get ctxt eta_contract ? eta_abs; |
42085 | 301 |
|
302 |
||
548 | 303 |
(* abstraction *) |
304 |
||
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
305 |
fun mark_bound_abs (x, T) = Const ("_bound", T --> T) $ Free (x, T); |
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
306 |
fun mark_bound_body (x, T) = Const ("_bound", dummyT) $ Free (x, T); |
2698 | 307 |
|
18958 | 308 |
fun bound_vars vars body = |
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
309 |
subst_bounds (map mark_bound_abs (Term.rename_wrt_term body vars), body); |
18958 | 310 |
|
548 | 311 |
fun strip_abss vars_of body_of tm = |
312 |
let |
|
313 |
val vars = vars_of tm; |
|
314 |
val body = body_of tm; |
|
29276 | 315 |
val rev_new_vars = Term.rename_wrt_term body vars; |
21750 | 316 |
fun subst (x, T) b = |
55948 | 317 |
if Name.is_internal x andalso not (Term.is_dependent b) |
21750 | 318 |
then (Const ("_idtdummy", T), incr_boundvars ~1 b) |
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
319 |
else (mark_bound_abs (x, T), Term.subst_bound (mark_bound_body (x, T), b)); |
21750 | 320 |
val (rev_vars', body') = fold_map subst rev_new_vars body; |
321 |
in (rev rev_vars', body') end; |
|
548 | 322 |
|
3928
787d2659ce4a
no longer tries bogus eta-contract involving aprops;
wenzelm
parents:
3777
diff
changeset
|
323 |
|
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
37216
diff
changeset
|
324 |
fun abs_tr' ctxt tm = |
42476 | 325 |
uncurry (fold_rev (fn x => fn t => Syntax.const "_abs" $ x $ t)) |
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
37216
diff
changeset
|
326 |
(strip_abss strip_abs_vars strip_abs_body (eta_contr ctxt tm)); |
548 | 327 |
|
14697 | 328 |
fun atomic_abs_tr' (x, T, t) = |
29276 | 329 |
let val [xT] = Term.rename_wrt_term t [(x, T)] |
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
330 |
in (mark_bound_abs xT, subst_bound (mark_bound_body xT, t)) end; |
13762 | 331 |
|
42085 | 332 |
fun abs_ast_tr' asts = |
5690 | 333 |
(case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of |
334 |
([], _) => raise Ast.AST ("abs_ast_tr'", asts) |
|
335 |
| (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]); |
|
548 | 336 |
|
42086
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
337 |
fun const_abs_tr' t = |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
338 |
(case eta_abs t of |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
339 |
Abs (_, _, t') => |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
340 |
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
|
341 |
else incr_boundvars ~1 t' |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
342 |
| _ => raise Match); |
74bf78db0d87
added Syntax.const_abs_tr' with proper eta_abs and Term.is_dependent;
wenzelm
parents:
42085
diff
changeset
|
343 |
|
32120
53a21a5e6889
attempt for more concise setup of non-etacontracting binders
haftmann
parents:
31542
diff
changeset
|
344 |
|
42085 | 345 |
(* binders *) |
548 | 346 |
|
21535 | 347 |
fun mk_binder_tr' (name, syn) = |
548 | 348 |
let |
349 |
fun mk_idts [] = raise Match (*abort translation*) |
|
350 |
| mk_idts [idt] = idt |
|
42476 | 351 |
| mk_idts (idt :: idts) = Syntax.const "_idts" $ idt $ mk_idts idts; |
548 | 352 |
|
353 |
fun tr' t = |
|
354 |
let |
|
355 |
val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t; |
|
42476 | 356 |
in Syntax.const syn $ mk_idts xs $ bd end; |
548 | 357 |
|
42476 | 358 |
fun binder_tr' (t :: ts) = Term.list_comb (tr' (Syntax.const name $ t), ts) |
21535 | 359 |
| binder_tr' [] = raise Match; |
52143 | 360 |
in (name, fn _ => binder_tr') end; |
548 | 361 |
|
52143 | 362 |
fun preserve_binder_abs_tr' name syn = (name, fn _ => fn Abs abs :: ts => |
42085 | 363 |
let val (x, t) = atomic_abs_tr' abs |
42476 | 364 |
in list_comb (Syntax.const syn $ x $ t, ts) end); |
42085 | 365 |
|
52143 | 366 |
fun preserve_binder_abs2_tr' name syn = (name, fn _ => fn A :: Abs abs :: ts => |
42085 | 367 |
let val (x, t) = atomic_abs_tr' abs |
42476 | 368 |
in list_comb (Syntax.const syn $ x $ A $ t, ts) end); |
42085 | 369 |
|
548 | 370 |
|
3691
f0396ac63e12
tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents:
2698
diff
changeset
|
371 |
(* idtyp constraints *) |
548 | 372 |
|
42045 | 373 |
fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant "_constrain", x, ty], xs] = |
374 |
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
|
375 |
| idtyp_ast_tr' _ _ = raise Match; |
548 | 376 |
|
377 |
||
378 |
(* meta implication *) |
|
379 |
||
42278 | 380 |
fun impl_ast_tr' asts = |
42262 | 381 |
if no_brackets () then raise Match |
10572 | 382 |
else |
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
383 |
(case Ast.unfold_ast_p "\<^const>Pure.imp" |
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
384 |
(Ast.Appl (Ast.Constant "\<^const>Pure.imp" :: asts)) of |
16612
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
385 |
(prems as _ :: _ :: _, concl) => |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
386 |
let |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
387 |
val (asms, asm) = split_last prems; |
48be8ef738df
transform_failure in translation functions: TRANSLATION_FAIL;
wenzelm
parents:
15570
diff
changeset
|
388 |
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
|
389 |
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
|
390 |
| _ => raise Match); |
fcf747c0b6b8
Syntax: last premise of "_bigimpl" is wrapped with "_asm", to have a hook for
schirmer
parents:
14981
diff
changeset
|
391 |
|
548 | 392 |
|
393 |
(* dependent / nondependent quantifiers *) |
|
394 |
||
20202 | 395 |
fun var_abs mark (x, T, b) = |
43326
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents:
42476
diff
changeset
|
396 |
let val (x', _) = Name.variant x (Term.declare_term_names b Name.context) |
20202 | 397 |
in (x', subst_bound (mark (x', T), b)) end; |
398 |
||
399 |
val variant_abs = var_abs Free; |
|
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
400 |
val variant_abs' = var_abs mark_bound_abs; |
2698 | 401 |
|
548 | 402 |
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
|
403 |
if Term.is_dependent B then |
2698 | 404 |
let val (x', B') = variant_abs' (x, dummyT, B); |
49660
de49d9b4d7bc
more explicit Syntax_Trans.mark_bound_abs/mark_bound_body: preserve type information for show_markup;
wenzelm
parents:
46236
diff
changeset
|
405 |
in Term.list_comb (Syntax.const q $ mark_bound_abs (x', T) $ A $ B', ts) end |
42476 | 406 |
else Term.list_comb (Syntax.const r $ A $ incr_boundvars ~1 B, ts) |
548 | 407 |
| dependent_tr' _ _ = raise Match; |
408 |
||
409 |
||
5084 | 410 |
(* quote / antiquote *) |
411 |
||
8577 | 412 |
fun antiquote_tr' name = |
413 |
let |
|
414 |
fun tr' i (t $ u) = |
|
42476 | 415 |
if u aconv Bound i then Syntax.const name $ tr' i t |
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
|
416 |
else tr' i t $ tr' i u |
8577 | 417 |
| tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t) |
18139 | 418 |
| tr' i a = if a aconv Bound i then raise Match else a; |
8577 | 419 |
in tr' 0 end; |
420 |
||
421 |
fun quote_tr' name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' name t) |
|
422 |
| quote_tr' _ _ = raise Match; |
|
423 |
||
5084 | 424 |
fun quote_antiquote_tr' quoteN antiquoteN name = |
425 |
let |
|
42476 | 426 |
fun tr' (t :: ts) = Term.list_comb (Syntax.const quoteN $ quote_tr' antiquoteN t, ts) |
8577 | 427 |
| tr' _ = raise Match; |
52143 | 428 |
in (name, fn _ => tr') end; |
5084 | 429 |
|
430 |
||
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
431 |
(* corresponding updates *) |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
432 |
|
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
|
433 |
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
|
434 |
|
58b465952287
update_name: more uniform treatment of type constraints (NB: type equality is hard to establish in parse trees);
wenzelm
parents:
42057
diff
changeset
|
435 |
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
|
436 |
| 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
|
437 |
|
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
438 |
fun upd_tr' (x_upd, T) = |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
439 |
(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
|
440 |
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
|
441 |
| NONE => raise Match); |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
442 |
|
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
|
443 |
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
|
444 |
|
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
445 |
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
|
446 |
| 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
|
447 |
| 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
|
448 |
| update_name_tr' _ = raise Match; |
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
449 |
|
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
|
450 |
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
|
451 |
|
35145
f132a4fd8679
moved generic update_name to Pure syntax -- not specific to HOL/record;
wenzelm
parents:
32786
diff
changeset
|
452 |
|
14697 | 453 |
(* indexed syntax *) |
548 | 454 |
|
42278 | 455 |
fun index_ast_tr' [Ast.Appl [Ast.Constant "_struct", ast]] = ast |
14697 | 456 |
| index_ast_tr' _ = raise Match; |
457 |
||
59152
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
458 |
fun struct_ast_tr' ctxt [Ast.Constant "_indexdefault"] = |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
459 |
(case #structs (get_idents ctxt) of |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
460 |
x :: _ => Ast.Appl [Ast.Constant "_free", Ast.Variable x] |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
461 |
| _ => raise Match) |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
462 |
| struct_ast_tr' _ _ = raise Match; |
14697 | 463 |
|
464 |
||
465 |
||
466 |
(** Pure translations **) |
|
548 | 467 |
|
52143 | 468 |
val pure_parse_ast_translation = |
469 |
[("_strip_positions", fn _ => strip_positions_ast_tr), |
|
470 |
("_constify", fn _ => constify_ast_tr), |
|
471 |
("_tapp", fn _ => tapp_ast_tr), |
|
472 |
("_tappl", fn _ => tappl_ast_tr), |
|
473 |
("_bracket", fn _ => bracket_ast_tr), |
|
474 |
("_appl", fn _ => appl_ast_tr), |
|
475 |
("_applC", fn _ => applC_ast_tr), |
|
476 |
("_lambda", fn _ => lambda_ast_tr), |
|
477 |
("_idtyp", fn _ => idtyp_ast_tr), |
|
478 |
("_bigimpl", fn _ => bigimpl_ast_tr), |
|
479 |
("_indexdefault", fn _ => indexdefault_ast_tr), |
|
480 |
("_indexvar", fn _ => indexvar_ast_tr), |
|
481 |
("_struct", fn _ => struct_ast_tr)]; |
|
548 | 482 |
|
52143 | 483 |
val pure_parse_translation = |
484 |
[("_abs", fn _ => abs_tr), |
|
485 |
("_aprop", fn _ => aprop_tr), |
|
486 |
("_ofclass", fn _ => ofclass_tr), |
|
487 |
("_sort_constraint", fn _ => sort_constraint_tr), |
|
488 |
("_TYPE", fn _ => type_tr), |
|
489 |
("_update_name", fn _ => update_name_tr), |
|
59152
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
490 |
("_index", fn _ => index_tr), |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
491 |
("_struct", struct_tr)]; |
52143 | 492 |
|
493 |
val pure_print_ast_translation = |
|
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
494 |
[("\<^type>fun", fn _ => fun_ast_tr'), |
52143 | 495 |
("_abs", fn _ => abs_ast_tr'), |
496 |
("_idts", fn _ => idtyp_ast_tr' "_idts"), |
|
497 |
("_pttrns", fn _ => idtyp_ast_tr' "_pttrns"), |
|
62529
8b7bdfc09f3b
clarified treatment of fragments of Isabelle symbols during bootstrap;
wenzelm
parents:
59152
diff
changeset
|
498 |
("\<^const>Pure.imp", fn _ => impl_ast_tr'), |
59152
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
499 |
("_index", fn _ => index_ast_tr'), |
66e6c539a36d
more frugal Local_Syntax.init -- maintain idents within context;
wenzelm
parents:
56438
diff
changeset
|
500 |
("_struct", struct_ast_tr')]; |
52143 | 501 |
|
548 | 502 |
end; |
42284 | 503 |
|
504 |
structure Basic_Syntax_Trans: BASIC_SYNTAX_TRANS = Syntax_Trans; |
|
505 |
open Basic_Syntax_Trans; |