| author | wenzelm | 
| Tue, 07 Nov 2017 16:44:25 +0100 | |
| changeset 67025 | 961285f581e6 | 
| parent 45395 | 830c9b9b0d66 | 
| child 67698 | 67caf783b9ee | 
| 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 | |
| 36766 | 20 | val instantiateT_same: ((indexname * sort) * typ) list -> typ Same.operation | 
| 21 | val instantiate_same: ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> | |
| 22 | term Same.operation | |
| 36620 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
35408diff
changeset | 23 | val instantiateT: ((indexname * sort) * typ) list -> typ -> typ | 
| 20513 | 24 | val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> | 
| 25 | term -> term | |
| 26 | val zero_var_indexes: term -> term | |
| 21607 | 27 | val zero_var_indexes_inst: term list -> | 
| 20513 | 28 | ((indexname * sort) * typ) list * ((indexname * typ) * term) list | 
| 29 | end; | |
| 30 | ||
| 31977 | 31 | structure Term_Subst: TERM_SUBST = | 
| 20513 | 32 | struct | 
| 33 | ||
| 31980 | 34 | (* generic mapping *) | 
| 35 | ||
| 36 | fun map_atypsT_same f = | |
| 37 | let | |
| 32020 | 38 | fun typ (Type (a, Ts)) = Type (a, Same.map typ Ts) | 
| 39 | | typ T = f T; | |
| 31980 | 40 | in typ end; | 
| 41 | ||
| 42 | fun map_types_same f = | |
| 43 | let | |
| 44 | fun term (Const (a, T)) = Const (a, f T) | |
| 45 | | term (Free (a, T)) = Free (a, f T) | |
| 46 | | term (Var (v, T)) = Var (v, f T) | |
| 32020 | 47 | | term (Bound _) = raise Same.SAME | 
| 31980 | 48 | | term (Abs (x, T, t)) = | 
| 32016 | 49 | (Abs (x, f T, Same.commit term t) | 
| 50 | handle Same.SAME => Abs (x, T, term t)) | |
| 51 | | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u); | |
| 31980 | 52 | in term end; | 
| 53 | ||
| 54 | fun map_aterms_same f = | |
| 55 | let | |
| 56 | fun term (Abs (x, T, t)) = Abs (x, T, term t) | |
| 32016 | 57 | | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u) | 
| 31980 | 58 | | term a = f a; | 
| 59 | in term end; | |
| 60 | ||
| 61 | ||
| 20513 | 62 | (* generalization of fixed variables *) | 
| 63 | ||
| 32016 | 64 | fun generalizeT_same [] _ _ = raise Same.SAME | 
| 20513 | 65 | | generalizeT_same tfrees idx ty = | 
| 66 | let | |
| 32020 | 67 | fun gen (Type (a, Ts)) = Type (a, Same.map gen Ts) | 
| 68 | | gen (TFree (a, S)) = | |
| 20513 | 69 | if member (op =) tfrees a then TVar ((a, idx), S) | 
| 32016 | 70 | else raise Same.SAME | 
| 32020 | 71 | | gen _ = raise Same.SAME; | 
| 72 | in gen ty end; | |
| 20513 | 73 | |
| 32016 | 74 | fun generalize_same ([], []) _ _ = raise Same.SAME | 
| 20513 | 75 | | generalize_same (tfrees, frees) idx tm = | 
| 76 | let | |
| 77 | val genT = generalizeT_same tfrees idx; | |
| 78 | fun gen (Free (x, T)) = | |
| 79 | if member (op =) frees x then | |
| 32016 | 80 | Var (Name.clean_index (x, idx), Same.commit genT T) | 
| 20513 | 81 | else Free (x, genT T) | 
| 82 | | gen (Var (xi, T)) = Var (xi, genT T) | |
| 83 | | gen (Const (c, T)) = Const (c, genT T) | |
| 32016 | 84 | | gen (Bound _) = raise Same.SAME | 
| 20513 | 85 | | gen (Abs (x, T, t)) = | 
| 32016 | 86 | (Abs (x, genT T, Same.commit gen t) | 
| 87 | handle Same.SAME => Abs (x, T, gen t)) | |
| 88 | | gen (t $ u) = (gen t $ Same.commit gen u handle Same.SAME => t $ gen u); | |
| 20513 | 89 | in gen tm end; | 
| 90 | ||
| 36766 | 91 | 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 | 92 | fun generalize names i tm = Same.commit (generalize_same names i) tm; | 
| 20513 | 93 | |
| 94 | ||
| 95 | (* instantiation of schematic variables (types before terms) -- recomputes maxidx *) | |
| 96 | ||
| 97 | local | |
| 98 | ||
| 99 | fun no_index (x, y) = (x, (y, ~1)); | |
| 100 | fun no_indexes1 inst = map no_index inst; | |
| 101 | fun no_indexes2 (inst1, inst2) = (map no_index inst1, map no_index inst2); | |
| 102 | ||
| 36620 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
35408diff
changeset | 103 | fun instT_same maxidx instT ty = | 
| 20513 | 104 | let | 
| 21184 | 105 | fun maxify i = if i > ! maxidx then maxidx := i else (); | 
| 106 | ||
| 20513 | 107 | fun subst_typ (Type (a, Ts)) = Type (a, subst_typs Ts) | 
| 108 | | subst_typ (TVar ((a, i), S)) = | |
| 109 | (case AList.lookup Term.eq_tvar instT ((a, i), S) of | |
| 110 | SOME (T, j) => (maxify j; T) | |
| 32016 | 111 | | NONE => (maxify i; raise Same.SAME)) | 
| 112 | | subst_typ _ = raise Same.SAME | |
| 20513 | 113 | and subst_typs (T :: Ts) = | 
| 32016 | 114 | (subst_typ T :: Same.commit subst_typs Ts | 
| 115 | handle Same.SAME => T :: subst_typs Ts) | |
| 116 | | subst_typs [] = raise Same.SAME; | |
| 20513 | 117 | in subst_typ ty end; | 
| 118 | ||
| 36620 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
35408diff
changeset | 119 | fun inst_same maxidx (instT, inst) tm = | 
| 20513 | 120 | let | 
| 21184 | 121 | fun maxify i = if i > ! maxidx then maxidx := i else (); | 
| 122 | ||
| 36620 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
35408diff
changeset | 123 | val substT = instT_same maxidx instT; | 
| 20513 | 124 | fun subst (Const (c, T)) = Const (c, substT T) | 
| 125 | | subst (Free (x, T)) = Free (x, substT T) | |
| 126 | | subst (Var ((x, i), T)) = | |
| 32016 | 127 | let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in | 
| 20513 | 128 | (case AList.lookup Term.eq_var inst ((x, i), T') of | 
| 129 | SOME (t, j) => (maxify j; t) | |
| 32016 | 130 | | NONE => (maxify i; if same then raise Same.SAME else Var ((x, i), T'))) | 
| 20513 | 131 | end | 
| 32016 | 132 | | subst (Bound _) = raise Same.SAME | 
| 20513 | 133 | | subst (Abs (x, T, t)) = | 
| 32016 | 134 | (Abs (x, substT T, Same.commit subst t) | 
| 135 | handle Same.SAME => Abs (x, T, subst t)) | |
| 136 | | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); | |
| 20513 | 137 | in subst tm end; | 
| 138 | ||
| 139 | in | |
| 140 | ||
| 141 | fun instantiateT_maxidx instT ty i = | |
| 32738 | 142 | 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 | 143 | in (Same.commit (instT_same maxidx instT) ty, ! maxidx) end; | 
| 20513 | 144 | |
| 145 | fun instantiate_maxidx insts tm i = | |
| 32738 | 146 | 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 | 147 | in (Same.commit (inst_same maxidx insts) tm, ! maxidx) end; | 
| 20513 | 148 | |
| 36620 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
35408diff
changeset | 149 | 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 | 150 | | instantiateT_same instT ty = instT_same (Unsynchronized.ref ~1) (no_indexes1 instT) ty; | 
| 20513 | 151 | |
| 36620 
e6bb250402b5
simplified/unified fundamental operations on types/terms/proofterms -- prefer Same.operation over "option" variant;
 wenzelm parents: 
35408diff
changeset | 152 | 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 | 153 | | instantiate_same insts tm = inst_same (Unsynchronized.ref ~1) (no_indexes2 insts) tm; | 
| 20513 | 154 | |
| 36766 | 155 | fun instantiateT instT ty = Same.commit (instantiateT_same instT) ty; | 
| 156 | fun instantiate inst tm = Same.commit (instantiate_same inst) tm; | |
| 157 | ||
| 20513 | 158 | end; | 
| 159 | ||
| 160 | ||
| 161 | (* zero var indexes *) | |
| 162 | ||
| 45395 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 163 | 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 | 164 | 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 | 165 | |
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 166 | 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 | 167 | let | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 168 | 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 | 169 | in if x = x' andalso i = 0 then (inst, used') else ((v, mk ((x', 0), X)) :: inst, used') end; | 
| 20513 | 170 | |
| 21607 | 171 | fun zero_var_indexes_inst ts = | 
| 20513 | 172 | let | 
| 45395 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 173 | val (instT, _) = | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 174 | TVars.fold (zero_var_inst TVar o #1) | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 175 | ((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 | 176 | TVars.insert (K true) (v, ()) | _ => I) ts TVars.empty) | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 177 | ([], Name.context); | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 178 | val (inst, _) = | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 179 | Vars.fold (zero_var_inst Var o #1) | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 180 | ((fold o fold_aterms) (fn Var (xi, T) => | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 181 | Vars.insert (K true) ((xi, instantiateT instT T), ()) | _ => I) ts Vars.empty) | 
| 
830c9b9b0d66
more scalable zero_var_indexes, depending on canonical order within table;
 wenzelm parents: 
43326diff
changeset | 182 | ([], Name.context); | 
| 20513 | 183 | in (instT, inst) end; | 
| 184 | ||
| 21607 | 185 | fun zero_var_indexes t = instantiate (zero_var_indexes_inst [t]) t; | 
| 20513 | 186 | |
| 187 | end; |