| author | Manuel Eberl <eberlm@in.tum.de> | 
| Wed, 27 Jul 2016 10:44:22 +0200 | |
| changeset 63552 | 2112e5fe9712 | 
| parent 61039 | 80f40d89dab6 | 
| child 74282 | c2ee8d993d6a | 
| permissions | -rw-r--r-- | 
| 47455 | 1  | 
(* Title: HOL/Matrix_LP/Compute_Oracle/linker.ML  | 
| 23663 | 2  | 
Author: Steven Obua  | 
3  | 
||
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
4  | 
This module solves the problem that the computing oracle does not  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
5  | 
instantiate polymorphic rules. By going through the PCompute  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
6  | 
interface, all possible instantiations are resolved by compiling new  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
7  | 
programs, if necessary. The obvious disadvantage of this approach is  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
8  | 
that in the worst case for each new term to be rewritten, a new  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
9  | 
program may be compiled.  | 
| 23663 | 10  | 
*)  | 
11  | 
||
| 24271 | 12  | 
(*  | 
| 23663 | 13  | 
Given constants/frees c_1::t_1, c_2::t_2, ...., c_n::t_n,  | 
14  | 
and constants/frees d_1::d_1, d_2::s_2, ..., d_m::s_m  | 
|
15  | 
||
16  | 
Find all substitutions S such that  | 
|
| 24271 | 17  | 
a) the domain of S is tvars (t_1, ..., t_n)  | 
| 23663 | 18  | 
b) there are indices i_1, ..., i_k, and j_1, ..., j_k with  | 
19  | 
1. S (c_i_1::t_i_1) = d_j_1::s_j_1, ..., S (c_i_k::t_i_k) = d_j_k::s_j_k  | 
|
20  | 
2. tvars (t_i_1, ..., t_i_k) = tvars (t_1, ..., t_n)  | 
|
21  | 
*)  | 
|
| 24271 | 22  | 
signature LINKER =  | 
| 23663 | 23  | 
sig  | 
24  | 
exception Link of string  | 
|
| 24271 | 25  | 
|
| 23663 | 26  | 
datatype constant = Constant of bool * string * typ  | 
27  | 
val constant_of : term -> constant  | 
|
28  | 
||
29  | 
type instances  | 
|
30  | 
type subst = Type.tyenv  | 
|
| 24271 | 31  | 
|
| 23663 | 32  | 
val empty : constant list -> instances  | 
33  | 
val typ_of_constant : constant -> typ  | 
|
| 24271 | 34  | 
val add_instances : theory -> instances -> constant list -> subst list * instances  | 
| 23663 | 35  | 
val substs_of : instances -> subst list  | 
36  | 
val is_polymorphic : constant -> bool  | 
|
37  | 
val distinct_constants : constant list -> constant list  | 
|
38  | 
val collect_consts : term list -> constant list  | 
|
39  | 
end  | 
|
40  | 
||
41  | 
structure Linker : LINKER = struct  | 
|
42  | 
||
43  | 
exception Link of string;  | 
|
44  | 
||
45  | 
type subst = Type.tyenv  | 
|
46  | 
||
47  | 
datatype constant = Constant of bool * string * typ  | 
|
48  | 
fun constant_of (Const (name, ty)) = Constant (false, name, ty)  | 
|
49  | 
| constant_of (Free (name, ty)) = Constant (true, name, ty)  | 
|
50  | 
| constant_of _ = raise Link "constant_of"  | 
|
51  | 
||
52  | 
fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL)  | 
|
| 35408 | 53  | 
fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term_Ord.typ_ord) (((x1,x2),x3), ((y1,y2),y3))  | 
| 23663 | 54  | 
fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2))  | 
55  | 
||
56  | 
||
| 
31971
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
57  | 
structure Consttab = Table(type key = constant val ord = constant_ord);  | 
| 
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
58  | 
structure ConsttabModTy = Table(type key = constant val ord = constant_modty_ord);  | 
| 23663 | 59  | 
|
60  | 
fun typ_of_constant (Constant (_, _, ty)) = ty  | 
|
61  | 
||
62  | 
val empty_subst = (Vartab.empty : Type.tyenv)  | 
|
63  | 
||
| 24271 | 64  | 
fun merge_subst (A:Type.tyenv) (B:Type.tyenv) =  | 
65  | 
SOME (Vartab.fold (fn (v, t) =>  | 
|
66  | 
fn tab =>  | 
|
67  | 
(case Vartab.lookup tab v of  | 
|
68  | 
NONE => Vartab.update (v, t) tab  | 
|
69  | 
| SOME t' => if t = t' then tab else raise Type.TYPE_MATCH)) A B)  | 
|
| 23663 | 70  | 
handle Type.TYPE_MATCH => NONE  | 
71  | 
||
| 24271 | 72  | 
fun subst_ord (A:Type.tyenv, B:Type.tyenv) =  | 
| 35408 | 73  | 
(list_ord (prod_ord Term_Ord.fast_indexname_ord (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord))) (Vartab.dest A, Vartab.dest B)  | 
| 23663 | 74  | 
|
| 
31971
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
75  | 
structure Substtab = Table(type key = Type.tyenv val ord = subst_ord);  | 
| 23663 | 76  | 
|
| 23768 | 77  | 
fun substtab_union c = Substtab.fold Substtab.update c  | 
| 23663 | 78  | 
fun substtab_unions [] = Substtab.empty  | 
79  | 
| substtab_unions [c] = c  | 
|
80  | 
| substtab_unions (c::cs) = substtab_union c (substtab_unions cs)  | 
|
81  | 
||
82  | 
datatype instances = Instances of unit ConsttabModTy.table * Type.tyenv Consttab.table Consttab.table * constant list list * unit Substtab.table  | 
|
83  | 
||
| 
29275
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
84  | 
fun is_polymorphic (Constant (_, _, ty)) = not (null (Term.add_tvarsT ty []))  | 
| 23663 | 85  | 
|
86  | 
fun distinct_constants cs =  | 
|
87  | 
Consttab.keys (fold (fn c => Consttab.update (c, ())) cs Consttab.empty)  | 
|
88  | 
||
| 24271 | 89  | 
fun empty cs =  | 
90  | 
let  | 
|
91  | 
val cs = distinct_constants (filter is_polymorphic cs)  | 
|
92  | 
val old_cs = cs  | 
|
| 44121 | 93  | 
(* fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (Misc_Legacy.typ_tvars ty) tab  | 
| 24271 | 94  | 
val tvars_count = length (Typtab.keys (fold (fn c => fn tab => collect_tvars (typ_of_constant c) tab) cs Typtab.empty))  | 
95  | 
fun tvars_of ty = collect_tvars ty Typtab.empty  | 
|
96  | 
val cs = map (fn c => (c, tvars_of (typ_of_constant c))) cs  | 
|
| 23663 | 97  | 
|
| 24271 | 98  | 
fun tyunion A B =  | 
99  | 
Typtab.fold  | 
|
100  | 
(fn (v,()) => fn tab => Typtab.update (v, case Typtab.lookup tab v of NONE => 1 | SOME n => n+1) tab)  | 
|
101  | 
A B  | 
|
| 23663 | 102  | 
|
103  | 
fun is_essential A B =  | 
|
| 24271 | 104  | 
Typtab.fold  | 
105  | 
(fn (v, ()) => fn essential => essential orelse (case Typtab.lookup B v of NONE => raise Link "is_essential" | SOME n => n=1))  | 
|
106  | 
A false  | 
|
| 23663 | 107  | 
|
| 24271 | 108  | 
fun add_minimal (c', tvs') (tvs, cs) =  | 
109  | 
let  | 
|
110  | 
val tvs = tyunion tvs' tvs  | 
|
111  | 
val cs = (c', tvs')::cs  | 
|
112  | 
in  | 
|
113  | 
if forall (fn (c',tvs') => is_essential tvs' tvs) cs then  | 
|
114  | 
SOME (tvs, cs)  | 
|
115  | 
else  | 
|
116  | 
NONE  | 
|
117  | 
end  | 
|
| 23663 | 118  | 
|
| 24271 | 119  | 
fun is_spanning (tvs, _) = (length (Typtab.keys tvs) = tvars_count)  | 
| 23663 | 120  | 
|
| 24271 | 121  | 
fun generate_minimal_subsets subsets [] = subsets  | 
122  | 
| generate_minimal_subsets subsets (c::cs) =  | 
|
123  | 
let  | 
|
124  | 
val subsets' = map_filter (add_minimal c) subsets  | 
|
125  | 
in  | 
|
126  | 
generate_minimal_subsets (subsets@subsets') cs  | 
|
127  | 
end*)  | 
|
| 23663 | 128  | 
|
| 24271 | 129  | 
val minimal_subsets = [old_cs] (*map (fn (tvs, cs) => map fst cs) (filter is_spanning (generate_minimal_subsets [(Typtab.empty, [])] cs))*)  | 
| 23663 | 130  | 
|
| 24271 | 131  | 
val constants = Consttab.keys (fold (fold (fn c => Consttab.update (c, ()))) minimal_subsets Consttab.empty)  | 
| 23663 | 132  | 
|
133  | 
in  | 
|
| 24271 | 134  | 
Instances (  | 
135  | 
fold (fn c => fn tab => ConsttabModTy.update (c, ()) tab) constants ConsttabModTy.empty,  | 
|
136  | 
Consttab.make (map (fn c => (c, Consttab.empty : Type.tyenv Consttab.table)) constants),  | 
|
137  | 
minimal_subsets, Substtab.empty)  | 
|
| 23663 | 138  | 
end  | 
139  | 
||
140  | 
local  | 
|
141  | 
fun calc ctab substtab [] = substtab  | 
|
| 24271 | 142  | 
| calc ctab substtab (c::cs) =  | 
| 23663 | 143  | 
let  | 
| 24271 | 144  | 
val csubsts = map snd (Consttab.dest (the (Consttab.lookup ctab c)))  | 
145  | 
fun merge_substs substtab subst =  | 
|
146  | 
Substtab.fold (fn (s,_) =>  | 
|
147  | 
fn tab =>  | 
|
148  | 
(case merge_subst subst s of NONE => tab | SOME s => Substtab.update (s, ()) tab))  | 
|
149  | 
substtab Substtab.empty  | 
|
150  | 
val substtab = substtab_unions (map (merge_substs substtab) csubsts)  | 
|
| 23663 | 151  | 
in  | 
| 24271 | 152  | 
calc ctab substtab cs  | 
| 23663 | 153  | 
end  | 
154  | 
in  | 
|
155  | 
fun calc_substs ctab (cs:constant list) = calc ctab (Substtab.update (empty_subst, ()) Substtab.empty) cs  | 
|
156  | 
end  | 
|
| 24271 | 157  | 
|
158  | 
fun add_instances thy (Instances (cfilter, ctab,minsets,substs)) cs =  | 
|
159  | 
let  | 
|
160  | 
(*      val _ = writeln (makestring ("add_instances: ", length_cs, length cs, length (Consttab.keys ctab)))*)
 | 
|
161  | 
fun calc_instantiations (constant as Constant (free, name, ty)) instantiations =  | 
|
162  | 
Consttab.fold (fn (constant' as Constant (free', name', ty'), insttab) =>  | 
|
163  | 
fn instantiations =>  | 
|
164  | 
if free <> free' orelse name <> name' then  | 
|
165  | 
instantiations  | 
|
166  | 
else case Consttab.lookup insttab constant of  | 
|
167  | 
SOME _ => instantiations  | 
|
168  | 
| NONE => ((constant', (constant, Sign.typ_match thy (ty', ty) empty_subst))::instantiations  | 
|
| 
39685
 
d8071cddb877
actually handle Type.TYPE_MATCH, not arbitrary exceptions;
 
wenzelm 
parents: 
37872 
diff
changeset
 | 
169  | 
handle Type.TYPE_MATCH => instantiations))  | 
| 24271 | 170  | 
ctab instantiations  | 
171  | 
val instantiations = fold calc_instantiations cs []  | 
|
172  | 
        (*val _ = writeln ("instantiations = "^(makestring (length instantiations)))*)
 | 
