src/Pure/Syntax/syn_trans.ML
author kleing
Mon, 21 Jun 2004 10:25:57 +0200
changeset 14981 e73f8140af78
parent 14868 617e4b19a2e5
child 15421 fcf747c0b6b8
permissions -rw-r--r--
Merged in license change from Isabelle2004
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
13762
9dd78dab72bc *** empty log message ***
nipkow
parents: 12785
diff changeset
    11
  val atomic_abs_tr': string * typ * term -> term * term
548
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 mk_binder_tr': string * string -> string * (term list -> term)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    14
  val dependent_tr': string * string -> term list -> term
8577
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
    15
  val antiquote_tr: string -> term -> term
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
    16
  val quote_tr: string -> term -> term
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
    17
  val quote_antiquote_tr: string -> string -> string -> string * (term list -> term)
8577
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
    18
  val antiquote_tr': string -> term -> term
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
    19
  val quote_tr': string -> term -> term
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
    20
  val quote_antiquote_tr': string -> string -> string -> string * (term list -> term)
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    21
  val mark_bound: string -> term
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    22
  val mark_boundT: string * typ -> term
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    23
  val variant_abs': string * typ * term -> string * term
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    24
end;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    25
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    26
signature SYN_TRANS1 =
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    27
sig
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    28
  include SYN_TRANS0
14647
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
    29
  val non_typed_tr': (term list -> term) -> bool -> typ -> term list -> term
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
    30
  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
    31
  val constrainAbsC: string
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    32
  val pure_trfuns:
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    33
      (string * (Ast.ast list -> Ast.ast)) list *
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    34
      (string * (term list -> term)) list *
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    35
      (string * (term list -> term)) list *
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    36
      (string * (Ast.ast list -> Ast.ast)) list
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
    37
  val pure_trfunsT: (string * (bool -> typ -> term list -> term)) list
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
    38
  val struct_trfuns: string list ->
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
    39
      (string * (Ast.ast list -> Ast.ast)) list *
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
    40
      (string * (term list -> term)) list *
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
    41
      (string * (bool -> typ -> term list -> term)) list *
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
    42
      (string * (Ast.ast list -> Ast.ast)) list
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    43
end;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    44
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    45
signature SYN_TRANS =
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    46
sig
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    47
  include SYN_TRANS1
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    48
  val abs_tr': term -> term
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
    49
  val prop_tr': term -> term
1511
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    50
  val appl_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
09354d37a5ab Elimination of fully-functorial style.
paulson
parents: 1326
diff changeset
    51
  val applC_ast_tr': Ast.ast * Ast.ast list -> Ast.ast
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
    52
  val pts_to_asts: (string -> (Ast.ast list -> Ast.ast) option) -> Parser.parsetree list -> Ast.ast list
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
    53
  val asts_to_terms: (string -> (term list -> term) option) -> Ast.ast list -> term list
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    54
end;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    55
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    56
structure SynTrans: SYN_TRANS =
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    57
struct
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    58
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
    59
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    60
(** parse (ast) translations **)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    61
11491
085a0d2857e8 added constify_ast_tr;
wenzelm
parents: 10572
diff changeset
    62
(* constify *)
085a0d2857e8 added constify_ast_tr;
wenzelm
parents: 10572
diff changeset
    63
085a0d2857e8 added constify_ast_tr;
wenzelm
parents: 10572
diff changeset
    64
fun constify_ast_tr [Ast.Variable c] = Ast.Constant c
085a0d2857e8 added constify_ast_tr;
wenzelm
parents: 10572
diff changeset
    65
  | constify_ast_tr asts = raise Ast.AST ("constify_ast_tr", asts);
085a0d2857e8 added constify_ast_tr;
wenzelm
parents: 10572
diff changeset
    66
