src/Pure/Syntax/syn_trans.ML
author wenzelm
Fri, 17 Oct 1997 18:19:14 +0200
changeset 3928 787d2659ce4a
parent 3777 434d875f4661
child 4148 e0e5a2820ac1
permissions -rw-r--r--
no longer tries bogus eta-contract involving aprops;
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
    ID:         $Id$
12208de7edfe added this file;
wenzelm
parents:
diff changeset
     3
    Author:     Tobias Nipkow and Markus Wenzel, TU Muenchen
12208de7edfe added this file;
wenzelm
parents:
diff changeset
     4
12208de7edfe added this file;
wenzelm
parents:
diff changeset
     5
Syntax translation functions.
12208de7edfe added this file;
wenzelm
parents:
diff changeset
     6
*)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
     7
12208de7edfe added this file;
wenzelm
parents:
diff changeset
     8
signature SYN_TRANS0 =
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
     9
sig
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    10
  val eta_contract: bool ref
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    11
  val mk_binder_tr: string * string -> string * (term list -> term)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    12
  val mk_binder_tr': string * string -> string * (term list -> term)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    13
  val dependent_tr': string * string -> term list -> term
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    14
  val mark_bound: string -> term
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    15
  val mark_boundT: string * typ -> term
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    16
  val variant_abs': string * typ * term -> string * term
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    17
end;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    18
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    19
signature SYN_TRANS1 =
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    20
sig
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    21
  include SYN_TRANS0
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    22
  val constrainAbsC: string
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    23
  val pure_trfuns:
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    24
      (string * (Ast.ast list -> Ast.ast)) list *
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    25
      (string * (term list -> term)) list *
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    26
      (string * (term list -> term)) list *
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    27
      (string * (Ast.ast list -> Ast.ast)) list
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    28
  val pure_trfunsT: (string * (typ -> term list -> term)) list
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    29
end;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    30
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    31
signature SYN_TRANS =
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    32
sig
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    33
  include SYN_TRANS1
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    34
  val abs_tr': term -> term
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    35
  val prop_tr': bool -> term -> term
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    36
  val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    37
  val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    38
  val pt_to_ast: (string -> (Ast.ast list -> Ast.ast) option) -> Parser.parsetree -> Ast.ast
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    39
  val ast_to_term: (string -> (term list -> term) option) -> Ast.ast -> term
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    40
end;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    41
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    42
structure SynTrans: SYN_TRANS =
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    43
struct
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    44
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    45
open TypeExt Lexicon Ast SynExt Parser;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    46
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    47
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    48
(** parse (ast) translations **)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    49
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    50
(* application *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    51
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
    52
fun appl_ast_tr [f, args] = Appl (f :: unfold_ast "_args" args)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    53
  | appl_ast_tr asts = raise AST ("appl_ast_tr", asts);
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
    54
1178
b28c6ecc3e6d added cargs for curried function application
clasohm
parents: 987
diff changeset
    55
fun applC_ast_tr [f, args] = Appl (f :: unfold_ast "_cargs" args)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    56
  | applC_ast_tr asts = raise AST ("applC_ast_tr", asts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    57
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    58
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    59
(* abstraction *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    60
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    61
fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Appl [Constant constrainC, x, ty]
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    62
  | idtyp_ast_tr (*"_idtyp"*) asts = raise AST ("idtyp_ast_tr", asts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    63
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
    64
fun lambda_ast_tr (*"_lambda"*) [pats, body] =
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
    65
      fold_ast_p "_abs" (unfold_ast "_pttrns" pats, body)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    66
  | lambda_ast_tr (*"_lambda"*) asts = raise AST ("lambda_ast_tr", asts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    67
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    68
val constrainAbsC = "_constrainAbs";
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    69
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    70
fun abs_tr (*"_abs"*) [Free (x, T), body] = absfree (x, T, body)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    71
  | abs_tr (*"_abs"*) (ts as [Const (c, _) $ Free (x, T) $ tT, body]) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    72
      if c = constrainC
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    73
        then const constrainAbsC $ absfree (x, T, body) $ tT
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    74
      else raise TERM ("abs_tr", ts)
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    75
  | abs_tr (*"_abs"*) ts = raise TERM ("abs_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    76
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    77
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    78
(* nondependent abstraction *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    79
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    80
fun k_tr (*"_K"*) [t] = Abs ("uu", dummyT, incr_boundvars 1 t)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    81
  | k_tr (*"_K"*) ts = raise TERM ("k_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    82
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    83
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    84
(* binder *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    85
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    86
fun mk_binder_tr (sy, name) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    87
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    88
    fun tr (Free (x, T), t) = const name $ absfree (x, T, t)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    89
      | tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t))
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    90
      | tr (t1 as Const (c, _) $ Free (x, T) $ tT, t) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    91
          if c = constrainC then
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    92
            const name $ (const constrainAbsC $ absfree (x, T, t) $ tT)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    93
          else raise TERM ("binder_tr", [t1, t])
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    94
      | tr (t1, t2) = raise TERM ("binder_tr", [t1, t2]);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    95
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    96
    fun binder_tr (*sy*) [idts, body] = tr (idts, body)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    97
      | binder_tr (*sy*) ts = raise TERM ("binder_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    98
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    99
    (sy, binder_tr)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   100
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   101
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   102
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   103
(* meta propositions *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   104
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   105
fun aprop_tr (*"_aprop"*) [t] = const constrainC $ t $ const "prop"
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   106
  | aprop_tr (*"_aprop"*) ts = raise TERM ("aprop_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   107
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   108
fun ofclass_tr (*"_ofclass"*) [ty, cls] =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   109
      cls $ (const constrainC $ const "TYPE" $ (const "itself" $ ty))
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   110
  | ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   111
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   112
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   113
(* meta implication *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   114
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   115
fun bigimpl_ast_tr (*"_bigimpl"*) [asms, concl] =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   116
      fold_ast_p "==>" (unfold_ast "_asms" asms, concl)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   117
  | bigimpl_ast_tr (*"_bigimpl"*) asts = raise AST ("bigimpl_ast_tr", asts);
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
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   121
(** print (ast) translations **)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   122
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   123
(* application *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   124
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   125
fun appl_ast_tr' (f, []) = raise AST ("appl_ast_tr'", [f])
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   126
  | appl_ast_tr' (f, args) = Appl [Constant "_appl", f, fold_ast "_args" args];
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   127
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   128
fun applC_ast_tr' (f, []) = raise AST ("applC_ast_tr'", [f])
1178
b28c6ecc3e6d added cargs for curried function application
clasohm
parents: 987
diff changeset
   129
  | applC_ast_tr' (f, args) =
b28c6ecc3e6d added cargs for curried function application
clasohm
parents: 987
diff changeset
   130
      Appl [Constant "_applC", f, fold_ast "_cargs" args];
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
   131
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   132
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   133
(* abstraction *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   134
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   135
fun mark_boundT x_T = const "_bound" $ Free x_T;
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   136
fun mark_bound x = mark_boundT (x, dummyT);
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   137
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   138
fun strip_abss vars_of body_of tm =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   139
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   140
    val vars = vars_of tm;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   141
    val body = body_of tm;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   142
    val rev_new_vars = rename_wrt_term body vars;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   143
  in
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   144
    (map mark_boundT (rev rev_new_vars),
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   145
      subst_bounds (map (mark_bound o #1) rev_new_vars, body))
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   146
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   147
3928
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   148
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   149
(*do (partial) eta-contraction before printing*)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   150
1326
1fbf9407757c Set eta_contract to true.
nipkow
parents: 1178
diff changeset
   151
val eta_contract = ref true;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   152
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   153
fun eta_contr tm =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   154
  let
3928
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   155
    fun is_aprop (Const ("_aprop", _)) = true
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   156
      | is_aprop _ = false;
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   157
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   158
    fun eta_abs (Abs (a, T, t)) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   159
          (case eta_abs t of
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   160
            t' as f $ u =>
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   161
              (case eta_abs u of
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   162
                Bound 0 =>
3928
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   163
                  if 0 mem loose_bnos f orelse is_aprop f then Abs (a, T, t')
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   164
                  else  incr_boundvars ~1 f
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   165
              | _ => Abs (a, T, t'))
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   166
          | t' => Abs (a, T, t'))
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   167
      | eta_abs t = t;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   168
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   169
    if ! eta_contract then eta_abs tm else tm
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   170
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   171
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   172
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   173
fun abs_tr' tm =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   174
  foldr (fn (x, t) => const "_abs" $ x $ t)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   175
    (strip_abss strip_abs_vars strip_abs_body (eta_contr tm));
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   176
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   177
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   178
fun abs_ast_tr' (*"_abs"*) asts =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   179
  (case unfold_ast_p "_abs" (Appl (Constant "_abs" :: asts)) of
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   180
    ([], _) => raise AST ("abs_ast_tr'", asts)
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   181
  | (xs, body) => Appl [Constant "_lambda", fold_ast "_pttrns" xs, body]);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   182
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   183
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   184
(* binder *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   185
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   186
fun mk_binder_tr' (name, sy) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   187
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   188
    fun mk_idts [] = raise Match    (*abort translation*)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   189
      | mk_idts [idt] = idt
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   190
      | mk_idts (idt :: idts) = const "_idts" $ idt $ mk_idts idts;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   191
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   192
    fun tr' t =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   193
      let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   194
        val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   195
      in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   196
        const sy $ mk_idts xs $ bd
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   197
      end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   198
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   199
    fun binder_tr' (*name*) (t :: ts) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   200
          list_comb (tr' (const name $ t), ts)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   201
      | binder_tr' (*name*) [] = raise Match;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   202
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   203
    (name, binder_tr')
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   204
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   205
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   206
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   207
(* idtyp constraints *)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   208
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   209
fun idtyp_ast_tr' a [Appl [Constant c, x, ty], xs] =
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   210
      if c = constrainC then
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   211
        Appl [Constant a, Appl [Constant "_idtyp", x, ty], xs]
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   212
      else raise Match
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   213
  | idtyp_ast_tr' _ _ = raise Match;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   214
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   215
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   216
(* meta propositions *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   217
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   218
fun prop_tr' show_sorts tm =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   219
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   220
    fun aprop t = const "_aprop" $ t;
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   221
    val mk_ofclass = if show_sorts then "_mk_ofclassS" else "_mk_ofclass";
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   222
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   223
    fun is_prop Ts t =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   224
      fastype_of1 (Ts, t) = propT handle TERM _ => false;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   225
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   226
    fun tr' _ (t as Const _) = t
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   227
      | tr' _ (t as Free (x, T)) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   228
          if T = propT then aprop (free x) else t
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   229
      | tr' _ (t as Var (xi, T)) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   230
          if T = propT then aprop (var xi) else t
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   231
      | tr' Ts (t as Bound _) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   232
          if is_prop Ts t then aprop t else t
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   233
      | tr' Ts (Abs (x, T, t)) = Abs (x, T, tr' (T :: Ts) t)
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   234
      | tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   235
          if is_prop Ts t then Const (mk_ofclass, T) $ tr' Ts t1
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   236
          else tr' Ts t1 $ tr' Ts t2
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   237
      | tr' Ts (t as t1 $ t2) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   238
          (if is_Const (head_of t) orelse not (is_prop Ts t)
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   239
            then I else aprop) (tr' Ts t1 $ tr' Ts t2);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   240
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   241
    tr' [] tm
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   242
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   243
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   244
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   245
fun mk_ofclass_tr' (*"_mk_ofclass"*) T [t] =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   246
      const "_ofclass" $ term_of_typ false T $ t
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   247
  | mk_ofclass_tr' (*"_mk_ofclass"*) T ts = raise TYPE ("mk_ofclass_tr'", [T], ts);
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   248
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   249
fun mk_ofclassS_tr' (*"_mk_ofclassS"*) T [t] =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   250
      const "_ofclass" $ term_of_typ true T $ t
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   251
  | mk_ofclassS_tr' (*"_mk_ofclassS"*) T ts = raise TYPE ("mk_ofclassS_tr'", [T], ts);
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   252
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   253
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   254
(* meta implication *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   255
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   256
fun impl_ast_tr' (*"==>"*) asts =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   257
  (case unfold_ast_p "==>" (Appl (Constant "==>" :: asts)) of
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   258
    (asms as _ :: _ :: _, concl)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   259
      => Appl [Constant "_bigimpl", fold_ast "_asms" asms, concl]
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   260
  | _ => raise Match);
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   261
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   262
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   263
(* dependent / nondependent quantifiers *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   264
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   265
fun variant_abs' (x, T, B) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   266
  let val x' = variant (add_term_names (B, [])) x in
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   267
    (x', subst_bound (mark_boundT (x', T), B))
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   268
  end;
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   269
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   270
fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   271
      if 0 mem (loose_bnos B) then
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   272
        let val (x', B') = variant_abs' (x, dummyT, B);
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   273
        in list_comb (const q $ mark_boundT (x', T) $ A $ B', ts) end
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   274
      else list_comb (const r $ A $ B, ts)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   275
  | dependent_tr' _ _ = raise Match;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   276
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   277
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   278
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   279
(** pure_trfuns **)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   280
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   281
val pure_trfuns =
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
   282
 ([("_appl", appl_ast_tr), ("_applC", applC_ast_tr),
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
   283
   ("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr),
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
   284
   ("_bigimpl", bigimpl_ast_tr)],
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
   285
  [("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr),
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
   286
   ("_K", k_tr)],
3700
3a8192e83579 pure_trfuns: added constraint;
wenzelm
parents: 3691
diff changeset
   287
  []: (string * (term list -> term)) list,
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   288
  [("_abs", abs_ast_tr'), ("_idts", idtyp_ast_tr' "_idts"),
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   289
   ("_pttrns", idtyp_ast_tr' "_pttrns"), ("==>", impl_ast_tr')]);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   290
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   291
val pure_trfunsT =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   292
  [("_mk_ofclass", mk_ofclass_tr'), ("_mk_ofclassS", mk_ofclassS_tr')];
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   293
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   294
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   295
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   296
(** pt_to_ast **)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   297
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   298
fun pt_to_ast trf pt =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   299
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   300
    fun trans a args =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   301
      (case trf a of
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   302
        None => mk_appl (Constant a) args
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   303
      | Some f => f args handle exn
987
32bb5a8d5aab changed translation of _applC
clasohm
parents: 922
diff changeset
   304
          => (writeln ("Error in parse ast translation for " ^ quote a);
32bb5a8d5aab changed translation of _applC
clasohm
parents: 922
diff changeset
   305
              raise exn));
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   306
987
32bb5a8d5aab changed translation of _applC
clasohm
parents: 922
diff changeset
   307
    (*translate pt bottom-up*)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   308
    fun ast_of (Node (a, pts)) = trans a (map ast_of pts)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   309
      | ast_of (Tip tok) = Variable (str_of_token tok);
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   310
  in
1178
b28c6ecc3e6d added cargs for curried function application
clasohm
parents: 987
diff changeset
   311
    ast_of pt
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   312
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   313
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   314
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   315
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   316
(** ast_to_term **)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   317
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   318
fun ast_to_term trf ast =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   319
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   320
    fun trans a args =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   321
      (case trf a of
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   322
        None => list_comb (const a, args)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   323
      | Some f => f args handle exn
987
32bb5a8d5aab changed translation of _applC
clasohm
parents: 922
diff changeset
   324
          => (writeln ("Error in parse translation for " ^ quote a);
32bb5a8d5aab changed translation of _applC
clasohm
parents: 922
diff changeset
   325
              raise exn));
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   326
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   327
    fun term_of (Constant a) = trans a []
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   328
      | term_of (Variable x) = scan_var x
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   329
      | term_of (Appl (Constant a :: (asts as _ :: _))) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   330
          trans a (map term_of asts)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   331
      | term_of (Appl (ast :: (asts as _ :: _))) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   332
          list_comb (term_of ast, map term_of asts)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   333
      | term_of (ast as Appl _) = raise AST ("ast_to_term: malformed ast", [ast]);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   334
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   335
    term_of ast
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   336
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   337
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   338
end;