|
173  | 
fun update_ctab (constant', entry) ctab =  | 
|
174  | 
(case Consttab.lookup ctab constant' of  | 
|
175  | 
NONE => raise Link "internal error: update_ctab"  | 
|
176  | 
| SOME tab => Consttab.update (constant', Consttab.update entry tab) ctab)  | 
|
177  | 
val ctab = fold update_ctab instantiations ctab  | 
|
178  | 
val new_substs = fold (fn minset => fn substs => substtab_union (calc_substs ctab minset) substs)  | 
|
179  | 
minsets Substtab.empty  | 
|
180  | 
val (added_substs, substs) =  | 
|
181  | 
Substtab.fold (fn (ns, _) =>  | 
|
182  | 
fn (added, substtab) =>  | 
|
183  | 
(case Substtab.lookup substs ns of  | 
|
184  | 
NONE => (ns::added, Substtab.update (ns, ()) substtab)  | 
|
185  | 
| SOME () => (added, substtab)))  | 
|
186  | 
new_substs ([], substs)  | 
|
| 23663 | 187  | 
in  | 
| 24271 | 188  | 
(added_substs, Instances (cfilter, ctab, minsets, substs))  | 
| 23663 | 189  | 
end  | 
| 24271 | 190  | 
|
| 23663 | 191  | 
fun substs_of (Instances (_,_,_,substs)) = Substtab.keys substs  | 
192  | 
||
| 24271 | 193  | 
|
| 23663 | 194  | 
local  | 
195  | 
||
| 46531 | 196  | 
fun collect (Var _) tab = tab  | 
| 23663 | 197  | 
| collect (Bound _) tab = tab  | 
198  | 
| collect (a $ b) tab = collect b (collect a tab)  | 
|
199  | 
| collect (Abs (_, _, body)) tab = collect body tab  | 
|
200  | 
| collect t tab = Consttab.update (constant_of t, ()) tab  | 
|
201  | 
||
202  | 
in  | 
|
203  | 
fun collect_consts tms = Consttab.keys (fold collect tms Consttab.empty)  | 
|
204  | 
end  | 
|
205  | 
||
206  | 
end  | 
|
207  | 
||
208  | 
signature PCOMPUTE =  | 
|
209  | 
sig  | 
|
210  | 
type pcomputer  | 
|
| 24271 | 211  | 
|
| 23663 | 212  | 
val make : Compute.machine -> theory -> thm list -> Linker.constant list -> pcomputer  | 
| 25520 | 213  | 
val make_with_cache : Compute.machine -> theory -> term list -> thm list -> Linker.constant list -> pcomputer  | 
| 25218 | 214  | 
|
215  | 
val add_instances : pcomputer -> Linker.constant list -> bool  | 
|
216  | 
val add_instances' : pcomputer -> term list -> bool  | 
|
| 23663 | 217  | 
|
218  | 
val rewrite : pcomputer -> cterm list -> thm list  | 
|
| 25218 | 219  | 
val simplify : pcomputer -> Compute.theorem -> thm  | 
220  | 
||
221  | 
val make_theorem : pcomputer -> thm -> string list -> Compute.theorem  | 
|
222  | 
val instantiate : pcomputer -> (string * cterm) list -> Compute.theorem -> Compute.theorem  | 
|
223  | 
val evaluate_prem : pcomputer -> int -> Compute.theorem -> Compute.theorem  | 
|
224  | 
val modus_ponens : pcomputer -> int -> thm -> Compute.theorem -> Compute.theorem  | 
|
| 23663 | 225  | 
|
226  | 
end  | 
|
227  | 
||
228  | 
structure PCompute : PCOMPUTE = struct  | 
|
229  | 
||
230  | 
exception PCompute of string  | 
|
231  | 
||
232  | 
datatype theorem = MonoThm of thm | PolyThm of thm * Linker.instances * thm list  | 
|
| 25520 | 233  | 
datatype pattern = MonoPattern of term | PolyPattern of term * Linker.instances * term list  | 
| 23663 | 234  | 
|
| 32740 | 235  | 
datatype pcomputer =  | 
| 52788 | 236  | 
PComputer of theory * Compute.computer * theorem list Unsynchronized.ref *  | 
| 32740 | 237  | 
pattern list Unsynchronized.ref  | 
| 23663 | 238  | 
|
239  | 
(*fun collect_consts (Var x) = []  | 
|
240  | 
| collect_consts (Bound _) = []  | 
|
241  | 
| collect_consts (a $ b) = (collect_consts a)@(collect_consts b)  | 
|
242  | 
| collect_consts (Abs (_, _, body)) = collect_consts body  | 
|
243  | 
| collect_consts t = [Linker.constant_of t]*)  | 
|
244  | 
||
| 25520 | 245  | 
fun computer_of (PComputer (_,computer,_,_)) = computer  | 
| 25218 | 246  | 
|
247  | 
fun collect_consts_of_thm th =  | 
|
| 23663 | 248  | 
let  | 
| 59582 | 249  | 
val th = Thm.prop_of th  | 
| 24271 | 250  | 
val (prems, th) = (Logic.strip_imp_prems th, Logic.strip_imp_concl th)  | 
251  | 
val (left, right) = Logic.dest_equals th  | 
|
| 23663 | 252  | 
in  | 
| 24271 | 253  | 
(Linker.collect_consts [left], Linker.collect_consts (right::prems))  | 
254  | 
end  | 
|
| 23663 | 255  | 
|
256  | 
fun create_theorem th =  | 
|
| 24271 | 257  | 
let  | 
| 23663 | 258  | 
val (left, right) = collect_consts_of_thm th  | 
259  | 
val polycs = filter Linker.is_polymorphic left  | 
|
| 44121 | 260  | 
val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (Misc_Legacy.typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty  | 
| 24271 | 261  | 
fun check_const (c::cs) cs' =  | 
262  | 
let  | 
|
| 44121 | 263  | 
val tvars = Misc_Legacy.typ_tvars (Linker.typ_of_constant c)  | 
| 24271 | 264  | 
val wrong = fold (fn n => fn wrong => wrong orelse is_none (Typtab.lookup tytab (TVar n))) tvars false  | 
265  | 
in  | 
|
266  | 
if wrong then raise PCompute "right hand side of theorem contains type variables which do not occur on the left hand side"  | 
|
267  | 
else  | 
|
268  | 
if null (tvars) then  | 
|
269  | 
check_const cs (c::cs')  | 
|
270  | 
else  | 
|
271  | 
check_const cs cs'  | 
|
272  | 
end  | 
|
| 23663 | 273  | 
| check_const [] cs' = cs'  | 
| 24271 | 274  | 
val monocs = check_const right []  | 
| 23663 | 275  | 
in  | 
| 24271 | 276  | 
if null (polycs) then  | 
277  | 
(monocs, MonoThm th)  | 
|
| 23663 | 278  | 
else  | 
| 24271 | 279  | 
(monocs, PolyThm (th, Linker.empty polycs, []))  | 
| 23663 | 280  | 
end  | 
281  | 
||
| 25520 | 282  | 
fun create_pattern pat =  | 
283  | 
let  | 
|
284  | 
val cs = Linker.collect_consts [pat]  | 
|
285  | 
val polycs = filter Linker.is_polymorphic cs  | 
|
286  | 
in  | 
|
287  | 
if null (polycs) then  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
288  | 
MonoPattern pat  | 
| 25520 | 289  | 
else  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
290  | 
PolyPattern (pat, Linker.empty polycs, [])  | 
| 25520 | 291  | 
end  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
292  | 
|
| 25520 | 293  | 
fun create_computer machine thy pats ths =  | 
| 23663 | 294  | 
let  | 
| 24271 | 295  | 
fun add (MonoThm th) ths = th::ths  | 
296  | 
| add (PolyThm (_, _, ths')) ths = ths'@ths  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
297  | 
fun addpat (MonoPattern p) pats = p::pats  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
298  | 
| addpat (PolyPattern (_, _, ps)) pats = ps@pats  | 
| 24271 | 299  | 
val ths = fold_rev add ths []  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
300  | 
val pats = fold_rev addpat pats []  | 
| 23663 | 301  | 
in  | 
| 25520 | 302  | 
Compute.make_with_cache machine thy pats ths  | 
| 23663 | 303  | 
end  | 
304  | 
||
| 25520 | 305  | 
fun update_computer computer pats ths =  | 
| 25218 | 306  | 
let  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
307  | 
fun add (MonoThm th) ths = th::ths  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
308  | 
| add (PolyThm (_, _, ths')) ths = ths'@ths  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
309  | 
fun addpat (MonoPattern p) pats = p::pats  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
310  | 
| addpat (PolyPattern (_, _, ps)) pats = ps@pats  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
311  | 
val ths = fold_rev add ths []  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
312  | 
val pats = fold_rev addpat pats []  | 
| 25218 | 313  | 
in  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
314  | 
Compute.update_with_cache computer pats ths  | 
| 25218 | 315  | 
end  | 
316  | 
||
| 23663 | 317  | 
fun conv_subst thy (subst : Type.tyenv) =  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
59621 
diff
changeset
 | 
318  | 
map (fn (iname, (sort, ty)) => ((iname, sort), Thm.global_ctyp_of thy ty))  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
59621 
diff
changeset
 | 
319  | 
(Vartab.dest subst)  | 
| 23663 | 320  | 
|
| 25520 | 321  | 
fun add_monos thy monocs pats ths =  | 
| 23663 | 322  | 
let  | 
| 32740 | 323  | 
val changed = Unsynchronized.ref false  | 
| 24271 | 324  | 
fun add monocs (th as (MonoThm _)) = ([], th)  | 
325  | 
| add monocs (PolyThm (th, instances, instanceths)) =  | 
|
326  | 
let  | 
|
327  | 
val (newsubsts, instances) = Linker.add_instances thy instances monocs  | 
|
328  | 
val _ = if not (null newsubsts) then changed := true else ()  | 
|
329  | 
val newths = map (fn subst => Thm.instantiate (conv_subst thy subst, []) th) newsubsts  | 
|
330  | 
(*              val _ = if not (null newths) then (print ("added new theorems: ", newths); ()) else ()*)
 | 
|
331  | 
val newmonos = fold (fn th => fn monos => (snd (collect_consts_of_thm th))@monos) newths []  | 
|
332  | 
in  | 
|
333  | 
(newmonos, PolyThm (th, instances, instanceths@newths))  | 
|
334  | 
end  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
335  | 
fun addpats monocs (pat as (MonoPattern _)) = pat  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
336  | 
| addpats monocs (PolyPattern (p, instances, instancepats)) =  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
337  | 
let  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
338  | 
val (newsubsts, instances) = Linker.add_instances thy instances monocs  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
339  | 
val _ = if not (null newsubsts) then changed := true else ()  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
340  | 
val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
341  | 
in  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
342  | 
PolyPattern (p, instances, instancepats@newpats)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
343  | 
end  | 
| 24271 | 344  | 
fun step monocs ths =  | 
345  | 
fold_rev (fn th =>  | 
|
346  | 
fn (newmonos, ths) =>  | 
|
| 25520 | 347  | 
let  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
348  | 
val (newmonos', th') = add monocs th  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
349  | 
in  | 
| 24271 | 350  | 
(newmonos'@newmonos, th'::ths)  | 
351  | 
end)  | 
|
352  | 
ths ([], [])  | 
|
| 25520 | 353  | 
fun loop monocs pats ths =  | 
354  | 
let  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
355  | 
val (monocs', ths') = step monocs ths  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
356  | 
val pats' = map (addpats monocs) pats  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
357  | 
in  | 
| 24271 | 358  | 
if null (monocs') then  | 
| 25520 | 359  | 
(pats', ths')  | 
| 24271 | 360  | 
else  | 
| 25520 | 361  | 
loop monocs' pats' ths'  | 
| 24271 | 362  | 
end  | 
| 25520 | 363  | 
val result = loop monocs pats ths  | 
| 23663 | 364  | 
in  | 
| 24271 | 365  | 
(!changed, result)  | 
366  | 
end  | 
|
| 23663 | 367  | 
|
368  | 
datatype cthm = ComputeThm of term list * sort list * term  | 
|
369  | 
||
| 61039 | 370  | 
fun thm2cthm th = ComputeThm (Thm.hyps_of th, Thm.shyps_of th, Thm.prop_of th)  | 
| 23663 | 371  | 
|
| 61039 | 372  | 
val cthm_ord' =  | 
373  | 
prod_ord (prod_ord (list_ord Term_Ord.term_ord) (list_ord Term_Ord.sort_ord)) Term_Ord.term_ord  | 
|
| 23663 | 374  | 
|
| 61039 | 375  | 
fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) =  | 
376  | 
cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2))  | 
|
| 23663 | 377  | 
|
| 
31971
 
8c1b845ed105
renamed functor TableFun to Table, and GraphFun to Graph;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
378  | 
structure CThmtab = Table(type key = cthm val ord = cthm_ord)  | 
| 24271 | 379  | 
|
| 23663 | 380  | 
fun remove_duplicates ths =  | 
381  | 
let  | 
|
| 32740 | 382  | 
val counter = Unsynchronized.ref 0  | 
383  | 
val tab = Unsynchronized.ref (CThmtab.empty : unit CThmtab.table)  | 
|
384  | 
val thstab = Unsynchronized.ref (Inttab.empty : thm Inttab.table)  | 
|
| 24271 | 385  | 
fun update th =  | 
386  | 
let  | 
|
387  | 
val key = thm2cthm th  | 
|
388  | 
in  | 
|
389  | 
case CThmtab.lookup (!tab) key of  | 
|
390  | 
NONE => ((tab := CThmtab.update_new (key, ()) (!tab)); thstab := Inttab.update_new (!counter, th) (!thstab); counter := !counter + 1)  | 
|
391  | 
| _ => ()  | 
|
392  | 
end  | 
|
393  | 
val _ = map update ths  | 
|
| 23663 | 394  | 
in  | 
| 24271 | 395  | 
map snd (Inttab.dest (!thstab))  | 
| 23663 | 396  | 
end  | 
| 24271 | 397  | 
|
| 25520 | 398  | 
fun make_with_cache machine thy pats ths cs =  | 
| 23663 | 399  | 
let  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
400  | 
val ths = remove_duplicates ths  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
401  | 
val (monocs, ths) = fold_rev (fn th =>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
402  | 
fn (monocs, ths) =>  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
403  | 
let val (m, t) = create_theorem th in  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
404  | 
(m@monocs, t::ths)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
405  | 
end)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
406  | 
ths (cs, [])  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
407  | 
val pats = map create_pattern pats  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
408  | 
val (_, (pats, ths)) = add_monos thy monocs pats ths  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
409  | 
val computer = create_computer machine thy pats ths  | 
| 23663 | 410  | 
in  | 
| 52788 | 411  | 
PComputer (thy, computer, Unsynchronized.ref ths, Unsynchronized.ref pats)  | 
| 23663 | 412  | 
end  | 
413  | 
||
| 25520 | 414  | 
fun make machine thy ths cs = make_with_cache machine thy [] ths cs  | 
415  | 
||
| 52788 | 416  | 
fun add_instances (PComputer (thy, computer, rths, rpats)) cs =  | 
| 23663 | 417  | 
let  | 
| 25520 | 418  | 
val (changed, (pats, ths)) = add_monos thy cs (!rpats) (!rths)  | 
| 23663 | 419  | 
in  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
420  | 
if changed then  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
421  | 
(update_computer computer pats ths;  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
422  | 
rths := ths;  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
423  | 
rpats := pats;  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
424  | 
true)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
425  | 
else  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
426  | 
false  | 
| 25218 | 427  | 
|
428  | 
end  | 
|
429  | 
||
430  | 
fun add_instances' pc ts = add_instances pc (Linker.collect_consts ts)  | 
|
431  | 
||
432  | 
fun rewrite pc cts =  | 
|
433  | 
let  | 
|
| 59582 | 434  | 
val _ = add_instances' pc (map Thm.term_of cts)  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
435  | 
val computer = (computer_of pc)  | 
| 25218 | 436  | 
in  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
437  | 
map (fn ct => Compute.rewrite computer ct) cts  | 
| 23663 | 438  | 
end  | 
439  | 
||
| 25218 | 440  | 
fun simplify pc th = Compute.simplify (computer_of pc) th  | 
441  | 
||
442  | 
fun make_theorem pc th vars =  | 
|
| 23663 | 443  | 
let  | 
| 59582 | 444  | 
val _ = add_instances' pc [Thm.prop_of th]  | 
| 25218 | 445  | 
|
| 23663 | 446  | 
in  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
447  | 
Compute.make_theorem (computer_of pc) th vars  | 
| 23663 | 448  | 
end  | 
| 24271 | 449  | 
|
| 25218 | 450  | 
fun instantiate pc insts th =  | 
451  | 
let  | 
|
| 59582 | 452  | 
val _ = add_instances' pc (map (Thm.term_of o snd) insts)  | 
| 25218 | 453  | 
in  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
454  | 
Compute.instantiate (computer_of pc) insts th  | 
| 25218 | 455  | 
end  | 
456  | 
||
457  | 
fun evaluate_prem pc prem_no th = Compute.evaluate_prem (computer_of pc) prem_no th  | 
|
458  | 
||
459  | 
fun modus_ponens pc prem_no th' th =  | 
|
460  | 
let  | 
|
| 59582 | 461  | 
val _ = add_instances' pc [Thm.prop_of th']  | 
| 25218 | 462  | 
in  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
463  | 
Compute.modus_ponens (computer_of pc) prem_no th' th  | 
| 25218 | 464  | 
end  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32740 
diff
changeset
 | 
465  | 
|
| 25218 | 466  | 
|
| 
24137
 
8d7896398147
replaced Theory.self_ref by Theory.check_thy, which now produces a checked ref;
 
wenzelm 
parents: 
23768 
diff
changeset
 | 
467  | 
end  |