085a0d2857e8 added constify_ast_tr;
wenzelm
parents: 10572
diff changeset
    67
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    68
(* application *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    69
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    70
fun appl_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_args" args)
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    71
  | 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
    72
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    73
fun applC_ast_tr [f, args] = Ast.Appl (f :: Ast.unfold_ast "_cargs" args)
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    74
  | applC_ast_tr asts = raise Ast.AST ("applC_ast_tr", asts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    75
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    76
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    77
(* abstraction *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    78
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    79
fun idtyp_ast_tr (*"_idtyp"*) [x, ty] = Ast.Appl [Ast.Constant SynExt.constrainC, x, ty]
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    80
  | idtyp_ast_tr (*"_idtyp"*) asts = raise Ast.AST ("idtyp_ast_tr", asts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    81
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
    82
fun lambda_ast_tr (*"_lambda"*) [pats, body] =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    83
      Ast.fold_ast_p "_abs" (Ast.unfold_ast "_pttrns" pats, body)
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    84
  | lambda_ast_tr (*"_lambda"*) asts = raise Ast.AST ("lambda_ast_tr", asts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    85
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    86
val constrainAbsC = "_constrainAbs";
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    87
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    88
fun abs_tr (*"_abs"*) [Free (x, T), body] = Term.absfree (x, T, body)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    89
  | abs_tr (*"_abs"*) (ts as [Const (c, _) $ Free (x, T) $ tT, body]) =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    90
      if c = SynExt.constrainC
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    91
        then Lexicon.const constrainAbsC $ Term.absfree (x, T, body) $ tT
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    92
      else raise TERM ("abs_tr", ts)
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    93
  | abs_tr (*"_abs"*) ts = raise TERM ("abs_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    94
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    95
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    96
(* nondependent abstraction *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
    97
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
    98
fun k_tr (*"_K"*) [t] = Abs ("uu", dummyT, Term.incr_boundvars 1 t)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
    99
  | k_tr (*"_K"*) ts = raise TERM ("k_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   100
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   101
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   102
(* binder *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   103
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   104
fun mk_binder_tr (sy, name) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   105
  let
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   106
    fun tr (Free (x, T), t) = Lexicon.const name $ Term.absfree (x, T, t)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   107
      | tr (Const ("_idts", _) $ idt $ idts, t) = tr (idt, tr (idts, t))
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   108
      | tr (t1 as Const (c, _) $ Free (x, T) $ tT, t) =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   109
          if c = SynExt.constrainC then
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   110
            Lexicon.const name $ (Lexicon.const constrainAbsC $ Term.absfree (x, T, t) $ tT)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   111
          else raise TERM ("binder_tr", [t1, t])
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   112
      | tr (t1, t2) = raise TERM ("binder_tr", [t1, t2]);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   113
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   114
    fun binder_tr (*sy*) [idts, body] = tr (idts, body)
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   115
      | binder_tr (*sy*) ts = raise TERM ("binder_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   116
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   117
    (sy, binder_tr)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   118
  end;
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
(* meta propositions *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   122
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   123
fun aprop_tr (*"_aprop"*) [t] = Lexicon.const SynExt.constrainC $ t $ Lexicon.const "prop"
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   124
  | aprop_tr (*"_aprop"*) ts = raise TERM ("aprop_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   125
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   126
fun ofclass_tr (*"_ofclass"*) [ty, cls] =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   127
      cls $ (Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   128
        (Lexicon.const "itself" $ ty))
3777
434d875f4661 eliminated raise_ast, raise_term, raise_typ;
wenzelm
parents: 3700
diff changeset
   129
  | ofclass_tr (*"_ofclass"*) ts = raise TERM ("ofclass_tr", ts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   130
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   131
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   132
(* meta implication *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   133
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   134
fun bigimpl_ast_tr (*"_bigimpl"*) [asms, concl] =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   135
      Ast.fold_ast_p "==>" (Ast.unfold_ast "_asms" asms, concl)
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   136
  | bigimpl_ast_tr (*"_bigimpl"*) asts = raise Ast.AST ("bigimpl_ast_tr", asts);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   137
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   138
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   139
(* type reflection *)
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   140
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   141
fun type_tr (*"_TYPE"*) [ty] =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   142
      Lexicon.const SynExt.constrainC $ Lexicon.const "TYPE" $ (Lexicon.const "itself" $ ty)
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   143
  | type_tr (*"_TYPE"*) ts = raise TERM ("type_tr", ts);
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   144
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   145
6761
aa71a04f4b93 added dddot_tr;
wenzelm
parents: 5690
diff changeset
   146
(* dddot *)
aa71a04f4b93 added dddot_tr;
wenzelm
parents: 5690
diff changeset
   147
8595
06874c5c3cfa fixed dddot_tr;
wenzelm
parents: 8577
diff changeset
   148
fun dddot_tr (*"_DDDOT"*) ts = Term.list_comb (Lexicon.var SynExt.dddot_indexname, ts);
6761
aa71a04f4b93 added dddot_tr;
wenzelm
parents: 5690
diff changeset
   149
aa71a04f4b93 added dddot_tr;
wenzelm
parents: 5690
diff changeset
   150
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   151
(* quote / antiquote *)
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   152
8577
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   153
fun antiquote_tr name =
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   154
  let
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   155
    fun tr i ((t as Const (c, _)) $ u) =
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   156
          if c = name then tr i u $ Bound i
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   157
          else tr i t $ tr i u
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   158
      | tr i (t $ u) = tr i t $ tr i u
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   159
      | tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t)
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   160
      | tr _ a = a;
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   161
  in tr 0 end;
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   162
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   163
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
   164
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   165
fun quote_antiquote_tr quoteN antiquoteN name =
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   166
  let
8577
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   167
    fun tr [t] = Lexicon.const name $ quote_tr antiquoteN t
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   168
      | tr ts = raise TERM ("quote_tr", ts);
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   169
  in (quoteN, tr) end;
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   170
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   171
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   172
(* indexed syntax *)
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   173
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   174
fun struct_ast_tr (*"_struct"*) [Ast.Appl [Ast.Constant "_index", ast]] = ast
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   175
  | struct_ast_tr (*"_struct"*) asts = Ast.mk_appl (Ast.Constant "_struct") asts;
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   176
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   177
fun index_ast_tr ast =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   178
  Ast.mk_appl (Ast.Constant "_index") [Ast.mk_appl (Ast.Constant "_struct") [ast]];
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   179
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   180
fun indexdefault_ast_tr (*"_indexdefault"*) [] =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   181
      index_ast_tr (Ast.Constant "_indexdefault")
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   182
  | indexdefault_ast_tr (*"_indexdefault"*) asts =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   183
      raise Ast.AST ("indexdefault_ast_tr", asts);
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   184
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   185
fun indexnum_ast_tr (*"_indexnum"*) [ast] =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   186
      index_ast_tr (Ast.mk_appl (Ast.Constant "_indexnum") [ast])
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   187
  | 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
   188
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   189
fun indexvar_ast_tr (*"_indexvar"*) [] =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   190
      Ast.mk_appl (Ast.Constant "_index") [Ast.Variable "some_index"]
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   191
  | indexvar_ast_tr (*"_indexvar"*) asts = raise Ast.AST ("indexvar_ast_tr", asts);
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   192
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   193
fun index_tr (*"_index"*) [t] = t
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   194
  | index_tr (*"_index"*) ts = raise TERM ("index_tr", ts);
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   195
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   196
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   197
(* implicit structures *)
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   198
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   199
fun the_struct structs i =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   200
  if 1 <= i andalso i <= length structs then Library.nth_elem (i - 1, structs)
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   201
  else raise ERROR_MESSAGE ("Illegal reference to implicit structure #" ^ string_of_int i);
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   202
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   203
fun struct_tr structs (*"_struct"*) [Const ("_indexdefault", _)] =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   204
      Lexicon.free (the_struct structs 1)
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   205
  | struct_tr structs (*"_struct"*) [t as (Const ("_indexnum", _) $ Const (s, _))] =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   206
      Lexicon.free (the_struct structs
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   207
        (case Lexicon.read_nat s of Some n => n | None => raise TERM ("struct_tr", [t])))
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   208
  | struct_tr _ (*"_struct"*) ts = raise TERM ("struct_tr", ts);
12122
7f8d88ed4f21 indexvar_ast_tr (for \<index> placeholder);
wenzelm
parents: 11491
diff changeset
   209
7f8d88ed4f21 indexvar_ast_tr (for \<index> placeholder);
wenzelm
parents: 11491
diff changeset
   210
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   211
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   212
(** print (ast) translations **)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   213
14647
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
   214
(* types *)
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
   215
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
   216
fun non_typed_tr' f _ _ ts = f ts;
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
   217
fun non_typed_tr'' f x _ _ ts = f x ts;
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
   218
3f9d3d5cd0cd non_typed_tr';
wenzelm
parents: 13762
diff changeset
   219
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   220
(* application *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   221
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   222
fun appl_ast_tr' (f, []) = raise Ast.AST ("appl_ast_tr'", [f])
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   223
  | 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
   224
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   225
fun applC_ast_tr' (f, []) = raise Ast.AST ("applC_ast_tr'", [f])
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   226
  | 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
   227
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   228
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   229
(* abstraction *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   230
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   231
fun mark_boundT x_T = Lexicon.const "_bound" $ Free x_T;
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   232
fun mark_bound x = mark_boundT (x, dummyT);
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   233
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   234
fun strip_abss vars_of body_of tm =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   235
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   236
    val vars = vars_of tm;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   237
    val body = body_of tm;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   238
    val rev_new_vars = rename_wrt_term body vars;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   239
  in
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   240
    (map mark_boundT (rev rev_new_vars),
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   241
      subst_bounds (map (mark_bound o #1) rev_new_vars, body))
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   242
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   243
3928
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   244
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   245
(*do (partial) eta-contraction before printing*)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   246
1326
1fbf9407757c Set eta_contract to true.
nipkow
parents: 1178
diff changeset
   247
val eta_contract = ref true;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   248
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   249
fun eta_contr tm =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   250
  let
3928
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   251
    fun is_aprop (Const ("_aprop", _)) = true
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   252
      | is_aprop _ = false;
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   253
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   254
    fun eta_abs (Abs (a, T, t)) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   255
          (case eta_abs t of
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   256
            t' as f $ u =>
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   257
              (case eta_abs u of
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   258
                Bound 0 =>
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   259
                  if Term.loose_bvar1 (f, 0) orelse is_aprop f then Abs (a, T, t')
3928
787d2659ce4a no longer tries bogus eta-contract involving aprops;
wenzelm
parents: 3777
diff changeset
   260
                  else  incr_boundvars ~1 f
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   261
              | _ => Abs (a, T, t'))
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   262
          | t' => Abs (a, T, t'))
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   263
      | eta_abs t = t;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   264
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   265
    if ! eta_contract then eta_abs tm else tm
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   266
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   267
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   268
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   269
fun abs_tr' tm =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   270
  foldr (fn (x, t) => Lexicon.const "_abs" $ x $ t)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   271
    (strip_abss strip_abs_vars strip_abs_body (eta_contr tm));
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   272
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   273
fun atomic_abs_tr' (x, T, t) =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   274
  let val [xT] = rename_wrt_term t [(x, T)]
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   275
  in (mark_boundT xT, subst_bound (mark_bound (fst xT), t)) end;
13762
9dd78dab72bc *** empty log message ***
nipkow
parents: 12785
diff changeset
   276
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   277
fun abs_ast_tr' (*"_abs"*) asts =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   278
  (case Ast.unfold_ast_p "_abs" (Ast.Appl (Ast.Constant "_abs" :: asts)) of
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   279
    ([], _) => raise Ast.AST ("abs_ast_tr'", asts)
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   280
  | (xs, body) => Ast.Appl [Ast.Constant "_lambda", Ast.fold_ast "_pttrns" xs, body]);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   281
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   282
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   283
(* binder *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   284
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   285
fun mk_binder_tr' (name, sy) =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   286
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   287
    fun mk_idts [] = raise Match    (*abort translation*)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   288
      | mk_idts [idt] = idt
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   289
      | mk_idts (idt :: idts) = Lexicon.const "_idts" $ idt $ mk_idts idts;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   290
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   291
    fun tr' t =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   292
      let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   293
        val (xs, bd) = strip_abss (strip_qnt_vars name) (strip_qnt_body name) t;
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   294
      in Lexicon.const sy $ mk_idts xs $ bd end;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   295
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   296
    fun binder_tr' (*name*) (t :: ts) = Term.list_comb (tr' (Lexicon.const name $ t), ts)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   297
      | binder_tr' (*name*) [] = raise Match;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   298
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   299
    (name, binder_tr')
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   300
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   301
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   302
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   303
(* idtyp constraints *)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   304
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   305
fun idtyp_ast_tr' a [Ast.Appl [Ast.Constant c, x, ty], xs] =
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   306
      if c = SynExt.constrainC then
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   307
        Ast.Appl [ Ast.Constant a,  Ast.Appl [Ast.Constant "_idtyp", x, ty], xs]
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   308
      else raise Match
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   309
  | idtyp_ast_tr' _ _ = raise Match;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   310
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   311
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   312
(* meta propositions *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   313
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   314
fun prop_tr' tm =
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   315
  let
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   316
    fun aprop t = Lexicon.const "_aprop" $ t;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   317
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   318
    fun is_prop Ts t =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   319
      fastype_of1 (Ts, t) = propT handle TERM _ => false;
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   320
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   321
    fun tr' _ (t as Const _) = t
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   322
      | tr' _ (t as Free (x, T)) =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   323
          if T = propT then aprop (Lexicon.free x) else t
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   324
      | tr' _ (t as Var (xi, T)) =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   325
          if T = propT then aprop (Lexicon.var xi) else t
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   326
      | tr' Ts (t as Bound _) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   327
          if is_prop Ts t then aprop t else t
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   328
      | 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
   329
      | tr' Ts (t as t1 $ (t2 as Const ("TYPE", Type ("itself", [T])))) =
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   330
          if is_prop Ts t then Const ("_mk_ofclass", T) $ tr' Ts t1
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   331
          else tr' Ts t1 $ tr' Ts t2
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   332
      | tr' Ts (t as t1 $ t2) =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   333
          (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
   334
            then I else aprop) (tr' Ts t1 $ tr' Ts t2);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   335
  in
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   336
    tr' [] tm
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   337
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   338
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   339
fun mk_ofclass_tr' show_sorts (*"_mk_ofclass"*) T [t] =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   340
      Lexicon.const "_ofclass" $ TypeExt.term_of_typ show_sorts T $ t
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   341
  | 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
   342
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   343
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   344
(* meta implication *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   345
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   346
fun impl_ast_tr' (*"==>"*) asts =
10572
b070825579b8 no_brackets mode;
wenzelm
parents: 8595
diff changeset
   347
  if TypeExt.no_brackets () then raise Match
b070825579b8 no_brackets mode;
wenzelm
parents: 8595
diff changeset
   348
  else
b070825579b8 no_brackets mode;
wenzelm
parents: 8595
diff changeset
   349
    (case Ast.unfold_ast_p "==>" (Ast.Appl (Ast.Constant "==>" :: asts)) of
b070825579b8 no_brackets mode;
wenzelm
parents: 8595
diff changeset
   350
      (asms as _ :: _ :: _, concl)
b070825579b8 no_brackets mode;
wenzelm
parents: 8595
diff changeset
   351
        => Ast.Appl [Ast.Constant "_bigimpl", Ast.fold_ast "_asms" asms, concl]
b070825579b8 no_brackets mode;
wenzelm
parents: 8595
diff changeset
   352
    | _ => raise Match);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   353
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   354
12150
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   355
(* meta conjunction *)
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   356
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   357
fun meta_conjunction_tr' (*"all"*)
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   358
      [Abs (_, _, Const ("==>", _) $
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   359
        (Const ("==>", _) $ A $ (Const ("==>", _) $ B $ (Const ("_aprop", _) $ Bound 0))) $
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   360
        (Const ("_aprop", _) $ Bound 0))] =
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   361
      if 0 mem_int Term.loose_bnos A orelse 0 mem_int Term.loose_bnos B then raise Match
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   362
      else Lexicon.const "_meta_conjunction" $ A $ B
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   363
  | meta_conjunction_tr' (*"all"*) ts = raise Match;
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   364
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   365
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   366
(* type reflection *)
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   367
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   368
fun type_tr' show_sorts (*"TYPE"*) (Type ("itself", [T])) ts =
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   369
      Term.list_comb (Lexicon.const "_TYPE" $ TypeExt.term_of_typ show_sorts T, ts)
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   370
  | type_tr' _ _ _ = raise Match;
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   371
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   372
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   373
(* dependent / nondependent quantifiers *)
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   374
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   375
fun variant_abs' (x, T, B) =
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   376
  let val x' = variant (add_term_names (B, [])) x in
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   377
    (x', subst_bound (mark_boundT (x', T), B))
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   378
  end;
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   379
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   380
fun dependent_tr' (q, r) (A :: Abs (x, T, B) :: ts) =
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   381
      if Term.loose_bvar1 (B, 0) then
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   382
        let val (x', B') = variant_abs' (x, dummyT, B);
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   383
        in Term.list_comb (Lexicon.const q $ mark_boundT (x', T) $ A $ B', ts) end
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   384
      else Term.list_comb (Lexicon.const r $ A $ B, ts)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   385
  | dependent_tr' _ _ = raise Match;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   386
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   387
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   388
(* quote / antiquote *)
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   389
8577
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   390
fun antiquote_tr' name =
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   391
  let
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   392
    fun tr' i (t $ u) =
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   393
      if u = Bound i then Lexicon.const name $ tr' i t
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   394
      else tr' i t $ tr' i u
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   395
      | tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t)
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   396
      | tr' i a = if a = Bound i then raise Match else a;
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   397
  in tr' 0 end;
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   398
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   399
fun quote_tr' name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' name t)
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   400
  | quote_tr' _ _ = raise Match;
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   401
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   402
fun quote_antiquote_tr' quoteN antiquoteN name =
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   403
  let
8577
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   404
    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
   405
      | tr' _ = raise Match;
4a3cdf01531b improved (anti)quote_tr(');
wenzelm
parents: 8575
diff changeset
   406
  in (name, tr') end;
5084
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   407
a676ada3b380 tuned loose bound vars check;
wenzelm
parents: 4700
diff changeset
   408
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   409
(* indexed syntax *)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   410
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   411
fun index_ast_tr' (*"_index"*) [Ast.Appl [Ast.Constant "_struct", ast]] = ast
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   412
  | index_ast_tr' _ = raise Match;
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   413
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   414
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   415
(* implicit structures *)
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   416
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   417
fun the_struct' structs s =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   418
  [(case Lexicon.read_nat s of
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   419
    Some i => Ast.Variable (the_struct structs i handle ERROR_MESSAGE _ => raise Match)
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   420
  | None => raise Match)] |> Ast.mk_appl (Ast.Constant "_free");
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   421
;
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   422
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   423
fun struct_ast_tr' structs (*"_struct"*) [Ast.Constant "_indexdefault"] =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   424
      the_struct' structs "1"
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   425
  | struct_ast_tr' structs (*"_struct"*) [Ast.Appl [Ast.Constant "_indexnum", Ast.Constant s]] =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   426
      the_struct' structs s
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   427
  | struct_ast_tr' _ _ = raise Match;
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   428
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   429
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   430
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   431
(** Pure translations **)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   432
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   433
val pure_trfuns =
11491
085a0d2857e8 added constify_ast_tr;
wenzelm
parents: 10572
diff changeset
   434
 ([("_constify", constify_ast_tr), ("_appl", appl_ast_tr), ("_applC", applC_ast_tr),
12122
7f8d88ed4f21 indexvar_ast_tr (for \<index> placeholder);
wenzelm
parents: 11491
diff changeset
   435
   ("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr), ("_bigimpl", bigimpl_ast_tr),
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   436
   ("_indexdefault", indexdefault_ast_tr), ("_indexnum", indexnum_ast_tr),
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   437
   ("_indexvar", indexvar_ast_tr), ("_struct", struct_ast_tr)],
922
196ca0973a6d added CPure (curried functions) and ProtoPure (ancestor of Pure and CPure)
clasohm
parents: 639
diff changeset
   438
  [("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr),
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   439
   ("_TYPE", type_tr), ("_DDDOT", dddot_tr), ("_K", k_tr),
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   440
   ("_index", index_tr)],
12150
f83dc4202b78 added meta_conjunction_tr';
wenzelm
parents: 12122
diff changeset
   441
  [("all", meta_conjunction_tr')],
3691
f0396ac63e12 tuned lambda_ast_tr, idtyp_ast_tr' to accomodate fix of idt/idts
wenzelm
parents: 2698
diff changeset
   442
  [("_abs", abs_ast_tr'), ("_idts", idtyp_ast_tr' "_idts"),
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   443
   ("_pttrns", idtyp_ast_tr' "_pttrns"), ("==>", impl_ast_tr'),
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   444
   ("_index", index_ast_tr')]);
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   445
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   446
val pure_trfunsT =
4148
e0e5a2820ac1 adapted pure_trfunsT;
wenzelm
parents: 3928
diff changeset
   447
  [("_mk_ofclass", mk_ofclass_tr'), ("TYPE", type_tr')];
2698
8bccb3ab4ca4 added mark_bound(T), variant_abs';
wenzelm
parents: 1511
diff changeset
   448
14697
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   449
fun struct_trfuns structs =
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   450
  ([], [("_struct", struct_tr structs)], [], [("_struct", struct_ast_tr' structs)]);
5ad13d05049b improved indexed syntax / implicit structures;
wenzelm
parents: 14647
diff changeset
   451
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   452
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   453
14868
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   454
(** pts_to_asts **)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   455
14868
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   456
exception TRANSLATION of string * exn;
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   457
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   458
fun pts_to_asts trf pts =
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   459
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   460
    fun trans a args =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   461
      (case trf a of
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   462
        None => Ast.mk_appl (Ast.Constant a) args
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   463
      | Some f => f args handle exn => raise TRANSLATION (a, exn));
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   464
987
32bb5a8d5aab changed translation of _applC
clasohm
parents: 922
diff changeset
   465
    (*translate pt bottom-up*)
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   466
    fun ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts)
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   467
      | ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok);
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   468
14868
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   469
    val exn_results = map (capture ast_of) pts;
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   470
    val exns = mapfilter get_exn exn_results;
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   471
    val results = mapfilter get_result exn_results
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   472
  in
14868
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   473
    (case (results, exns) of
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   474
      ([], TRANSLATION (a, exn) :: _) =>
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   475
        (writeln ("Error in parse ast translation for " ^ quote a); raise exn)
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   476
    | _ => results)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   477
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   478
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   479
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   480
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   481
(** asts_to_terms **)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   482
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   483
fun asts_to_terms trf asts =
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   484
  let
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   485
    fun trans a args =
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   486
      (case trf a of
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   487
        None => Term.list_comb (Lexicon.const a, args)
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   488
      | Some f => f args handle exn => raise TRANSLATION (a, exn));
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   489
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   490
    fun term_of (Ast.Constant a) = trans a []
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   491
      | term_of (Ast.Variable x) = Lexicon.read_var x
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   492
      | term_of (Ast.Appl (Ast.Constant a :: (asts as _ :: _))) =
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   493
          trans a (map term_of asts)
5690
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   494
      | term_of (Ast.Appl (ast :: (asts as _ :: _))) =
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   495
          Term.list_comb (term_of ast, map term_of asts)
4b056ee5435c no open;
wenzelm
parents: 5288
diff changeset
   496
      | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]);
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   497
14868
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   498
    val exn_results = map (capture term_of) asts;
14798
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   499
    val exns = mapfilter get_exn exn_results;
702cb4859cab Modified functions pt_to_ast and ast_to_term to improve handling
berghofe
parents: 14697
diff changeset
   500
    val results = mapfilter get_result exn_results
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   501
  in
14868
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   502
    (case (results, exns) of
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   503
      ([], TRANSLATION (a, exn) :: _) =>
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   504
        (writeln ("Error in parse translation for " ^ quote a); raise exn)
617e4b19a2e5 tuned exeption handling (capture/release);
wenzelm
parents: 14798
diff changeset
   505
    | _ => results)
548
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   506
  end;
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   507
12208de7edfe added this file;
wenzelm
parents:
diff changeset
   508
end;