src/Pure/term_subst.ML
author wenzelm
Wed, 22 Apr 2020 19:22:43 +0200
changeset 71787 acfe72ff00c2
parent 69023 cef000855cf4
child 74200 17090e27aae9
permissions -rw-r--r--
merged
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
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
     4
Efficient type/term substitution.
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
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
     9
  val map_atypsT_same: typ Same.operation -> typ Same.operation
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    10
  val map_types_same: typ Same.operation -> term Same.operation
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    11
  val map_aterms_same: term Same.operation -> term Same.operation
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
    12
  val generalizeT_same: string list -> int -> typ Same.operation
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
    13
  val generalize_same: string list * string list -> int -> term Same.operation
36766
wenzelm
parents: 36765
diff changeset
    14
  val generalizeT: string list -> int -> typ -> typ
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    15
  val generalize: string list * string list -> int -> term -> term
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    16
  val instantiateT_maxidx: ((indexname * sort) * (typ * int)) list -> typ -> int -> typ * int
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    17
  val instantiate_maxidx:
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    18
    ((indexname * sort) * (typ * int)) list * ((indexname * typ) * (term * int)) list ->
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    19
    term -> int -> term * int
67698
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
    20
  val instantiateT_frees_same: ((string * sort) * typ) list -> typ Same.operation
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
    21
  val instantiate_frees_same: ((string * sort) * typ) list * ((string * typ) * term) list ->
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
    22
    term Same.operation
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
    23
  val instantiateT_frees: ((string * sort) * typ) list -> typ -> typ
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
    24
  val instantiate_frees: ((string * sort) * typ) list * ((string * typ) * term) list ->
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
    25
    term -> term
36766
wenzelm
parents: 36765
diff changeset
    26
  val instantiateT_same: ((indexname * sort) * typ) list -> typ Same.operation
wenzelm
parents: 36765
diff changeset
    27
  val instantiate_same: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
wenzelm
parents: 36765
diff changeset
    28
    term Same.operation
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
    29
  val instantiateT: ((indexname * sort) * typ) list -> typ -> typ
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    30
  val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    31
    term -> term
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 68234
diff changeset
    32
  val zero_var_indexes_inst: Name.context -> term list ->
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    33
    ((indexname * sort) * typ) list * ((indexname * typ) * term) list
68234
07eb13eb4065 tuned signature;
wenzelm
parents: 67698
diff changeset
    34
  val zero_var_indexes: term -> term
07eb13eb4065 tuned signature;
wenzelm
parents: 67698
diff changeset
    35
  val zero_var_indexes_list: term list -> term list
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    36
end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    37
31977
e03059ae2d82 renamed structure TermSubst to Term_Subst;
wenzelm
parents: 29269
diff changeset
    38
structure Term_Subst: TERM_SUBST =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    39
struct
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    40
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    41
(* generic mapping *)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    42
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    43
fun map_atypsT_same f =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    44
  let
32020
9abf5d66606e use structure Same;
wenzelm
parents: 32016
diff changeset
    45
    fun typ (Type (a, Ts)) = Type (a, Same.map typ Ts)
9abf5d66606e use structure Same;
wenzelm
parents: 32016
diff changeset
    46
      | typ T = f T;
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    47
  in typ end;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    48
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    49
fun map_types_same f =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    50
  let
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    51
    fun term (Const (a, T)) = Const (a, f T)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    52
      | term (Free (a, T)) = Free (a, f T)
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    53
      | term (Var (v, T)) = Var (v, f T)
32020
9abf5d66606e use structure Same;
wenzelm
parents: 32016
diff changeset
    54
      | term (Bound _) = raise Same.SAME
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    55
      | term (Abs (x, T, t)) =
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    56
          (Abs (x, f T, Same.commit term t)
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    57
            handle Same.SAME => Abs (x, T, term t))
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    58
      | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u);
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    59
  in term end;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    60
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    61
fun map_aterms_same f =
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    62
  let
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    63
    fun term (Abs (x, T, t)) = Abs (x, T, term t)
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    64
      | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u)
31980
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    65
      | term a = f a;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    66
  in term end;
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    67
c7c1d545007e added some generic mapping combinators;
wenzelm
parents: 31977
diff changeset
    68
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    69
(* generalization of fixed variables *)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    70
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    71
fun generalizeT_same [] _ _ = raise Same.SAME
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    72
  | generalizeT_same tfrees idx ty =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    73
      let
32020
9abf5d66606e use structure Same;
wenzelm
parents: 32016
diff changeset
    74
        fun gen (Type (a, Ts)) = Type (a, Same.map gen Ts)
