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