author | wenzelm |
Thu, 30 Nov 2006 14:17:36 +0100 | |
changeset 21610 | 52c0d3280798 |
parent 21571 | 6096c956a11f |
child 21683 | b90fa6c8e062 |
permissions | -rw-r--r-- |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/variable.ML |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
2 |
ID: $Id$ |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
3 |
Author: Makarius |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
4 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
5 |
Fixed type/term variables and polymorphic term abbreviations. |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
6 |
*) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
7 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
8 |
signature VARIABLE = |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
9 |
sig |
20303 | 10 |
val is_body: Proof.context -> bool |
11 |
val set_body: bool -> Proof.context -> Proof.context |
|
12 |
val restore_body: Proof.context -> Proof.context -> Proof.context |
|
13 |
val names_of: Proof.context -> Name.context |
|
14 |
val fixes_of: Proof.context -> (string * string) list |
|
15 |
val binds_of: Proof.context -> (typ * term) Vartab.table |
|
16 |
val constraints_of: Proof.context -> typ Vartab.table * sort Vartab.table |
|
17 |
val is_declared: Proof.context -> string -> bool |
|
18 |
val is_fixed: Proof.context -> string -> bool |
|
19 |
val newly_fixed: Proof.context -> Proof.context -> string -> bool |
|
21571 | 20 |
val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list |
20303 | 21 |
val default_type: Proof.context -> string -> typ option |
22 |
val def_type: Proof.context -> bool -> indexname -> typ option |
|
23 |
val def_sort: Proof.context -> indexname -> sort option |
|
24 |
val declare_constraints: term -> Proof.context -> Proof.context |
|
25 |
val declare_internal: term -> Proof.context -> Proof.context |
|
26 |
val declare_term: term -> Proof.context -> Proof.context |
|
27 |
val declare_prf: Proofterm.proof -> Proof.context -> Proof.context |
|
28 |
val declare_thm: thm -> Proof.context -> Proof.context |
|
29 |
val thm_context: thm -> Proof.context |
|
30 |
val variant_frees: Proof.context -> term list -> (string * 'a) list -> (string * 'a) list |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
31 |
val hidden_polymorphism: term -> typ -> (indexname * sort) list |
20303 | 32 |
val add_binds: (indexname * term option) list -> Proof.context -> Proof.context |
33 |
val expand_binds: Proof.context -> term -> term |
|
34 |
val add_fixes: string list -> Proof.context -> string list * Proof.context |
|
35 |
val add_fixes_direct: string list -> Proof.context -> Proof.context |
|
21369 | 36 |
val auto_fixes: term -> Proof.context -> Proof.context |
20797
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents:
20579
diff
changeset
|
37 |
val variant_fixes: string list -> Proof.context -> string list * Proof.context |
20303 | 38 |
val invent_types: sort list -> Proof.context -> (string * sort) list * Proof.context |
39 |
val export_inst: Proof.context -> Proof.context -> string list * string list |
|
40 |
val exportT_inst: Proof.context -> Proof.context -> string list |
|
41 |
val export_terms: Proof.context -> Proof.context -> term list -> term list |
|
42 |
val exportT_terms: Proof.context -> Proof.context -> term list -> term list |
|
43 |
val exportT: Proof.context -> Proof.context -> thm list -> thm list |
|
44 |
val export_prf: Proof.context -> Proof.context -> Proofterm.proof -> Proofterm.proof |
|
45 |
val export: Proof.context -> Proof.context -> thm list -> thm list |
|
21522 | 46 |
val export_morphism: Proof.context -> Proof.context -> morphism |
20303 | 47 |
val importT_inst: term list -> Proof.context -> ((indexname * sort) * typ) list * Proof.context |
48 |
val import_inst: bool -> term list -> Proof.context -> |
|
49 |
(((indexname * sort) * typ) list * ((indexname * typ) * term) list) * Proof.context |
|
50 |
val importT_terms: term list -> Proof.context -> term list * Proof.context |
|
51 |
val import_terms: bool -> term list -> Proof.context -> term list * Proof.context |
|
52 |
val importT: thm list -> Proof.context -> (ctyp list * thm list) * Proof.context |
|
53 |
val import_prf: bool -> Proofterm.proof -> Proof.context -> Proofterm.proof * Proof.context |
|
54 |
val import: bool -> thm list -> Proof.context -> |
|
55 |
((ctyp list * cterm list) * thm list) * Proof.context |
|
21287 | 56 |
val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list |
57 |
val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list |
|
20303 | 58 |
val focus: cterm -> Proof.context -> (cterm list * cterm) * Proof.context |
59 |
val focus_subgoal: int -> thm -> Proof.context -> (cterm list * cterm) * Proof.context |
|
60 |
val warn_extra_tfrees: Proof.context -> Proof.context -> unit |
|
61 |
val polymorphic: Proof.context -> term list -> term list |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
62 |
end; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
63 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
64 |
structure Variable: VARIABLE = |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
65 |
struct |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
66 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
67 |
(** local context data **) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
68 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
69 |
datatype data = Data of |
20102 | 70 |
{is_body: bool, (*inner body mode*) |
20162 | 71 |
names: Name.context, (*type/term variable names*) |
72 |
fixes: (string * string) list, (*term fixes -- extern/intern*) |
|
20102 | 73 |
binds: (typ * term) Vartab.table, (*term bindings*) |
20162 | 74 |
type_occs: string list Symtab.table, (*type variables -- possibly within term variables*) |
75 |
constraints: |
|
20102 | 76 |
typ Vartab.table * (*type constraints*) |
20162 | 77 |
sort Vartab.table}; (*default sorts*) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
78 |
|
20162 | 79 |
fun make_data (is_body, names, fixes, binds, type_occs, constraints) = |
80 |
Data {is_body = is_body, names = names, fixes = fixes, binds = binds, |
|
81 |
type_occs = type_occs, constraints = constraints}; |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
82 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
83 |
structure Data = ProofDataFun |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
84 |
( |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
85 |
val name = "Pure/variable"; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
86 |
type T = data; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
87 |
fun init thy = |
20162 | 88 |
make_data (false, Name.context, [], Vartab.empty, Symtab.empty, (Vartab.empty, Vartab.empty)); |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
89 |
fun print _ _ = (); |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
90 |
); |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
91 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
92 |
val _ = Context.add_setup Data.init; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
93 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
94 |
fun map_data f = |
20162 | 95 |
Data.map (fn Data {is_body, names, fixes, binds, type_occs, constraints} => |
96 |
make_data (f (is_body, names, fixes, binds, type_occs, constraints))); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
97 |
|
20162 | 98 |
fun map_names f = map_data (fn (is_body, names, fixes, binds, type_occs, constraints) => |
99 |
(is_body, f names, fixes, binds, type_occs, constraints)); |
|
100 |
||
101 |
fun map_fixes f = map_data (fn (is_body, names, fixes, binds, type_occs, constraints) => |
|
102 |
(is_body, names, f fixes, binds, type_occs, constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
103 |
|
20162 | 104 |
fun map_binds f = map_data (fn (is_body, names, fixes, binds, type_occs, constraints) => |
105 |
(is_body, names, fixes, f binds, type_occs, constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
106 |
|
20162 | 107 |
fun map_type_occs f = map_data (fn (is_body, names, fixes, binds, type_occs, constraints) => |
108 |
(is_body, names, fixes, binds, f type_occs, constraints)); |
|
20102 | 109 |
|
20162 | 110 |
fun map_constraints f = map_data (fn (is_body, names, fixes, binds, type_occs, constraints) => |
111 |
(is_body, names, fixes, binds, type_occs, f constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
112 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
113 |
fun rep_data ctxt = Data.get ctxt |> (fn Data args => args); |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
114 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
115 |
val is_body = #is_body o rep_data; |
20162 | 116 |
fun set_body b = map_data (fn (_, names, fixes, binds, type_occs, constraints) => |
117 |
(b, names, fixes, binds, type_occs, constraints)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
118 |
fun restore_body ctxt = set_body (is_body ctxt); |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
119 |
|
20102 | 120 |
val names_of = #names o rep_data; |
121 |
val fixes_of = #fixes o rep_data; |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
122 |
val binds_of = #binds o rep_data; |
20162 | 123 |
val type_occs_of = #type_occs o rep_data; |
124 |
val constraints_of = #constraints o rep_data; |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
125 |
|
20162 | 126 |
val is_declared = Name.is_declared o names_of; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
127 |
fun is_fixed ctxt x = exists (fn (_, y) => x = y) (fixes_of ctxt); |
20149 | 128 |
fun newly_fixed inner outer x = is_fixed inner x andalso not (is_fixed outer x); |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
129 |
|
21571 | 130 |
fun add_fixed ctxt = Term.fold_aterms |
131 |
(fn Free (x, T) => if is_fixed ctxt x then insert (op =) (x, T) else I | _ => I); |
|
132 |
||
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
133 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
134 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
135 |
(** declarations **) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
136 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
137 |
(* default sorts and types *) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
138 |
|
20162 | 139 |
fun default_type ctxt x = Vartab.lookup (#1 (constraints_of ctxt)) (x, ~1); |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
140 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
141 |
fun def_type ctxt pattern xi = |
20162 | 142 |
let val {binds, constraints = (types, _), ...} = rep_data ctxt in |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
143 |
(case Vartab.lookup types xi of |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
144 |
NONE => |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
145 |
if pattern then NONE |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
146 |
else Vartab.lookup binds xi |> Option.map (TypeInfer.polymorphicT o #1) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
147 |
| some => some) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
148 |
end; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
149 |
|
20162 | 150 |
val def_sort = Vartab.lookup o #2 o constraints_of; |
151 |
||
152 |
||
153 |
(* names *) |
|
154 |
||
155 |
val declare_type_names = map_names o |
|
156 |
fold_types (fold_atyps (fn TFree (a, _) => Name.declare a | _ => I)); |
|
157 |
||
158 |
fun declare_names t = |
|
159 |
declare_type_names t #> |
|
160 |
map_names (fold_aterms (fn Free (x, _) => Name.declare x | _ => I) t); |
|
161 |
||
162 |
||
163 |
(* type occurrences *) |
|
164 |
||
165 |
val declare_type_occs = map_type_occs o fold_term_types |
|
166 |
(fn Free (x, _) => fold_atyps (fn TFree (a, _) => Symtab.insert_list (op =) (a, x) | _ => I) |
|
167 |
| _ => fold_atyps (fn TFree (a, _) => Symtab.default (a, []) | _ => I)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
168 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
169 |
|
20162 | 170 |
(* constraints *) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
171 |
|
20162 | 172 |
fun redeclare_skolems ctxt = ctxt |> map_constraints (apfst (fn types => |
173 |
let |
|
174 |
fun decl (x, x') = |
|
175 |
(case default_type ctxt x' of |
|
176 |
SOME T => Vartab.update ((x, ~1), T) |
|
177 |
| NONE => I); |
|
178 |
in fold_rev decl (fixes_of ctxt) types end)); |
|
20084 | 179 |
|
21355 | 180 |
fun constrain_tvar (xi, S) = |
181 |
if S = dummyS then Vartab.delete_safe xi else Vartab.update (xi, S); |
|
182 |
||
20162 | 183 |
fun declare_constraints t = map_constraints (fn (types, sorts) => |
184 |
let |
|
185 |
val types' = fold_aterms |
|
186 |
(fn Free (x, T) => Vartab.update ((x, ~1), T) |
|
187 |
| Var v => Vartab.update v |
|
188 |
| _ => I) t types; |
|
189 |
val sorts' = fold_types (fold_atyps |
|
21355 | 190 |
(fn TFree (x, S) => constrain_tvar ((x, ~1), S) |
191 |
| TVar v => constrain_tvar v |
|
20162 | 192 |
| _ => I)) t sorts; |
193 |
in (types', sorts') end) |
|
194 |
#> declare_type_names t |
|
195 |
#> redeclare_skolems; |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
196 |
|
20162 | 197 |
|
198 |
(* common declarations *) |
|
199 |
||
200 |
fun declare_internal t = |
|
201 |
declare_names t #> |
|
202 |
declare_type_occs t; |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
203 |
|
20162 | 204 |
fun declare_term t = |
205 |
declare_internal t #> |
|
206 |
declare_constraints t; |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
207 |
|
20303 | 208 |
val declare_prf = Proofterm.fold_proof_terms declare_internal (declare_internal o Logic.mk_type); |
209 |
||
20797
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents:
20579
diff
changeset
|
210 |
val declare_thm = Drule.fold_terms declare_internal; |
21522 | 211 |
fun thm_context th = declare_thm th (ProofContext.init (Thm.theory_of_thm th)); |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
212 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
213 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
214 |
(* renaming term/type frees *) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
215 |
|
20162 | 216 |
fun variant_frees ctxt ts frees = |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
217 |
let |
20162 | 218 |
val names = names_of (fold declare_names ts ctxt); |
20084 | 219 |
val xs = fst (Name.variants (map #1 frees) names); |
220 |
in xs ~~ map snd frees end; |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
221 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
222 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
223 |
|
20303 | 224 |
(** term bindings **) |
225 |
||
226 |
fun hidden_polymorphism t T = |
|
227 |
let |
|
228 |
val tvarsT = Term.add_tvarsT T []; |
|
229 |
val extra_tvars = Term.fold_types (Term.fold_atyps |
|
230 |
(fn TVar v => if member (op =) tvarsT v then I else insert (op =) v | _ => I)) t []; |
|
231 |
in extra_tvars end; |
|
232 |
||
233 |
fun add_bind (xi, NONE) = map_binds (Vartab.delete_safe xi) |
|
234 |
| add_bind ((x, i), SOME t) = |
|
235 |
let |
|
236 |
val T = Term.fastype_of t; |
|
237 |
val t' = |
|
238 |
if null (hidden_polymorphism t T) then t |
|
239 |
else Var ((x ^ "_has_extra_type_vars_on_rhs", i), T); |
|
240 |
in declare_term t' #> map_binds (Vartab.update ((x, i), (T, t'))) end; |
|
241 |
||
242 |
val add_binds = fold add_bind; |
|
243 |
||
244 |
fun expand_binds ctxt = |
|
245 |
let |
|
246 |
val binds = binds_of ctxt; |
|
247 |
fun expand (t as Var (xi, T)) = |
|
248 |
(case Vartab.lookup binds xi of |
|
249 |
SOME u => Envir.expand_atom T u |
|
250 |
| NONE => t) |
|
251 |
| expand t = t; |
|
252 |
in Envir.beta_norm o Term.map_aterms expand end; |
|
253 |
||
254 |
||
255 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
256 |
(** fixes **) |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
257 |
|
20084 | 258 |
local |
259 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
260 |
fun no_dups [] = () |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
261 |
| no_dups dups = error ("Duplicate fixed variable(s): " ^ commas_quote dups); |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
262 |
|
20102 | 263 |
fun new_fixes names' xs xs' = |
264 |
map_names (K names') #> |
|
265 |
map_fixes (fn fixes => (rev (xs ~~ xs') @ fixes)) #> |
|
20162 | 266 |
fold (declare_constraints o Syntax.free) xs' #> |
20084 | 267 |
pair xs'; |
268 |
||
269 |
in |
|
270 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
271 |
fun add_fixes xs ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
272 |
let |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
273 |
val _ = |
20084 | 274 |
(case filter (can Name.dest_skolem) xs of [] => () |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
275 |
| bads => error ("Illegal internal Skolem constant(s): " ^ commas_quote bads)); |
20084 | 276 |
val _ = no_dups (duplicates (op =) xs); |
277 |
val (ys, zs) = split_list (fixes_of ctxt); |
|
20102 | 278 |
val names = names_of ctxt; |
20084 | 279 |
val (xs', names') = |
20149 | 280 |
if is_body ctxt then Name.variants xs names |>> map Name.skolem |
20084 | 281 |
else (no_dups (xs inter_string ys); no_dups (xs inter_string zs); |
282 |
(xs, fold Name.declare xs names)); |
|
20102 | 283 |
in ctxt |> new_fixes names' xs xs' end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
284 |
|
20797
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents:
20579
diff
changeset
|
285 |
fun variant_fixes raw_xs ctxt = |
20084 | 286 |
let |
20102 | 287 |
val names = names_of ctxt; |
20149 | 288 |
val xs = map Name.clean raw_xs; |
289 |
val (xs', names') = Name.variants xs names |>> map Name.skolem; |
|
20102 | 290 |
in ctxt |> new_fixes names' xs xs' end; |
20084 | 291 |
|
292 |
end; |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
293 |
|
20251 | 294 |
|
295 |
fun add_fixes_direct xs ctxt = ctxt |
|
296 |
|> set_body false |
|
297 |
|> (snd o add_fixes xs) |
|
298 |
|> restore_body ctxt; |
|
299 |
||
300 |
fun fix_frees t ctxt = ctxt |
|
301 |
|> add_fixes_direct |
|
302 |
(rev (fold_aterms (fn Free (x, _) => |
|
21369 | 303 |
if is_fixed ctxt x then I else insert (op =) x | _ => I) t [])); |
304 |
||
305 |
fun auto_fixes t ctxt = |
|
306 |
(if is_body ctxt then ctxt else fix_frees t ctxt) |
|
20251 | 307 |
|> declare_term t; |
308 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
309 |
fun invent_types Ss ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
310 |
let |
20102 | 311 |
val tfrees = Name.invents (names_of ctxt) "'a" (length Ss) ~~ Ss; |
20162 | 312 |
val ctxt' = fold (declare_constraints o Logic.mk_type o TFree) tfrees ctxt; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
313 |
in (tfrees, ctxt') end; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
314 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
315 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
316 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
317 |
(** export -- generalize type/term variables **) |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
318 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
319 |
fun export_inst inner outer = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
320 |
let |
20162 | 321 |
val declared_outer = is_declared outer; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
322 |
val fixes_inner = fixes_of inner; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
323 |
val fixes_outer = fixes_of outer; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
324 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
325 |
val gen_fixes = map #2 (Library.take (length fixes_inner - length fixes_outer, fixes_inner)); |
20162 | 326 |
val still_fixed = not o member (op =) gen_fixes; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
327 |
val gen_fixesT = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
328 |
Symtab.fold (fn (a, xs) => |
20162 | 329 |
if declared_outer a orelse exists still_fixed xs |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
330 |
then I else cons a) (type_occs_of inner) []; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
331 |
in (gen_fixesT, gen_fixes) end; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
332 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
333 |
fun exportT_inst inner outer = #1 (export_inst inner outer); |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
334 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
335 |
fun exportT_terms inner outer ts = |
20509 | 336 |
map (TermSubst.generalize (exportT_inst (fold declare_type_occs ts inner) outer, []) |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
337 |
(fold (Term.fold_types Term.maxidx_typ) ts ~1 + 1)) ts; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
338 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
339 |
fun export_terms inner outer ts = |
20509 | 340 |
map (TermSubst.generalize (export_inst (fold declare_type_occs ts inner) outer) |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
341 |
(fold Term.maxidx_term ts ~1 + 1)) ts; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
342 |
|
20303 | 343 |
fun export_prf inner outer prf = |
344 |
let |
|
345 |
val insts = export_inst (declare_prf prf inner) outer; |
|
346 |
val idx = Proofterm.maxidx_proof prf ~1 + 1; |
|
20509 | 347 |
val gen_term = TermSubst.generalize_option insts idx; |
348 |
val gen_typ = TermSubst.generalizeT_option (#1 insts) idx; |
|
20303 | 349 |
in Proofterm.map_proof_terms_option gen_term gen_typ prf end; |
350 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
351 |
fun gen_export inst inner outer ths = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
352 |
let |
20262 | 353 |
val inner' = fold (declare_type_occs o Thm.full_prop_of) ths inner; |
354 |
in map (Thm.generalize (inst inner' outer) (fold Thm.maxidx_thm ths ~1 + 1)) ths end; |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
355 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
356 |
val exportT = gen_export (rpair [] oo exportT_inst); |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
357 |
val export = gen_export export_inst; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
358 |
|
21522 | 359 |
fun export_morphism inner outer = |
360 |
let |
|
361 |
val fact = export inner outer; |
|
362 |
val term = singleton (export_terms inner outer); |
|
363 |
val typ = Logic.type_map term; |
|
364 |
in Morphism.morphism {name = I, var = I, typ = typ, term = term, fact = fact} end; |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
365 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
366 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
367 |
(** import -- fix schematic type/term variables **) |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
368 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
369 |
fun importT_inst ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
370 |
let |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
371 |
val tvars = rev (fold Term.add_tvars ts []); |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
372 |
val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
373 |
in (tvars ~~ map TFree tfrees, ctxt') end; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
374 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
375 |
fun import_inst is_open ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
376 |
let |
20198 | 377 |
val ren = if is_open then I else Name.internal; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
378 |
val (instT, ctxt') = importT_inst ts ctxt; |
20509 | 379 |
val vars = map (apsnd (TermSubst.instantiateT instT)) (rev (fold Term.add_vars ts [])); |
20797
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents:
20579
diff
changeset
|
380 |
val (xs, ctxt'') = variant_fixes (map (ren o #1 o #1) vars) ctxt'; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
381 |
val inst = vars ~~ map Free (xs ~~ map #2 vars); |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
382 |
in ((instT, inst), ctxt'') end; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
383 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
384 |
fun importT_terms ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
385 |
let val (instT, ctxt') = importT_inst ts ctxt |
20509 | 386 |
in (map (TermSubst.instantiate (instT, [])) ts, ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
387 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
388 |
fun import_terms is_open ts ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
389 |
let val (inst, ctxt') = import_inst is_open ts ctxt |
20509 | 390 |
in (map (TermSubst.instantiate inst) ts, ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
391 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
392 |
fun importT ths ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
393 |
let |
21522 | 394 |
val thy = ProofContext.theory_of ctxt; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
395 |
val certT = Thm.ctyp_of thy; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
396 |
val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
397 |
val instT' = map (fn (v, T) => (certT (TVar v), certT T)) instT; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
398 |
val ths' = map (Thm.instantiate (instT', [])) ths; |
20220 | 399 |
in ((map #2 instT', ths'), ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
400 |
|
20303 | 401 |
fun import_prf is_open prf ctxt = |
402 |
let |
|
403 |
val ts = rev (Proofterm.fold_proof_terms cons (cons o Logic.mk_type) prf []); |
|
404 |
val (insts, ctxt') = import_inst is_open ts ctxt; |
|
405 |
in (Proofterm.instantiate insts prf, ctxt') end; |
|
406 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
407 |
fun import is_open ths ctxt = |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
408 |
let |
21522 | 409 |
val thy = ProofContext.theory_of ctxt; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
410 |
val cert = Thm.cterm_of thy; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
411 |
val certT = Thm.ctyp_of thy; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
412 |
val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
413 |
val instT' = map (fn (v, T) => (certT (TVar v), certT T)) instT; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
414 |
val inst' = map (fn (v, t) => (cert (Var v), cert t)) inst; |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
415 |
val ths' = map (Thm.instantiate (instT', inst')) ths; |
20220 | 416 |
in (((map #2 instT', map #2 inst'), ths'), ctxt') end; |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
417 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
418 |
|
19926 | 419 |
(* import/export *) |
420 |
||
21287 | 421 |
fun gen_trade imp exp f ctxt ths = |
20220 | 422 |
let val ((_, ths'), ctxt') = imp ths ctxt |
21287 | 423 |
in exp ctxt' ctxt (f ctxt' ths') end; |
19926 | 424 |
|
425 |
val tradeT = gen_trade importT exportT; |
|
426 |
val trade = gen_trade (import true) export; |
|
427 |
||
428 |
||
20162 | 429 |
(* focus on outermost parameters *) |
20149 | 430 |
|
431 |
fun forall_elim_prop t prop = |
|
20579 | 432 |
Thm.beta_conversion false (Thm.capply (Thm.dest_arg prop) t) |
433 |
|> Thm.cprop_of |> Thm.dest_arg; |
|
20149 | 434 |
|
435 |
fun focus goal ctxt = |
|
436 |
let |
|
437 |
val cert = Thm.cterm_of (Thm.theory_of_cterm goal); |
|
438 |
val t = Thm.term_of goal; |
|
20162 | 439 |
val ps = Term.variant_frees t (Term.strip_all_vars t); (*as they are printed :-*) |
20149 | 440 |
val (xs, Ts) = split_list ps; |
20797
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents:
20579
diff
changeset
|
441 |
val (xs', ctxt') = variant_fixes xs ctxt; |
20220 | 442 |
val ps' = ListPair.map (cert o Free) (xs', Ts); |
443 |
val goal' = fold forall_elim_prop ps' goal; |
|
20149 | 444 |
in ((ps', goal'), ctxt') end; |
445 |
||
20303 | 446 |
fun focus_subgoal i st = |
447 |
let |
|
448 |
val all_vars = Drule.fold_terms Term.add_vars st []; |
|
449 |
val no_binds = map (fn (xi, _) => (xi, NONE)) all_vars; |
|
450 |
in |
|
451 |
add_binds no_binds #> |
|
452 |
fold (declare_constraints o Var) all_vars #> |
|
453 |
focus (Thm.cprem_of st i) |
|
454 |
end; |
|
455 |
||
456 |
||
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
457 |
|
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
458 |
(** implicit polymorphism **) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
459 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
460 |
(* warn_extra_tfrees *) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
461 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
462 |
fun warn_extra_tfrees ctxt1 ctxt2 = |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
463 |
let |
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
464 |
fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); |
20162 | 465 |
fun occs_free a x = |
466 |
(case def_type ctxt1 false (x, ~1) of |
|
467 |
SOME T => if occs_typ a T then I else cons (a, x) |
|
468 |
| NONE => cons (a, x)); |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
469 |
|
20162 | 470 |
val occs1 = type_occs_of ctxt1; |
471 |
val occs2 = type_occs_of ctxt2; |
|
19911
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
472 |
val extras = Symtab.fold (fn (a, xs) => |
300bc6ce970d
major reworking of export functionality -- based on Term/Thm.generalize;
wenzelm
parents:
19899
diff
changeset
|
473 |
if Symtab.defined occs1 a then I else fold (occs_free a) xs) occs2 []; |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
474 |
val tfrees = map #1 extras |> sort_distinct string_ord; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
475 |
val frees = map #2 extras |> sort_distinct string_ord; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
476 |
in |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
477 |
if null extras then () |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
478 |
else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^ |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
479 |
space_implode " or " (map quote frees)) |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
480 |
end; |
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
481 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
482 |
|
20149 | 483 |
(* polymorphic terms *) |
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
484 |
|
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
485 |
fun polymorphic ctxt ts = |
20003 | 486 |
let |
487 |
val ctxt' = fold declare_term ts ctxt; |
|
20162 | 488 |
val occs = type_occs_of ctxt; |
489 |
val occs' = type_occs_of ctxt'; |
|
490 |
val types = Symtab.fold (fn (a, _) => if Symtab.defined occs a then I else cons a) occs' []; |
|
20003 | 491 |
val idx = fold (Term.fold_types Term.maxidx_typ) ts ~1 + 1; |
20509 | 492 |
in map (TermSubst.generalize (types, []) idx) ts end; |
20003 | 493 |
|
19899
b7385ca02d79
Fixed type/term variables and polymorphic term abbreviations.
wenzelm
parents:
diff
changeset
|
494 |
end; |