| author | wenzelm | 
| Tue, 25 Jul 2006 23:17:41 +0200 | |
| changeset 20205 | 7b2958d3d575 | 
| parent 20156 | 7a7898b1cfa4 | 
| child 21648 | c8a0370c9b93 | 
| permissions | -rw-r--r-- | 
| 13270 | 1  | 
(* Title: Pure/fact_index.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Markus Wenzel, TU Muenchen  | 
|
4  | 
||
| 16020 | 5  | 
Facts indexed by consts or frees.  | 
| 13270 | 6  | 
*)  | 
7  | 
||
8  | 
signature FACT_INDEX =  | 
|
9  | 
sig  | 
|
| 16020 | 10  | 
type spec  | 
| 13270 | 11  | 
type T  | 
| 16020 | 12  | 
val facts: T -> (string * thm list) list  | 
| 20010 | 13  | 
val props: T -> thm list  | 
| 18026 | 14  | 
val could_unify: T -> term -> thm list  | 
| 13270 | 15  | 
val empty: T  | 
| 18026 | 16  | 
val add_global: (string * thm list) -> T -> T  | 
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
17  | 
val add_local: bool -> (string * thm list) -> T -> T  | 
| 16020 | 18  | 
val find: T -> spec -> (string * thm list) list  | 
| 13270 | 19  | 
end;  | 
20  | 
||
21  | 
structure FactIndex: FACT_INDEX =  | 
|
22  | 
struct  | 
|
23  | 
||
| 16020 | 24  | 
|
| 13270 | 25  | 
(* indexing items *)  | 
26  | 
||
| 18026 | 27  | 
type spec = string list * string list;  | 
28  | 
||
| 16874 | 29  | 
val add_consts = fold_aterms  | 
30  | 
(fn Const (c, _) => if c = Term.dummy_patternN then I else insert (op =) c | _ => I);  | 
|
| 13270 | 31  | 
|
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
32  | 
val add_frees = fold_aterms (fn Free (x, _) => insert (op =) x | _ => I);  | 
| 16874 | 33  | 
|
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
34  | 
fun index_term t (cs, xs) = (add_consts t cs, add_frees t xs);  | 
| 13270 | 35  | 
|
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
36  | 
fun index_thm th =  | 
| 16020 | 37  | 
  let val {tpairs, hyps, prop, ...} = Thm.rep_thm th in
 | 
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
38  | 
index_term prop  | 
| 
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
39  | 
#> fold index_term hyps  | 
| 
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
40  | 
#> fold (fn (t, u) => index_term t #> index_term u) tpairs  | 
| 16020 | 41  | 
end;  | 
| 13270 | 42  | 
|
43  | 
||
44  | 
(* build index *)  | 
|
45  | 
||
| 18026 | 46  | 
type fact = string * thm list;  | 
47  | 
||
48  | 
datatype T = Index of  | 
|
49  | 
 {facts: fact list,
 | 
|
50  | 
consts: (serial * fact) list Symtab.table,  | 
|
51  | 
frees: (serial * fact) list Symtab.table,  | 
|
52  | 
props: thm Net.net};  | 
|
| 13270 | 53  | 
|
| 16020 | 54  | 
fun facts (Index {facts, ...}) = facts;
 | 
| 20010 | 55  | 
fun props (Index {props, ...}) =
 | 
56  | 
sort_distinct (Term.term_ord o pairself Thm.full_prop_of) (Net.content props);  | 
|
| 18026 | 57  | 
fun could_unify (Index {props, ...}) = Net.unify_term props;
 | 
| 16020 | 58  | 
|
| 13270 | 59  | 
val empty =  | 
| 18026 | 60  | 
  Index {facts = [], consts = Symtab.empty, frees = Symtab.empty, props = Net.empty};
 | 
| 13270 | 61  | 
|
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
62  | 
fun add do_props do_index (fact as (_, ths)) (Index {facts, consts, frees, props}) =
 | 
| 13270 | 63  | 
let  | 
| 18026 | 64  | 
val entry = (serial (), fact);  | 
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
65  | 
val (cs, xs) = fold index_thm ths ([], []);  | 
| 18026 | 66  | 
|
| 19141 | 67  | 
val (facts', consts', frees') =  | 
68  | 
if do_index then  | 
|
69  | 
(fact :: facts,  | 
|
70  | 
consts |> fold (fn c => Symtab.update_list (c, entry)) cs,  | 
|
71  | 
frees |> fold (fn x => Symtab.update_list (x, entry)) xs)  | 
|
72  | 
else (facts, consts, frees);  | 
|
73  | 
val props' =  | 
|
74  | 
if do_props then props |> fold (fn th => Net.insert_term (K false) (Thm.prop_of th, th)) ths  | 
|
75  | 
else props;  | 
|
| 18026 | 76  | 
  in Index {facts = facts', consts = consts', frees = frees', props = props'} end;
 | 
77  | 
||
| 
20156
 
7a7898b1cfa4
add_local: simplified interface, all frees are known'';
 
wenzelm 
parents: 
20010 
diff
changeset
 | 
78  | 
val add_global = add false true;  | 
| 18026 | 79  | 
val add_local = add true;  | 
| 13270 | 80  | 
|
81  | 
||
| 18026 | 82  | 
(* find by consts/frees *)  | 
83  | 
||
84  | 
local  | 
|
| 13270 | 85  | 
|
| 16491 | 86  | 
fun fact_ord ((i, _), (j, _)) = int_ord (j, i);  | 
| 13270 | 87  | 
|
88  | 
fun intersects [xs] = xs  | 
|
| 16491 | 89  | 
| intersects xss =  | 
90  | 
if exists null xss then []  | 
|
91  | 
else fold (OrdList.inter fact_ord) (tl xss) (hd xss);  | 
|
| 13270 | 92  | 
|
| 18026 | 93  | 
in  | 
94  | 
||
| 16020 | 95  | 
fun find idx ([], []) = facts idx  | 
96  | 
  | find (Index {consts, frees, ...}) (cs, xs) =
 | 
|
| 18931 | 97  | 
(map (Symtab.lookup_list consts) cs @  | 
98  | 
map (Symtab.lookup_list frees) xs) |> intersects |> map #2;  | 
|
| 13270 | 99  | 
|
| 18026 | 100  | 
end  | 
101  | 
||
| 13270 | 102  | 
end;  |