9abf5d66606e use structure Same;
wenzelm
parents: 32016
diff changeset
    75
          | gen (TFree (a, S)) =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    76
              if member (op =) tfrees a then TVar ((a, idx), S)
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    77
              else raise Same.SAME
32020
9abf5d66606e use structure Same;
wenzelm
parents: 32016
diff changeset
    78
          | gen _ = raise Same.SAME;
9abf5d66606e use structure Same;
wenzelm
parents: 32016
diff changeset
    79
      in gen ty end;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    80
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    81
fun generalize_same ([], []) _ _ = raise Same.SAME
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    82
  | generalize_same (tfrees, frees) idx tm =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    83
      let
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    84
        val genT = generalizeT_same tfrees idx;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    85
        fun gen (Free (x, T)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    86
              if member (op =) frees x then
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    87
                Var (Name.clean_index (x, idx), Same.commit genT T)
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    88
              else Free (x, genT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    89
          | gen (Var (xi, T)) = Var (xi, genT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    90
          | gen (Const (c, T)) = Const (c, genT T)
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    91
          | gen (Bound _) = raise Same.SAME
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    92
          | gen (Abs (x, T, t)) =
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    93
              (Abs (x, genT T, Same.commit gen t)
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    94
                handle Same.SAME => Abs (x, T, gen t))
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
    95
          | gen (t $ u) = (gen t $ Same.commit gen u handle Same.SAME => t $ gen u);
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    96
      in gen tm end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
    97
36766
wenzelm
parents: 36765
diff changeset
    98
fun generalizeT names i ty = Same.commit (generalizeT_same names i) ty;
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
    99
fun generalize names i tm = Same.commit (generalize_same names i) tm;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   100
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   101
67698
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   102
(* instantiation of free variables (types before terms) *)
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   103
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   104
fun instantiateT_frees_same [] _ = raise Same.SAME
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   105
  | instantiateT_frees_same instT ty =
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   106
      let
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   107
        fun subst (Type (a, Ts)) = Type (a, Same.map subst Ts)
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   108
          | subst (TFree v) =
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   109
              (case AList.lookup (op =) instT v of
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   110
                SOME T => T
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   111
              | NONE => raise Same.SAME)
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   112
          | subst _ = raise Same.SAME;
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   113
      in subst ty end;
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   114
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   115
fun instantiate_frees_same ([], []) _ = raise Same.SAME
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   116
  | instantiate_frees_same (instT, inst) tm =
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   117
      let
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   118
        val substT = instantiateT_frees_same instT;
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   119
        fun subst (Const (c, T)) = Const (c, substT T)
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   120
          | subst (Free (x, T)) =
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   121
              let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   122
                (case AList.lookup (op =) inst (x, T') of
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   123
                   SOME t => t
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   124
                 | NONE => if same then raise Same.SAME else Free (x, T'))
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   125
              end
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   126
          | subst (Var (xi, T)) = Var (xi, substT T)
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   127
          | subst (Bound _) = raise Same.SAME
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   128
          | subst (Abs (x, T, t)) =
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   129
              (Abs (x, substT T, Same.commit subst t)
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   130
                handle Same.SAME => Abs (x, T, subst t))
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   131
          | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u);
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   132
      in subst tm end;
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   133
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   134
fun instantiateT_frees instT ty = Same.commit (instantiateT_frees_same instT) ty;
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   135
fun instantiate_frees inst tm = Same.commit (instantiate_frees_same inst) tm;
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   136
67caf783b9ee explicit operations to instantiate frees: typ, term, thm, morphism;
wenzelm
parents: 45395
diff changeset
   137
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   138
(* instantiation of schematic variables (types before terms) -- recomputes maxidx *)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   139
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   140
local
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   141
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   142
fun no_index (x, y) = (x, (y, ~1));
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   143
fun no_indexes1 inst = map no_index inst;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   144
fun no_indexes2 (inst1, inst2) = (map no_index inst1, map no_index inst2);
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   145
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   146
fun instT_same maxidx instT ty =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   147
  let
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   148
    fun maxify i = if i > ! maxidx then maxidx := i else ();
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   149
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   150
    fun subst_typ (Type (a, Ts)) = Type (a, subst_typs Ts)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   151
      | subst_typ (TVar ((a, i), S)) =
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   152
          (case AList.lookup Term.eq_tvar instT ((a, i), S) of
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   153
            SOME (T, j) => (maxify j; T)
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   154
          | NONE => (maxify i; raise Same.SAME))
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   155
      | subst_typ _ = raise Same.SAME
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   156
    and subst_typs (T :: Ts) =
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   157
        (subst_typ T :: Same.commit subst_typs Ts
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   158
          handle Same.SAME => T :: subst_typs Ts)
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   159
      | subst_typs [] = raise Same.SAME;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   160
  in subst_typ ty end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   161
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   162
fun inst_same maxidx (instT, inst) tm =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   163
  let
21184
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   164
    fun maxify i = if i > ! maxidx then maxidx := i else ();
35baf14cfb6d instantiate: avoid global references;
wenzelm
parents: 20513
diff changeset
   165
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   166
    val substT = instT_same maxidx instT;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   167
    fun subst (Const (c, T)) = Const (c, substT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   168
      | subst (Free (x, T)) = Free (x, substT T)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   169
      | subst (Var ((x, i), T)) =
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   170
          let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   171
            (case AList.lookup Term.eq_var inst ((x, i), T') of
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   172
               SOME (t, j) => (maxify j; t)
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   173
             | NONE => (maxify i; if same then raise Same.SAME else Var ((x, i), T')))
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   174
          end
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   175
      | subst (Bound _) = raise Same.SAME
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   176
      | subst (Abs (x, T, t)) =
32016
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   177
          (Abs (x, substT T, Same.commit subst t)
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   178
            handle Same.SAME => Abs (x, T, subst t))
597b3b69b8d8 use structure Same;
wenzelm
parents: 31980
diff changeset
   179
      | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u);
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   180
  in subst tm end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   181
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   182
in
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   183
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   184
fun instantiateT_maxidx instT ty i =
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32020
diff changeset
   185
  let val maxidx = Unsynchronized.ref i
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   186
  in (Same.commit (instT_same maxidx instT) ty, ! maxidx) end;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   187
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   188
fun instantiate_maxidx insts tm i =
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32020
diff changeset
   189
  let val maxidx = Unsynchronized.ref i
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   190
  in (Same.commit (inst_same maxidx insts) tm, ! maxidx) end;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   191
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   192
fun instantiateT_same [] _ = raise Same.SAME
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   193
  | instantiateT_same instT ty = instT_same (Unsynchronized.ref ~1) (no_indexes1 instT) ty;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   194
36620
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   195
fun instantiate_same ([], []) _ = raise Same.SAME
e6bb250402b5 simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
wenzelm
parents: 35408
diff changeset
   196
  | instantiate_same insts tm = inst_same (Unsynchronized.ref ~1) (no_indexes2 insts) tm;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   197
36766
wenzelm
parents: 36765
diff changeset
   198
fun instantiateT instT ty = Same.commit (instantiateT_same instT) ty;
wenzelm
parents: 36765
diff changeset
   199
fun instantiate inst tm = Same.commit (instantiate_same inst) tm;
wenzelm
parents: 36765
diff changeset
   200
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   201
end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   202
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   203
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   204
(* zero var indexes *)
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   205
45395
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   206
structure TVars = Table(type key = indexname * sort val ord = Term_Ord.tvar_ord);
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   207
structure Vars = Table(type key = indexname * typ val ord = Term_Ord.var_ord);
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   208
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   209
fun zero_var_inst mk (v as ((x, i), X)) (inst, used) =
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   210
  let
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   211
    val (x', used') = Name.variant (if Name.is_bound x then "u" else x) used;
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   212
  in if x = x' andalso i = 0 then (inst, used') else ((v, mk ((x', 0), X)) :: inst, used') end;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   213
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 68234
diff changeset
   214
fun zero_var_indexes_inst used ts =
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   215
  let
45395
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   216
    val (instT, _) =
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   217
      TVars.fold (zero_var_inst TVar o #1)
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   218
        ((fold o fold_types o fold_atyps) (fn TVar v =>
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   219
          TVars.insert (K true) (v, ()) | _ => I) ts TVars.empty)
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 68234
diff changeset
   220
        ([], used);
45395
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   221
    val (inst, _) =
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   222
      Vars.fold (zero_var_inst Var o #1)
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   223
        ((fold o fold_aterms) (fn Var (xi, T) =>
830c9b9b0d66 more scalable zero_var_indexes, depending on canonical order within table;
wenzelm
parents: 43326
diff changeset
   224
          Vars.insert (K true) ((xi, instantiateT instT T), ()) | _ => I) ts Vars.empty)
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 68234
diff changeset
   225
        ([], used);
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   226
  in (instT, inst) end;
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   227
69023
cef000855cf4 clarified standardization of variables, with proper treatment of local variables;
wenzelm
parents: 68234
diff changeset
   228
fun zero_var_indexes_list ts = map (instantiate (zero_var_indexes_inst Name.context ts)) ts;
68234
07eb13eb4065 tuned signature;
wenzelm
parents: 67698
diff changeset
   229
val zero_var_indexes = singleton zero_var_indexes_list;
20513
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   230
4294eb252283 Efficient term substitution -- avoids copying.
wenzelm
parents:
diff changeset
   231
end;