src/Pure/term_subst.ML
author wenzelm
Fri, 10 Jul 2009 00:08:38 +0200
changeset 31980 c7c1d545007e
parent 31977 e03059ae2d82
child 32016 597b3b69b8d8
permissions -rw-r--r--
added some generic mapping combinators; share private exception SAME locally;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/term_subst.ML
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
     3
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
     4
Efficient type/term substitution -- avoids copying.
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
     5
*)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
     6
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
     7
signature TERM_SUBST =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
     8
sig
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
     9
  val map_atypsT_option: (typ -> typ option) -> typ -> typ option
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    10
  val map_atyps_option: (typ -> typ option) -> term -> term option
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    11
  val map_types_option: (typ -> typ option) -> term -> term option
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    12
  val map_aterms_option: (term -> term option) -> term -> term option
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    13
  val generalize: string list * string list -> int -> term -> term
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    14
  val generalizeT: string list -> int -> typ -> typ
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    15
  val generalize_option: string list * string list -> int -> term -> term option
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    16
  val generalizeT_option: string list -> int -> typ -> typ option
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    17
  val instantiateT_maxidx: ((indexname * sort) * (typ * int)) list -> typ -> int -> typ * int
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    18
  val instantiate_maxidx:
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    19
    ((indexname * sort) * (typ * int)) list * ((indexname * typ) * (term * int)) list ->
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    20
    term -> int -> term * int
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    21
  val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    22
    term -> term
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    23
  val instantiateT: ((indexname * sort) * typ) list -> typ -> typ
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    24
  val instantiate_option: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    25
    term -> term option
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    26
  val instantiateT_option: ((indexname * sort) * typ) list -> typ -> typ option
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    27
  val zero_var_indexes: term -> term
21607
3698319f6503 zero_var_indexes_inst: multiple terms;
wenzelm
parents: 21315
diff changeset
    28
  val zero_var_indexes_inst: term list ->
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    29
    ((indexname * sort) * typ) list * ((indexname * typ) * term) list
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    30
end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    31
31977
e03059ae2d82 renamed structure TermSubst to Term_Subst;
wenzelm
parents: 29269
diff changeset
    32
