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