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