structure Term_Subst: TERM_SUBST =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    33
struct
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    34
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    35
(* indication of same result *)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    36
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    37
exception SAME;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    38
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    39
fun same_fn f x =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    40
  (case f x of
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    41
    NONE => raise SAME
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    42
  | SOME y => y);
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    43
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    44
fun option_fn f x =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    45
  SOME (f x) handle SAME => NONE;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    46
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    47
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    48
(* generic mapping *)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    49
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    50
local
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    51
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    52
fun map_atypsT_same f =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    53
  let
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    54
    fun typ (Type (a, Ts)) = Type (a, typs Ts)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    55
      | typ T = f T
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    56
    and typs (T :: Ts) = (typ T :: (typs Ts handle SAME => Ts) handle SAME => T :: typs Ts)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    57
      | typs [] = raise SAME;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    58
  in typ end;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    59
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    60
fun map_types_same f =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    61
  let
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    62
    fun term (Const (a, T)) = Const (a, f T)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    63
      | term (Free (a, T)) = Free (a, f T)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    64
      | term (Var (v, T)) = Var (v, f T)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    65
      | term (Bound _)  = raise SAME
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    66
      | term (Abs (x, T, t)) =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    67
          (Abs (x, f T, term t handle SAME => t)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    68
            handle SAME => Abs (x, T, term t))
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    69
      | term (t $ u) = (term t $ (term u handle SAME => u) handle SAME => t $ term u);
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    70
  in term end;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    71
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    72
fun map_aterms_same f =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    73
  let
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    74
    fun term (Abs (x, T, t)) = Abs (x, T, term t)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    75
      | term (t $ u) = (term t $ (term u handle SAME => u) handle SAME => t $ term u)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    76
      | term a = f a;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    77
  in term end;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    78
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    79
in
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    80
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    81
val map_atypsT_option = option_fn o map_atypsT_same o same_fn;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    82
val map_atyps_option = option_fn o map_types_same o map_atypsT_same o same_fn;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    83
val map_types_option = option_fn o map_types_same o same_fn;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    84
val map_aterms_option = option_fn o map_aterms_same o same_fn;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    85
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    86
end;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    87
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    88
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    89
(* generalization of fixed variables *)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    90
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    91
local
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    92
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    93
fun generalizeT_same [] _ _ = raise SAME
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    94
  | generalizeT_same tfrees idx ty =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    95
      let
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    96
        fun gen_typ (Type (a, Ts)) = Type (a, gen_typs Ts)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    97
          | gen_typ (TFree (a, S)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    98
              if member (op =) tfrees a then TVar ((a, idx), S)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    99
              else raise SAME
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   100
          | gen_typ _ = raise SAME
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   101
        and gen_typs (T :: Ts) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   102
            (gen_typ T :: (gen_typs Ts handle SAME => Ts)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   103
              handle SAME => T :: gen_typs Ts)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   104
          | gen_typs [] = raise SAME;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   105
      in gen_typ ty end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   106
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   107
fun generalize_same ([], []) _ _ = raise SAME
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   108
  | generalize_same (tfrees, frees) idx tm =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   109
      let
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   110
        val genT = generalizeT_same tfrees idx;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   111
        fun gen (Free (x, T)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   112
              if member (op =) frees x then
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   113
                Var (Name.clean_index (x, idx), genT T handle SAME => T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   114
              else Free (x, genT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   115
          | gen (Var (xi, T)) = Var (xi, genT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   116
          | gen (Const (c, T)) = Const (c, genT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   117
          | gen (Bound _) = raise SAME
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   118
          | gen (Abs (x, T, t)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   119
              (Abs (x, genT T, gen t handle SAME => t)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   120
                handle SAME => Abs (x, T, gen t))
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   121
          | gen (t $ u) = (gen t $ (gen u handle SAME => u) handle SAME => t $ gen u);
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   122
      in gen tm end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   123
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   124
in
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   125
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   126
fun generalize names i tm = generalize_same names i tm handle SAME => tm;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   127
fun generalizeT names i ty = generalizeT_same names i ty handle SAME => ty;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   128
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   129
fun generalize_option names i tm = SOME (generalize_same names i tm) handle SAME => NONE;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   130
fun generalizeT_option names i ty = SOME (generalizeT_same names i ty) handle SAME => NONE;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   131
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   132
end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   133
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   134
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   135
(* instantiation of schematic variables (types before terms) -- recomputes maxidx *)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   136
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   137
local
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   138
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   139
fun no_index (x, y) = (x, (y, ~1));
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   140
fun no_indexes1 inst = map no_index inst;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   141
fun no_indexes2 (inst1, inst2) = (map no_index inst1, map no_index inst2);
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   142
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   143
fun instantiateT_same maxidx instT ty =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   144
  let
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   145
    fun maxify i = if i > ! maxidx then maxidx := i else ();
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   146
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   147
    fun subst_typ (Type (a, Ts)) = Type (a, subst_typs Ts)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   148
      | subst_typ (TVar ((a, i), S)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   149
          (case AList.lookup Term.eq_tvar instT ((a, i), S) of
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   150
            SOME (T, j) => (maxify j; T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   151
          | NONE => (maxify i; raise SAME))
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   152
      | subst_typ _ = raise SAME
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   153
    and subst_typs (T :: Ts) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   154
        (subst_typ T :: (subst_typs Ts handle SAME => Ts)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   155
          handle SAME => T :: subst_typs Ts)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   156
      | subst_typs [] = raise SAME;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   157
  in subst_typ ty end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   158
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   159
fun instantiate_same maxidx (instT, inst) tm =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   160
  let
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   161
    fun maxify i = if i > ! maxidx then maxidx := i else ();
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   162
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   163
    val substT = instantiateT_same maxidx instT;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   164
    fun subst (Const (c, T)) = Const (c, substT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   165
      | subst (Free (x, T)) = Free (x, substT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   166
      | subst (Var ((x, i), T)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   167
          let val (T', same) = (substT T, false) handle SAME => (T, true) in
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   168
            (case AList.lookup Term.eq_var inst ((x, i), T') of
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   169
               SOME (t, j) => (maxify j; t)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   170
             | NONE => (maxify i; if same then raise SAME else Var ((x, i), T')))
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   171
          end
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   172
      | subst (Bound _) = raise SAME
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   173
      | subst (Abs (x, T, t)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   174
          (Abs (x, substT T, subst t handle SAME => t)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   175
            handle SAME => Abs (x, T, subst t))
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   176
      | subst (t $ u) = (subst t $ (subst u handle SAME => u) handle SAME => t $ subst u);
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   177
  in subst tm end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   178
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   179
in
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   180
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   181
fun instantiateT_maxidx instT ty i =
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   182
  let val maxidx = ref i
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   183
  in (instantiateT_same maxidx instT ty handle SAME => ty, ! maxidx) end;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   184
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   185
fun instantiate_maxidx insts tm i =
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   186
  let val maxidx = ref i
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   187
  in (instantiate_same maxidx insts tm handle SAME => tm, ! maxidx) end;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   188
21315
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   189
fun instantiateT [] ty = ty
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   190
  | instantiateT instT ty =
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   191
      (instantiateT_same (ref ~1) (no_indexes1 instT) ty handle SAME => ty);
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   192
21315
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   193
fun instantiate ([], []) tm = tm
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   194
  | instantiate insts tm =
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   195
      (instantiate_same (ref ~1) (no_indexes2 insts) tm handle SAME => tm);
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   196
21315
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   197
fun instantiateT_option [] _ = NONE
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   198
  | instantiateT_option instT ty =
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   199
      (SOME (instantiateT_same (ref ~1) (no_indexes1 instT) ty) handle SAME => NONE);
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   200
21315
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   201
fun instantiate_option ([], []) _ = NONE
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   202
  | instantiate_option insts tm =
be2669fe8363 instantiate: tuned indentity case;
wenzelm
parents: 21184
diff changeset
   203
      (SOME (instantiate_same (ref ~1) (no_indexes2 insts) tm) handle SAME => NONE);
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   204
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   205
end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   206
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   207
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   208
(* zero var indexes *)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   209
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   210
fun zero_var_inst vars =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   211
  fold (fn v as ((x, i), X) => fn (inst, used) =>
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   212
    let
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   213
      val ([x'], used') = Name.variants [if Name.is_bound x then "u" else x] used;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   214
    in if x = x' andalso i = 0 then (inst, used') else ((v, ((x', 0), X)) :: inst, used') end)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   215
  vars ([], Name.context) |> #1;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   216
21607
3698319f6503 zero_var_indexes_inst: multiple terms;
wenzelm
parents: 21315
diff changeset
   217
fun zero_var_indexes_inst ts =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   218
  let
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 21607
diff changeset
   219
    val tvars = sort TermOrd.tvar_ord (fold Term.add_tvars ts []);
21607
3698319f6503 zero_var_indexes_inst: multiple terms;
wenzelm
parents: 21315
diff changeset
   220
    val instT = map (apsnd TVar) (zero_var_inst tvars);
29269
5c25a2012975 moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
wenzelm
parents: 21607
diff changeset
   221
    val vars = sort TermOrd.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts []));
21607
3698319f6503 zero_var_indexes_inst: multiple terms;
wenzelm
parents: 21315
diff changeset
   222
    val inst = map (apsnd Var) (zero_var_inst vars);
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   223
  in (instT, inst) end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   224
21607
3698319f6503 zero_var_indexes_inst: multiple terms;
wenzelm
parents: 21315
diff changeset
   225
fun zero_var_indexes t = instantiate (zero_var_indexes_inst [t]) t;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   226
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   227
end;