| author | wenzelm | 
| Tue, 21 Jul 2009 20:37:32 +0200 | |
| changeset 32104 | d1d98a02473e | 
| parent 32020 | 9abf5d66606e | 
| child 32738 | 15bb09ca0378 | 
| 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 | |
| 31980 | 12 | val map_atypsT_option: (typ -> typ option) -> typ -> typ option | 
| 13 | val map_atyps_option: (typ -> typ option) -> term -> term option | |
| 14 | val map_types_option: (typ -> typ option) -> term -> term option | |
| 15 | val map_aterms_option: (term -> term option) -> term -> term option | |
| 20513 | 16 | val generalize: string list * string list -> int -> term -> term | 
| 17 | val generalizeT: string list -> int -> typ -> typ | |
| 18 | val generalize_option: string list * string list -> int -> term -> term option | |
| 19 | val generalizeT_option: string list -> int -> typ -> typ option | |
| 20 | val instantiateT_maxidx: ((indexname * sort) * (typ * int)) list -> typ -> int -> typ * int | |
| 21 | val instantiate_maxidx: | |
| 22 | ((indexname * sort) * (typ * int)) list * ((indexname * typ) * (term * int)) list -> | |
| 23 | term -> int -> term * int | |
| 24 | val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> | |
| 25 | term -> term | |
| 26 | val instantiateT: ((indexname * sort) * typ) list -> typ -> typ | |
| 27 | val instantiate_option: ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> | |
| 28 | term -> term option | |
| 29 | val instantiateT_option: ((indexname * sort) * typ) list -> typ -> typ option | |
| 30 | val zero_var_indexes: term -> term | |
| 21607 | 31 | val zero_var_indexes_inst: term list -> | 
| 20513 | 32 | ((indexname * sort) * typ) list * ((indexname * typ) * term) list | 
| 33 | end; | |
| 34 | ||
| 31977 | 35 | structure Term_Subst: TERM_SUBST = | 
| 20513 | 36 | struct | 
| 37 | ||
| 31980 | 38 | (* generic mapping *) | 
| 39 | ||
| 40 | fun map_atypsT_same f = | |
| 41 | let | |
| 32020 | 42 | fun typ (Type (a, Ts)) = Type (a, Same.map typ Ts) | 
| 43 | | typ T = f T; | |
| 31980 | 44 | in typ end; | 
| 45 | ||
| 46 | fun map_types_same f = | |
| 47 | let | |
| 48 | fun term (Const (a, T)) = Const (a, f T) | |
| 49 | | term (Free (a, T)) = Free (a, f T) | |
| 50 | | term (Var (v, T)) = Var (v, f T) | |
| 32020 | 51 | | term (Bound _) = raise Same.SAME | 
| 31980 | 52 | | term (Abs (x, T, t)) = | 
| 32016 | 53 | (Abs (x, f T, Same.commit term t) | 
| 54 | handle Same.SAME => Abs (x, T, term t)) | |
| 55 | | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u); | |
| 31980 | 56 | in term end; | 
| 57 | ||
| 58 | fun map_aterms_same f = | |
| 59 | let | |
| 60 | fun term (Abs (x, T, t)) = Abs (x, T, term t) | |
| 32016 | 61 | | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u) | 
| 31980 | 62 | | term a = f a; | 
| 63 | in term end; | |
| 64 | ||
| 32016 | 65 | val map_atypsT_option = Same.capture o map_atypsT_same o Same.function; | 
| 66 | val map_atyps_option = Same.capture o map_types_same o map_atypsT_same o Same.function; | |
| 67 | val map_types_option = Same.capture o map_types_same o Same.function; | |
| 68 | val map_aterms_option = Same.capture o map_aterms_same o Same.function; | |
| 31980 | 69 | |
| 70 | ||
| 20513 | 71 | (* generalization of fixed variables *) | 
| 72 | ||
| 73 | local | |
| 74 | ||
| 32016 | 75 | fun generalizeT_same [] _ _ = raise Same.SAME | 
| 20513 | 76 | | generalizeT_same tfrees idx ty = | 
| 77 | let | |
| 32020 | 78 | fun gen (Type (a, Ts)) = Type (a, Same.map gen Ts) | 
| 79 | | gen (TFree (a, S)) = | |
| 20513 | 80 | if member (op =) tfrees a then TVar ((a, idx), S) | 
| 32016 | 81 | else raise Same.SAME | 
| 32020 | 82 | | gen _ = raise Same.SAME; | 
| 83 | in gen ty end; | |
| 20513 | 84 | |
| 32016 | 85 | fun generalize_same ([], []) _ _ = raise Same.SAME | 
| 20513 | 86 | | generalize_same (tfrees, frees) idx tm = | 
| 87 | let | |
| 88 | val genT = generalizeT_same tfrees idx; | |
| 89 | fun gen (Free (x, T)) = | |
| 90 | if member (op =) frees x then | |
| 32016 | 91 | Var (Name.clean_index (x, idx), Same.commit genT T) | 
| 20513 | 92 | else Free (x, genT T) | 
| 93 | | gen (Var (xi, T)) = Var (xi, genT T) | |
| 94 | | gen (Const (c, T)) = Const (c, genT T) | |
| 32016 | 95 | | gen (Bound _) = raise Same.SAME | 
| 20513 | 96 | | gen (Abs (x, T, t)) = | 
| 32016 | 97 | (Abs (x, genT T, Same.commit gen t) | 
| 98 | handle Same.SAME => Abs (x, T, gen t)) | |
| 99 | | gen (t $ u) = (gen t $ Same.commit gen u handle Same.SAME => t $ gen u); | |
| 20513 | 100 | in gen tm end; | 
| 101 | ||
| 102 | in | |
| 103 | ||
| 32016 | 104 | fun generalize names i tm = generalize_same names i tm handle Same.SAME => tm; | 
| 105 | fun generalizeT names i ty = generalizeT_same names i ty handle Same.SAME => ty; | |
| 20513 | 106 | |
| 32016 | 107 | fun generalize_option names i tm = SOME (generalize_same names i tm) handle Same.SAME => NONE; | 
| 108 | fun generalizeT_option names i ty = SOME (generalizeT_same names i ty) handle Same.SAME => NONE; | |
| 20513 | 109 | |
| 110 | end; | |
| 111 | ||
| 112 | ||
| 113 | (* instantiation of schematic variables (types before terms) -- recomputes maxidx *) | |
| 114 | ||
| 115 | local | |
| 116 | ||
| 117 | fun no_index (x, y) = (x, (y, ~1)); | |
| 118 | fun no_indexes1 inst = map no_index inst; | |
| 119 | fun no_indexes2 (inst1, inst2) = (map no_index inst1, map no_index inst2); | |
| 120 | ||
| 21184 | 121 | fun instantiateT_same maxidx instT ty = | 
| 20513 | 122 | let | 
| 21184 | 123 | fun maxify i = if i > ! maxidx then maxidx := i else (); | 
| 124 | ||
| 20513 | 125 | fun subst_typ (Type (a, Ts)) = Type (a, subst_typs Ts) | 
| 126 | | subst_typ (TVar ((a, i), S)) = | |
| 127 | (case AList.lookup Term.eq_tvar instT ((a, i), S) of | |
| 128 | SOME (T, j) => (maxify j; T) | |
| 32016 | 129 | | NONE => (maxify i; raise Same.SAME)) | 
| 130 | | subst_typ _ = raise Same.SAME | |
| 20513 | 131 | and subst_typs (T :: Ts) = | 
| 32016 | 132 | (subst_typ T :: Same.commit subst_typs Ts | 
| 133 | handle Same.SAME => T :: subst_typs Ts) | |
| 134 | | subst_typs [] = raise Same.SAME; | |
| 20513 | 135 | in subst_typ ty end; | 
| 136 | ||
| 21184 | 137 | fun instantiate_same maxidx (instT, inst) tm = | 
| 20513 | 138 | let | 
| 21184 | 139 | fun maxify i = if i > ! maxidx then maxidx := i else (); | 
| 140 | ||
| 141 | val substT = instantiateT_same maxidx instT; | |
| 20513 | 142 | fun subst (Const (c, T)) = Const (c, substT T) | 
| 143 | | subst (Free (x, T)) = Free (x, substT T) | |
| 144 | | subst (Var ((x, i), T)) = | |
| 32016 | 145 | let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in | 
| 20513 | 146 | (case AList.lookup Term.eq_var inst ((x, i), T') of | 
| 147 | SOME (t, j) => (maxify j; t) | |
| 32016 | 148 | | NONE => (maxify i; if same then raise Same.SAME else Var ((x, i), T'))) | 
| 20513 | 149 | end | 
| 32016 | 150 | | subst (Bound _) = raise Same.SAME | 
| 20513 | 151 | | subst (Abs (x, T, t)) = | 
| 32016 | 152 | (Abs (x, substT T, Same.commit subst t) | 
| 153 | handle Same.SAME => Abs (x, T, subst t)) | |
| 154 | | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); | |
| 20513 | 155 | in subst tm end; | 
| 156 | ||
| 157 | in | |
| 158 | ||
| 159 | fun instantiateT_maxidx instT ty i = | |
| 21184 | 160 | let val maxidx = ref i | 
| 32016 | 161 | in (instantiateT_same maxidx instT ty handle Same.SAME => ty, ! maxidx) end; | 
| 20513 | 162 | |
| 163 | fun instantiate_maxidx insts tm i = | |
| 21184 | 164 | let val maxidx = ref i | 
| 32016 | 165 | in (instantiate_same maxidx insts tm handle Same.SAME => tm, ! maxidx) end; | 
| 20513 | 166 | |
| 21315 | 167 | fun instantiateT [] ty = ty | 
| 168 | | instantiateT instT ty = | |
| 32016 | 169 | (instantiateT_same (ref ~1) (no_indexes1 instT) ty handle Same.SAME => ty); | 
| 21184 | 170 | |
| 21315 | 171 | fun instantiate ([], []) tm = tm | 
| 172 | | instantiate insts tm = | |
| 32016 | 173 | (instantiate_same (ref ~1) (no_indexes2 insts) tm handle Same.SAME => tm); | 
| 20513 | 174 | |
| 21315 | 175 | fun instantiateT_option [] _ = NONE | 
| 176 | | instantiateT_option instT ty = | |
| 32016 | 177 | (SOME (instantiateT_same (ref ~1) (no_indexes1 instT) ty) handle Same.SAME => NONE); | 
| 20513 | 178 | |
| 21315 | 179 | fun instantiate_option ([], []) _ = NONE | 
| 180 | | instantiate_option insts tm = | |
| 32016 | 181 | (SOME (instantiate_same (ref ~1) (no_indexes2 insts) tm) handle Same.SAME => NONE); | 
| 20513 | 182 | |
| 183 | end; | |
| 184 | ||
| 185 | ||
| 186 | (* zero var indexes *) | |
| 187 | ||
| 188 | fun zero_var_inst vars = | |
| 189 | fold (fn v as ((x, i), X) => fn (inst, used) => | |
| 190 | let | |
| 191 | val ([x'], used') = Name.variants [if Name.is_bound x then "u" else x] used; | |
| 192 | in if x = x' andalso i = 0 then (inst, used') else ((v, ((x', 0), X)) :: inst, used') end) | |
| 193 | vars ([], Name.context) |> #1; | |
| 194 | ||
| 21607 | 195 | fun zero_var_indexes_inst ts = | 
| 20513 | 196 | let | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
21607diff
changeset | 197 | val tvars = sort TermOrd.tvar_ord (fold Term.add_tvars ts []); | 
| 21607 | 198 | 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: 
21607diff
changeset | 199 | val vars = sort TermOrd.var_ord (map (apsnd (instantiateT instT)) (fold Term.add_vars ts [])); | 
| 21607 | 200 | val inst = map (apsnd Var) (zero_var_inst vars); | 
| 20513 | 201 | in (instT, inst) end; | 
| 202 | ||
| 21607 | 203 | fun zero_var_indexes t = instantiate (zero_var_indexes_inst [t]) t; | 
| 20513 | 204 | |
| 205 | end; |