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