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