| author | wenzelm | 
| Sat, 04 Jul 2009 22:22:34 +0200 | |
| changeset 31941 | d3a94ae9936f | 
| parent 31910 | a8e9ccfc427a | 
| child 31971 | 8c1b845ed105 | 
| permissions | -rw-r--r-- | 
| 29267 | 1 | (* Author: Jia Meng, Cambridge University Computer Laboratory, NICTA *) | 
| 15452 | 2 | |
| 16802 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 3 | signature RES_ATP = | 
| 
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
 wenzelm parents: 
16767diff
changeset | 4 | sig | 
| 19194 | 5 | datatype mode = Auto | Fol | Hol | 
| 22989 | 6 | val tvar_classes_of_terms : term list -> string list | 
| 7 | val tfree_classes_of_terms : term list -> string list | |
| 8 | val type_consts_of_terms : theory -> term list -> string list | |
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
30536diff
changeset | 9 | val get_relevant : int -> bool -> Proof.context * (thm list * 'a) -> thm list -> | 
| 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
30536diff
changeset | 10 | (thm * (string * int)) list | 
| 31410 | 11 | val prepare_clauses : bool -> thm list -> thm list -> | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 12 | (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list -> | 
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
30536diff
changeset | 13 | (thm * (ResHolClause.axiom_name * ResHolClause.clause_id)) list -> theory -> | 
| 31865 
5e97c4abd18e
fixed: count constants with supplementary lemmas
 immler@in.tum.de parents: 
31837diff
changeset | 14 | ResHolClause.axiom_name vector * | 
| 
5e97c4abd18e
fixed: count constants with supplementary lemmas
 immler@in.tum.de parents: 
31837diff
changeset | 15 | (ResHolClause.clause list * ResHolClause.clause list * ResHolClause.clause list * | 
| 
5e97c4abd18e
fixed: count constants with supplementary lemmas
 immler@in.tum.de parents: 
31837diff
changeset | 16 | ResHolClause.clause list * ResClause.classrelClause list * ResClause.arityClause list) | 
| 15347 | 17 | end; | 
| 18 | ||
| 22865 | 19 | structure ResAtp: RES_ATP = | 
| 15347 | 20 | struct | 
| 21 | ||
| 21070 | 22 | |
| 19194 | 23 | (********************************************************************) | 
| 24 | (* some settings for both background automatic ATP calling procedure*) | |
| 25 | (* and also explicit ATP invocation methods *) | |
| 26 | (********************************************************************) | |
| 27 | ||
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 28 | (*Translation mode can be auto-detected, or forced to be first-order or higher-order*) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 29 | datatype mode = Auto | Fol | Hol; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 30 | |
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 31 | val linkup_logic_mode = Auto; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 32 | |
| 19194 | 33 | (*** background linkup ***) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 34 | val run_blacklist_filter = true; | 
| 21224 | 35 | |
| 24287 | 36 | (*** relevance filter parameters ***) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 37 | val run_relevance_filter = true; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 38 | val pass_mark = 0.5; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 39 | val convergence = 3.2; (*Higher numbers allow longer inference chains*) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 40 | val follow_defs = false; (*Follow definitions. Makes problems bigger.*) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 41 | val include_all = true; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 42 | val include_atpset = true; | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 43 | |
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 44 | (***************************************************************) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 45 | (* Relevance Filtering *) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 46 | (***************************************************************) | 
| 19194 | 47 | |
| 24958 | 48 | fun strip_Trueprop (Const("Trueprop",_) $ t) = t
 | 
| 49 | | strip_Trueprop t = t; | |
| 19194 | 50 | |
| 24287 | 51 | (*A surprising number of theorems contain only a few significant constants. | 
| 52 | These include all induction rules, and other general theorems. Filtering | |
| 53 | theorems in clause form reveals these complexities in the form of Skolem | |
| 54 | functions. If we were instead to filter theorems in their natural form, | |
| 55 | some other method of measuring theorem complexity would become necessary.*) | |
| 56 | ||
| 57 | fun log_weight2 (x:real) = 1.0 + 2.0/Math.ln (x+1.0); | |
| 58 | ||
| 59 | (*The default seems best in practice. A constant function of one ignores | |
| 60 | the constant frequencies.*) | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 61 | val weight_fn = log_weight2; | 
| 24287 | 62 | |
| 63 | ||
| 64 | (*Including equality in this list might be expected to stop rules like subset_antisym from | |
| 65 | being chosen, but for some reason filtering works better with them listed. The | |
| 66 | logical signs All, Ex, &, and --> are omitted because any remaining occurrrences | |
| 67 | must be within comprehensions.*) | |
| 68 | val standard_consts = ["Trueprop","==>","all","==","op |","Not","op ="]; | |
| 69 | ||
| 70 | ||
| 71 | (*** constants with types ***) | |
| 72 | ||
| 73 | (*An abstraction of Isabelle types*) | |
| 74 | datatype const_typ = CTVar | CType of string * const_typ list | |
| 75 | ||
| 76 | (*Is the second type an instance of the first one?*) | |
| 77 | fun match_type (CType(con1,args1)) (CType(con2,args2)) = | |
| 78 | con1=con2 andalso match_types args1 args2 | |
| 79 | | match_type CTVar _ = true | |
| 80 | | match_type _ CTVar = false | |
| 81 | and match_types [] [] = true | |
| 82 | | match_types (a1::as1) (a2::as2) = match_type a1 a2 andalso match_types as1 as2; | |
| 83 | ||
| 84 | (*Is there a unifiable constant?*) | |
| 85 | fun uni_mem gctab (c,c_typ) = | |
| 86 | case Symtab.lookup gctab c of | |
| 87 | NONE => false | |
| 88 | | SOME ctyps_list => exists (match_types c_typ) ctyps_list; | |
| 89 | ||
| 90 | (*Maps a "real" type to a const_typ*) | |
| 91 | fun const_typ_of (Type (c,typs)) = CType (c, map const_typ_of typs) | |
| 92 | | const_typ_of (TFree _) = CTVar | |
| 93 | | const_typ_of (TVar _) = CTVar | |
| 94 | ||
| 95 | (*Pairs a constant with the list of its type instantiations (using const_typ)*) | |
| 96 | fun const_with_typ thy (c,typ) = | |
| 97 | let val tvars = Sign.const_typargs thy (c,typ) | |
| 98 | in (c, map const_typ_of tvars) end | |
| 99 | handle TYPE _ => (c,[]); (*Variable (locale constant): monomorphic*) | |
| 100 | ||
| 101 | (*Add a const/type pair to the table, but a [] entry means a standard connective, | |
| 102 | which we ignore.*) | |
| 103 | fun add_const_typ_table ((c,ctyps), tab) = | |
| 104 | Symtab.map_default (c, [ctyps]) (fn [] => [] | ctyps_list => insert (op =) ctyps ctyps_list) | |
| 105 | tab; | |
| 106 | ||
| 107 | (*Free variables are included, as well as constants, to handle locales*) | |
| 108 | fun add_term_consts_typs_rm thy (Const(c, typ), tab) = | |
| 109 | add_const_typ_table (const_with_typ thy (c,typ), tab) | |
| 110 | | add_term_consts_typs_rm thy (Free(c, typ), tab) = | |
| 111 | add_const_typ_table (const_with_typ thy (c,typ), tab) | |
| 112 | | add_term_consts_typs_rm thy (t $ u, tab) = | |
| 113 | add_term_consts_typs_rm thy (t, add_term_consts_typs_rm thy (u, tab)) | |
| 114 | | add_term_consts_typs_rm thy (Abs(_,_,t), tab) = add_term_consts_typs_rm thy (t, tab) | |
| 115 | | add_term_consts_typs_rm thy (_, tab) = tab; | |
| 116 | ||
| 117 | (*The empty list here indicates that the constant is being ignored*) | |
| 118 | fun add_standard_const (s,tab) = Symtab.update (s,[]) tab; | |
| 119 | ||
| 120 | val null_const_tab : const_typ list list Symtab.table = | |
| 30190 | 121 | List.foldl add_standard_const Symtab.empty standard_consts; | 
| 24287 | 122 | |
| 30190 | 123 | fun get_goal_consts_typs thy = List.foldl (add_term_consts_typs_rm thy) null_const_tab; | 
| 24287 | 124 | |
| 125 | (*Inserts a dummy "constant" referring to the theory name, so that relevance | |
| 126 | takes the given theory into account.*) | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 127 | fun const_prop_of theory_const th = | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 128 | if theory_const then | 
| 24287 | 129 | let val name = Context.theory_name (theory_of_thm th) | 
| 130 | val t = Const (name ^ ". 1", HOLogic.boolT) | |
| 131 | in t $ prop_of th end | |
| 132 | else prop_of th; | |
| 133 | ||
| 134 | (**** Constant / Type Frequencies ****) | |
| 135 | ||
| 136 | (*A two-dimensional symbol table counts frequencies of constants. It's keyed first by | |
| 137 | constant name and second by its list of type instantiations. For the latter, we need | |
| 138 | a linear ordering on type const_typ list.*) | |
| 139 | ||
| 140 | local | |
| 141 | ||
| 142 | fun cons_nr CTVar = 0 | |
| 143 | | cons_nr (CType _) = 1; | |
| 144 | ||
| 145 | in | |
| 146 | ||
| 147 | fun const_typ_ord TU = | |
| 148 | case TU of | |
| 149 | (CType (a, Ts), CType (b, Us)) => | |
| 150 | (case fast_string_ord(a,b) of EQUAL => dict_ord const_typ_ord (Ts,Us) | ord => ord) | |
| 151 | | (T, U) => int_ord (cons_nr T, cons_nr U); | |
| 152 | ||
| 153 | end; | |
| 154 | ||
| 155 | structure CTtab = TableFun(type key = const_typ list val ord = dict_ord const_typ_ord); | |
| 156 | ||
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 157 | fun count_axiom_consts theory_const thy ((thm,_), tab) = | 
| 24287 | 158 | let fun count_const (a, T, tab) = | 
| 159 | let val (c, cts) = const_with_typ thy (a,T) | |
| 160 | in (*Two-dimensional table update. Constant maps to types maps to count.*) | |
| 161 | Symtab.map_default (c, CTtab.empty) | |
| 162 | (CTtab.map_default (cts,0) (fn n => n+1)) tab | |
| 163 | end | |
| 164 | fun count_term_consts (Const(a,T), tab) = count_const(a,T,tab) | |
| 165 | | count_term_consts (Free(a,T), tab) = count_const(a,T,tab) | |
| 166 | | count_term_consts (t $ u, tab) = | |
| 167 | count_term_consts (t, count_term_consts (u, tab)) | |
| 168 | | count_term_consts (Abs(_,_,t), tab) = count_term_consts (t, tab) | |
| 169 | | count_term_consts (_, tab) = tab | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 170 | in count_term_consts (const_prop_of theory_const thm, tab) end; | 
| 24287 | 171 | |
| 172 | ||
| 173 | (**** Actual Filtering Code ****) | |
| 174 | ||
| 175 | (*The frequency of a constant is the sum of those of all instances of its type.*) | |
| 176 | fun const_frequency ctab (c, cts) = | |
| 177 | let val pairs = CTtab.dest (the (Symtab.lookup ctab c)) | |
| 178 | fun add ((cts',m), n) = if match_types cts cts' then m+n else n | |
| 179 | in List.foldl add 0 pairs end; | |
| 180 | ||
| 181 | (*Add in a constant's weight, as determined by its frequency.*) | |
| 182 | fun add_ct_weight ctab ((c,T), w) = | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 183 | w + weight_fn (real (const_frequency ctab (c,T))); | 
| 24287 | 184 | |
| 185 | (*Relevant constants are weighted according to frequency, | |
| 186 | but irrelevant constants are simply counted. Otherwise, Skolem functions, | |
| 187 | which are rare, would harm a clause's chances of being picked.*) | |
| 188 | fun clause_weight ctab gctyps consts_typs = | |
| 189 | let val rel = filter (uni_mem gctyps) consts_typs | |
| 190 | val rel_weight = List.foldl (add_ct_weight ctab) 0.0 rel | |
| 191 | in | |
| 192 | rel_weight / (rel_weight + real (length consts_typs - length rel)) | |
| 193 | end; | |
| 194 | ||
| 195 | (*Multiplies out to a list of pairs: 'a * 'b list -> ('a * 'b) list -> ('a * 'b) list*)
 | |
| 30190 | 196 | fun add_expand_pairs (x,ys) xys = List.foldl (fn (y,acc) => (x,y)::acc) xys ys; | 
| 24287 | 197 | |
| 198 | fun consts_typs_of_term thy t = | |
| 199 | let val tab = add_term_consts_typs_rm thy (t, null_const_tab) | |
| 200 | in Symtab.fold add_expand_pairs tab [] end; | |
| 201 | ||
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 202 | fun pair_consts_typs_axiom theory_const thy (thm,name) = | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 203 | ((thm,name), (consts_typs_of_term thy (const_prop_of theory_const thm))); | 
| 24287 | 204 | |
| 205 | exception ConstFree; | |
| 206 | fun dest_ConstFree (Const aT) = aT | |
| 207 | | dest_ConstFree (Free aT) = aT | |
| 208 | | dest_ConstFree _ = raise ConstFree; | |
| 209 | ||
| 210 | (*Look for definitions of the form f ?x1 ... ?xn = t, but not reversed.*) | |
| 211 | fun defines thy (thm,(name,n)) gctypes = | |
| 212 | let val tm = prop_of thm | |
| 213 | fun defs lhs rhs = | |
| 214 | let val (rator,args) = strip_comb lhs | |
| 215 | val ct = const_with_typ thy (dest_ConstFree rator) | |
| 216 | in forall is_Var args andalso uni_mem gctypes ct andalso | |
| 217 | Term.add_vars rhs [] subset Term.add_vars lhs [] | |
| 218 | end | |
| 219 | handle ConstFree => false | |
| 220 | in | |
| 221 | 	case tm of Const ("Trueprop",_) $ (Const("op =",_) $ lhs $ rhs) => 
 | |
| 31910 | 222 | defs lhs rhs | 
| 24287 | 223 | | _ => false | 
| 224 | end; | |
| 225 | ||
| 226 | type annotd_cls = (thm * (string * int)) * ((string * const_typ list) list); | |
| 227 | ||
| 228 | (*For a reverse sort, putting the largest values first.*) | |
| 229 | fun compare_pairs ((_,w1),(_,w2)) = Real.compare (w2,w1); | |
| 230 | ||
| 231 | (*Limit the number of new clauses, to prevent runaway acceptance.*) | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 232 | fun take_best max_new (newpairs : (annotd_cls*real) list) = | 
| 24287 | 233 | let val nnew = length newpairs | 
| 234 | in | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 235 | if nnew <= max_new then (map #1 newpairs, []) | 
| 24287 | 236 | else | 
| 237 | let val cls = sort compare_pairs newpairs | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 238 | val accepted = List.take (cls, max_new) | 
| 24287 | 239 | in | 
| 240 |         Output.debug (fn () => ("Number of candidates, " ^ Int.toString nnew ^ 
 | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 241 | ", exceeds the limit of " ^ Int.toString (max_new))); | 
| 24287 | 242 |         Output.debug (fn () => ("Effective pass mark: " ^ Real.toString (#2 (List.last accepted))));
 | 
| 243 | Output.debug (fn () => "Actually passed: " ^ | |
| 244 | space_implode ", " (map (fn (((_,(name,_)),_),_) => name) accepted)); | |
| 245 | ||
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 246 | (map #1 accepted, map #1 (List.drop (cls, max_new))) | 
| 24287 | 247 | end | 
| 248 | end; | |
| 249 | ||
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 250 | fun relevant_clauses max_new thy ctab p rel_consts = | 
| 24287 | 251 | let fun relevant ([],_) [] = [] : (thm * (string * int)) list (*Nothing added this iteration*) | 
| 252 | | relevant (newpairs,rejects) [] = | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 253 | let val (newrels,more_rejects) = take_best max_new newpairs | 
| 24287 | 254 | val new_consts = List.concat (map #2 newrels) | 
| 30190 | 255 | val rel_consts' = List.foldl add_const_typ_table rel_consts new_consts | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 256 | val newp = p + (1.0-p) / convergence | 
| 24287 | 257 | in | 
| 258 |               Output.debug (fn () => ("relevant this iteration: " ^ Int.toString (length newrels)));
 | |
| 259 | (map #1 newrels) @ | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 260 | (relevant_clauses max_new thy ctab newp rel_consts' (more_rejects@rejects)) | 
| 24287 | 261 | end | 
| 262 | | relevant (newrels,rejects) ((ax as (clsthm as (_,(name,n)),consts_typs)) :: axs) = | |
| 263 | let val weight = clause_weight ctab rel_consts consts_typs | |
| 264 | in | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 265 | if p <= weight orelse (follow_defs andalso defines thy clsthm rel_consts) | 
| 24287 | 266 | then (Output.debug (fn () => (name ^ " clause " ^ Int.toString n ^ | 
| 267 | " passes: " ^ Real.toString weight)); | |
| 268 | relevant ((ax,weight)::newrels, rejects) axs) | |
| 269 | else relevant (newrels, ax::rejects) axs | |
| 270 | end | |
| 271 |     in  Output.debug (fn () => ("relevant_clauses, current pass mark = " ^ Real.toString p));
 | |
| 272 | relevant ([],[]) | |
| 273 | end; | |
| 274 | ||
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 275 | fun relevance_filter max_new theory_const thy axioms goals = | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 276 | if run_relevance_filter andalso pass_mark >= 0.1 | 
| 24287 | 277 | then | 
| 31910 | 278 | let val const_tab = List.foldl (count_axiom_consts theory_const thy) Symtab.empty axioms | 
| 24287 | 279 | val goal_const_tab = get_goal_consts_typs thy goals | 
| 280 |       val _ = Output.debug (fn () => ("Initial constants: " ^
 | |
| 281 | space_implode ", " (Symtab.keys goal_const_tab))); | |
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 282 | val rels = relevant_clauses max_new thy const_tab (pass_mark) | 
| 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 283 | goal_const_tab (map (pair_consts_typs_axiom theory_const thy) axioms) | 
| 24287 | 284 | in | 
| 285 |       Output.debug (fn () => ("Total relevant: " ^ Int.toString (length rels)));
 | |
| 286 | rels | |
| 287 | end | |
| 288 | else axioms; | |
| 289 | ||
| 290 | (***************************************************************) | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 291 | (* Retrieving and filtering lemmas *) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 292 | (***************************************************************) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 293 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 294 | (*** white list and black list of lemmas ***) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 295 | |
| 24215 | 296 | (*The rule subsetI is frequently omitted by the relevance filter. This could be theory data | 
| 297 | or identified with ATPset (which however is too big currently)*) | |
| 31837 | 298 | val whitelist_fo = [subsetI]; | 
| 299 | (* ext has been a 'helper clause', always given to the atps. | |
| 300 | As such it was not passed to metis, but metis does not include ext (in contrast | |
| 301 | to the other helper_clauses *) | |
| 302 | val whitelist_ho = [ResHolClause.ext]; | |
| 21588 | 303 | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 304 | (*** retrieve lemmas from clasimpset and atpset, may filter them ***) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 305 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 306 | (*Hashing to detect duplicate and variant clauses, e.g. from the [iff] attribute*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 307 | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 308 | fun setinsert (x,s) = Symtab.update (x,()) s; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 309 | |
| 20757 
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
 paulson parents: 
20661diff
changeset | 310 | (*Reject theorems with names like "List.filter.filter_list_def" or | 
| 21690 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 311 | "Accessible_Part.acc.defs", as these are definitions arising from packages.*) | 
| 20757 
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
 paulson parents: 
20661diff
changeset | 312 | fun is_package_def a = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30291diff
changeset | 313 | let val names = Long_Name.explode a | 
| 21690 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 314 | in | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 315 | length names > 2 andalso | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 316 | not (hd names = "local") andalso | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 317 | String.isSuffix "_def" a orelse String.isSuffix "_defs" a | 
| 
552d20ff9a95
Removal of theorem tagging, which the ATP linkup no longer requires.
 paulson parents: 
21588diff
changeset | 318 | end; | 
| 20757 
fe84fe0dfd30
Definitions produced by packages are now blacklisted.
 paulson parents: 
20661diff
changeset | 319 | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 320 | (** a hash function from Term.term to int, and also a hash table **) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 321 | val xor_words = List.foldl Word.xorb 0w0; | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 322 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 323 | fun hashw_term ((Const(c,_)), w) = Polyhash.hashw_string (c,w) | 
| 20661 
46832fee1215
Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
 paulson parents: 
20643diff
changeset | 324 | | hashw_term ((Free(a,_)), w) = Polyhash.hashw_string (a,w) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 325 | | hashw_term ((Var(_,_)), w) = w | 
| 20661 
46832fee1215
Yet another version of fake_thm_name. "Full" hashing ensures that there are no collisions
 paulson parents: 
20643diff
changeset | 326 | | hashw_term ((Bound i), w) = Polyhash.hashw_int(i,w) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 327 | | hashw_term ((Abs(_,_,t)), w) = hashw_term (t, w) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 328 | | hashw_term ((P$Q), w) = hashw_term (Q, (hashw_term (P, w))); | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 329 | |
| 21070 | 330 | fun hash_literal (Const("Not",_)$P) = Word.notb(hashw_term(P,0w0))
 | 
| 331 | | hash_literal P = hashw_term(P,0w0); | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 332 | |
| 24958 | 333 | fun hash_term t = Word.toIntX (xor_words (map hash_literal (HOLogic.disjuncts (strip_Trueprop t)))); | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 334 | |
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 335 | fun equal_thm (thm1,thm2) = Term.aconv(prop_of thm1, prop_of thm2); | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 336 | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 337 | exception HASH_CLAUSE; | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 338 | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 339 | (*Create a hash table for clauses, of the given size*) | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 340 | fun mk_clause_table n = | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 341 | Polyhash.mkTable (hash_term o prop_of, equal_thm) | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 342 | (n, HASH_CLAUSE); | 
| 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 343 | |
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 344 | (*Use a hash table to eliminate duplicates from xs. Argument is a list of | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 345 | (thm * (string * int)) tuples. The theorems are hashed into the table. *) | 
| 21588 | 346 | fun make_unique xs = | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 347 | let val ht = mk_clause_table 7000 | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 348 | in | 
| 22130 | 349 |       Output.debug (fn () => ("make_unique gets " ^ Int.toString (length xs) ^ " clauses"));
 | 
| 21588 | 350 | app (ignore o Polyhash.peekInsert ht) xs; | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 351 | Polyhash.listItems ht | 
| 20457 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 paulson parents: 
20446diff
changeset | 352 | end; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 353 | |
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 354 | (*Remove existing axiom clauses from the conjecture clauses, as this can dramatically | 
| 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 355 | boost an ATP's performance (for some reason)*) | 
| 21588 | 356 | fun subtract_cls c_clauses ax_clauses = | 
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 357 | let val ht = mk_clause_table 2200 | 
| 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 358 | fun known x = isSome (Polyhash.peek ht x) | 
| 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 359 | in | 
| 21588 | 360 | app (ignore o Polyhash.peekInsert ht) ax_clauses; | 
| 361 | filter (not o known) c_clauses | |
| 20868 
724ab9896068
Improved detection of identical clauses, also in the conjecture
 paulson parents: 
20854diff
changeset | 362 | end; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 363 | |
| 26675 | 364 | fun valid_facts facts = | 
| 26691 
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
 wenzelm parents: 
26675diff
changeset | 365 | Facts.fold_static (fn (name, ths) => | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 366 | if run_blacklist_filter andalso is_package_def name then I | 
| 26691 
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
 wenzelm parents: 
26675diff
changeset | 367 | else | 
| 
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
 wenzelm parents: 
26675diff
changeset | 368 | let val xname = Facts.extern facts name in | 
| 
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
 wenzelm parents: 
26675diff
changeset | 369 | if NameSpace.is_hidden xname then I | 
| 
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
 wenzelm parents: 
26675diff
changeset | 370 | else cons (xname, filter_out ResAxioms.bad_for_atp ths) | 
| 
520c99e0b9a0
valid_facts: more elementary version using Facts.fold_static;
 wenzelm parents: 
26675diff
changeset | 371 | end) facts []; | 
| 26675 | 372 | |
| 21224 | 373 | fun all_valid_thms ctxt = | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 374 | let | 
| 26675 | 375 | val global_facts = PureThy.facts_of (ProofContext.theory_of ctxt); | 
| 26278 | 376 | val local_facts = ProofContext.facts_of ctxt; | 
| 26675 | 377 | in valid_facts global_facts @ valid_facts local_facts end; | 
| 21224 | 378 | |
| 21588 | 379 | fun multi_name a (th, (n,pairs)) = | 
| 21224 | 380 |   (n+1, (a ^ "(" ^ Int.toString n ^ ")", th) :: pairs)
 | 
| 381 | ||
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 382 | fun add_single_names ((a, []), pairs) = pairs | 
| 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 383 | | add_single_names ((a, [th]), pairs) = (a,th)::pairs | 
| 30190 | 384 | | add_single_names ((a, ths), pairs) = #2 (List.foldl (multi_name a) (1,pairs) ths); | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 385 | |
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 386 | (*Ignore blacklisted basenames*) | 
| 21588 | 387 | fun add_multi_names ((a, ths), pairs) = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30291diff
changeset | 388 | if (Long_Name.base_name a) mem_string ResAxioms.multi_base_blacklist then pairs | 
| 22382 
dbf09db0a40d
New code to store theorem names in a concise form rather than as fully-qualified.
 paulson parents: 
22217diff
changeset | 389 | else add_single_names ((a, ths), pairs); | 
| 21224 | 390 | |
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 391 | fun is_multi (a, ths) = length ths > 1 orelse String.isSuffix ".axioms" a; | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 392 | |
| 24286 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24215diff
changeset | 393 | (*The single theorems go BEFORE the multiple ones. Blacklist is applied to all.*) | 
| 21588 | 394 | fun name_thm_pairs ctxt = | 
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 395 | let val (mults,singles) = List.partition is_multi (all_valid_thms ctxt) | 
| 24286 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24215diff
changeset | 396 | val ht = mk_clause_table 900 (*ht for blacklisted theorems*) | 
| 28477 
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
 wenzelm parents: 
28065diff
changeset | 397 | fun blacklisted x = run_blacklist_filter andalso isSome (Polyhash.peek ht x) | 
| 24286 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24215diff
changeset | 398 | in | 
| 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24215diff
changeset | 399 | app (fn th => ignore (Polyhash.peekInsert ht (th,()))) (ResBlacklist.get ctxt); | 
| 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24215diff
changeset | 400 | filter (not o blacklisted o #2) | 
| 30190 | 401 | (List.foldl add_single_names (List.foldl add_multi_names [] mults) singles) | 
| 24286 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
24215diff
changeset | 402 | end; | 
| 21224 | 403 | |
| 26928 | 404 | fun check_named ("",th) = (warning ("No name for theorem " ^ Display.string_of_thm th); false)
 | 
| 21224 | 405 | | check_named (_,th) = true; | 
| 19894 | 406 | |
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 407 | (* get lemmas from claset, simpset, atpset and extra supplied rules *) | 
| 31410 | 408 | fun get_clasimp_atp_lemmas ctxt = | 
| 19894 | 409 | let val included_thms = | 
| 31410 | 410 | if include_all | 
| 411 | then (tap (fn ths => Output.debug | |
| 412 |                  (fn () => ("Including all " ^ Int.toString (length ths) ^ " theorems")))
 | |
| 413 | (name_thm_pairs ctxt)) | |
| 414 | else | |
| 415 | let val atpset_thms = | |
| 416 | if include_atpset then ResAxioms.atpset_rules_of ctxt | |
| 417 | else [] | |
| 418 | in atpset_thms end | |
| 19894 | 419 | in | 
| 31410 | 420 | filter check_named included_thms | 
| 19894 | 421 | end; | 
| 19768 
9afd9b9c47d0
ATP/res_clasimpset.ML has been merged into res_atp.ML.
 mengj parents: 
19746diff
changeset | 422 | |
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 423 | (***************************************************************) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 424 | (* Type Classes Present in the Axiom or Conjecture Clauses *) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 425 | (***************************************************************) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 426 | |
| 30190 | 427 | fun add_classes (sorts, cset) = List.foldl setinsert cset (List.concat sorts); | 
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 428 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 429 | (*Remove this trivial type class*) | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 430 | fun delete_type cset = Symtab.delete_safe "HOL.type" cset; | 
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 431 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 432 | fun tvar_classes_of_terms ts = | 
| 29270 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 wenzelm parents: 
29267diff
changeset | 433 | let val sorts_list = map (map #2 o OldTerm.term_tvars) ts | 
| 30190 | 434 | in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end; | 
| 21290 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 435 | |
| 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 paulson parents: 
21243diff
changeset | 436 | fun tfree_classes_of_terms ts = | 
| 29270 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 wenzelm parents: 
29267diff
changeset | 437 | let val sorts_list = map (map #2 o OldTerm.term_tfrees) ts | 
| 30190 | 438 | in Symtab.keys (delete_type (List.foldl add_classes Symtab.empty sorts_list)) end; | 
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 439 | |
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 440 | (*fold type constructors*) | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 441 | fun fold_type_consts f (Type (a, Ts)) x = fold (fold_type_consts f) Ts (f (a,x)) | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 442 | | fold_type_consts f T x = x; | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 443 | |
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 444 | val add_type_consts_in_type = fold_type_consts setinsert; | 
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 445 | |
| 21397 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 446 | (*Type constructors used to instantiate overloaded constants are the only ones needed.*) | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 447 | fun add_type_consts_in_term thy = | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 448 | let val const_typargs = Sign.const_typargs thy | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 449 | fun add_tcs (Const cT) x = fold add_type_consts_in_type (const_typargs cT) x | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 450 | | add_tcs (Abs (_, T, u)) x = add_tcs u x | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 451 | | add_tcs (t $ u) x = add_tcs t (add_tcs u x) | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 452 | | add_tcs _ x = x | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 453 | in add_tcs end | 
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 454 | |
| 21397 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 455 | fun type_consts_of_terms thy ts = | 
| 
2134b81a0b37
Now includes only types used to instantiate overloaded constants in arity clauses
 paulson parents: 
21373diff
changeset | 456 | Symtab.keys (fold (add_type_consts_in_term thy) ts Symtab.empty); | 
| 21373 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 457 | |
| 
18f519614978
Arity clauses are now produced only for types and type classes actually used.
 paulson parents: 
21311diff
changeset | 458 | |
| 19194 | 459 | (***************************************************************) | 
| 460 | (* ATP invocation methods setup *) | |
| 461 | (***************************************************************) | |
| 462 | ||
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 463 | (*Ensures that no higher-order theorems "leak out"*) | 
| 24958 | 464 | fun restrict_to_logic thy true cls = filter (Meson.is_fol_term thy o prop_of o fst) cls | 
| 465 | | restrict_to_logic thy false cls = cls; | |
| 20526 
756c4f1fd04c
Extended the blacklist with higher-order theorems. Restructured. Added checks to
 paulson parents: 
20457diff
changeset | 466 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 467 | (**** Predicates to detect unwanted clauses (prolific or likely to cause unsoundness) ****) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 468 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 469 | (** Too general means, positive equality literal with a variable X as one operand, | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 470 | when X does not occur properly in the other operand. This rules out clearly | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 471 | inconsistent clauses such as V=a|V=b, though it by no means guarantees soundness. **) | 
| 21588 | 472 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 473 | fun occurs ix = | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 474 | let fun occ(Var (jx,_)) = (ix=jx) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 475 | | occ(t1$t2) = occ t1 orelse occ t2 | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 476 | | occ(Abs(_,_,t)) = occ t | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 477 | | occ _ = false | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 478 | in occ end; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 479 | |
| 31723 
f5cafe803b55
discontinued ancient tradition to suffix certain ML module names with "_package"
 haftmann parents: 
31410diff
changeset | 480 | fun is_recordtype T = not (null (Record.dest_recTs T)); | 
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 481 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 482 | (*Unwanted equalities include | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 483 | (1) those between a variable that does not properly occur in the second operand, | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 484 | (2) those between a variable and a record, since these seem to be prolific "cases" thms | 
| 21588 | 485 | *) | 
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 486 | fun too_general_eqterms (Var (ix,T), t) = not (occurs ix t) orelse is_recordtype T | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 487 | | too_general_eqterms _ = false; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 488 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 489 | fun too_general_equality (Const ("op =", _) $ x $ y) =
 | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 490 | too_general_eqterms (x,y) orelse too_general_eqterms(y,x) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 491 | | too_general_equality _ = false; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 492 | |
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 493 | (* tautologous? *) | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 494 | fun is_taut (Const ("Trueprop", _) $ Const ("True", _)) = true
 | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 495 | | is_taut _ = false; | 
| 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 496 | |
| 29267 | 497 | fun has_typed_var tycons = exists_subterm | 
| 498 | (fn Var (_, Type (a, _)) => member (op =) tycons a | _ => false); | |
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 499 | |
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 500 | (*Clauses are forbidden to contain variables of these types. The typical reason is that | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 501 | they lead to unsoundness. Note that "unit" satisfies numerous equations like ?X=(). | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 502 | The resulting clause will have no type constraint, yielding false proofs. Even "bool" | 
| 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 503 | leads to many unsound proofs, though (obviously) only for higher-order problems.*) | 
| 24215 | 504 | val unwanted_types = ["Product_Type.unit","bool"]; | 
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 505 | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 506 | fun unwanted t = | 
| 24958 | 507 | is_taut t orelse has_typed_var unwanted_types t orelse | 
| 508 | forall too_general_equality (HOLogic.disjuncts (strip_Trueprop t)); | |
| 21470 
7c1b59ddcd56
Consolidation of code to "blacklist" unhelpful theorems, including record
 paulson parents: 
21431diff
changeset | 509 | |
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 510 | (*Clauses containing variables of type "unit" or "bool" are unlikely to be useful and | 
| 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 511 | likely to lead to unsound proofs.*) | 
| 22217 
a5d983f7113f
Tidying; more debugging information. New reference unwanted_types.
 paulson parents: 
22193diff
changeset | 512 | fun remove_unwanted_clauses cls = filter (not o unwanted o prop_of o fst) cls; | 
| 21431 
ef9080e7dbbc
Outputs a minimal number of arity clauses. Tidying of blacklist, fixing the blacklisting of thm lists
 paulson parents: 
21397diff
changeset | 513 | |
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 514 | fun isFO thy goal_cls = case linkup_logic_mode of | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 515 | Auto => forall (Meson.is_fol_term thy) (map prop_of goal_cls) | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 516 | | Fol => true | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 517 | | Hol => false | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 518 | |
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
30536diff
changeset | 519 | fun get_relevant max_new theory_const (ctxt, (chain_ths, th)) goal_cls = | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 520 | let | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 521 | val thy = ProofContext.theory_of ctxt | 
| 31410 | 522 | val included_thms = get_clasimp_atp_lemmas ctxt | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 523 | val included_cls = included_thms |> ResAxioms.cnf_rules_pairs thy |> make_unique | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 524 | |> restrict_to_logic thy (isFO thy goal_cls) | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 525 | |> remove_unwanted_clauses | 
| 19194 | 526 | in | 
| 31410 | 527 | relevance_filter max_new theory_const thy included_cls (map prop_of goal_cls) | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 528 | end; | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 529 | |
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 530 | (* prepare for passing to writer, | 
| 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 531 | create additional clauses based on the information from extra_cls *) | 
| 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 532 | fun prepare_clauses dfg goal_cls chain_ths axcls extra_cls thy = | 
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
30536diff
changeset | 533 | let | 
| 31837 | 534 | val isFO = isFO thy goal_cls | 
| 535 | val white_thms = filter check_named (map ResAxioms.pairname | |
| 536 | (whitelist_fo @ (if isFO then [] else whitelist_ho) @ chain_ths)) | |
| 31410 | 537 | val white_cls = ResAxioms.cnf_rules_pairs thy white_thms | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 538 | val extra_cls = white_cls @ extra_cls | 
| 31836 | 539 | val axcls = white_cls @ axcls | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 540 | val ccls = subtract_cls goal_cls extra_cls | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 541 | val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm th)) ccls | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 542 | val ccltms = map prop_of ccls | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 543 | and axtms = map (prop_of o #1) extra_cls | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 544 | val subs = tfree_classes_of_terms ccltms | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 545 | and supers = tvar_classes_of_terms axtms | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 546 | and tycons = type_consts_of_terms thy (ccltms@axtms) | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 547 | (*TFrees in conjecture clauses; TVars in axiom clauses*) | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 548 | val conjectures = ResHolClause.make_conjecture_clauses dfg thy ccls | 
| 31865 
5e97c4abd18e
fixed: count constants with supplementary lemmas
 immler@in.tum.de parents: 
31837diff
changeset | 549 | val (_, extra_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy extra_cls) | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 550 | val (clnames,axiom_clauses) = ListPair.unzip (ResHolClause.make_axiom_clauses dfg thy axcls) | 
| 31837 | 551 | val helper_clauses = ResHolClause.get_helper_clauses dfg thy isFO (conjectures, extra_cls, []) | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 552 | val (supers',arity_clauses) = ResClause.make_arity_clauses_dfg dfg thy tycons supers | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 553 | val classrel_clauses = ResClause.make_classrel_clauses thy subs supers' | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 554 | in | 
| 31752 
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
 immler@in.tum.de parents: 
31749diff
changeset | 555 | (Vector.fromList clnames, | 
| 31865 
5e97c4abd18e
fixed: count constants with supplementary lemmas
 immler@in.tum.de parents: 
31837diff
changeset | 556 | (conjectures, axiom_clauses, extra_clauses, helper_clauses, classrel_clauses, arity_clauses)) | 
| 31409 
d8537ba165b5
split preparing clauses and writing problemfile;
 immler@in.tum.de parents: 
30536diff
changeset | 557 | end | 
| 15644 | 558 | |
| 15347 | 559 | end; | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30364diff
changeset | 560 |