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