blanchet@35826
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_fact_filter.ML
|
blanchet@38027
|
2 |
Author: Jia Meng, Cambridge University Computer Laboratory and NICTA
|
blanchet@36393
|
3 |
Author: Jasmin Blanchette, TU Muenchen
|
wenzelm@33309
|
4 |
*)
|
paulson@15452
|
5 |
|
blanchet@35826
|
6 |
signature SLEDGEHAMMER_FACT_FILTER =
|
wenzelm@16802
|
7 |
sig
|
blanchet@38937
|
8 |
datatype locality = General | Intro | Elim | Simp | Local | Chained
|
blanchet@38752
|
9 |
|
blanchet@35966
|
10 |
type relevance_override =
|
blanchet@35966
|
11 |
{add: Facts.ref list,
|
blanchet@35966
|
12 |
del: Facts.ref list,
|
blanchet@35966
|
13 |
only: bool}
|
blanchet@35966
|
14 |
|
blanchet@37616
|
15 |
val trace : bool Unsynchronized.ref
|
blanchet@38938
|
16 |
val worse_irrel_freq : real Unsynchronized.ref
|
blanchet@38899
|
17 |
val abs_rel_weight : real Unsynchronized.ref
|
blanchet@38899
|
18 |
val abs_irrel_weight : real Unsynchronized.ref
|
blanchet@38899
|
19 |
val skolem_irrel_weight : real Unsynchronized.ref
|
blanchet@38937
|
20 |
val intro_bonus : real Unsynchronized.ref
|
blanchet@38937
|
21 |
val elim_bonus : real Unsynchronized.ref
|
blanchet@38937
|
22 |
val simp_bonus : real Unsynchronized.ref
|
blanchet@38899
|
23 |
val local_bonus : real Unsynchronized.ref
|
blanchet@38899
|
24 |
val chained_bonus : real Unsynchronized.ref
|
blanchet@38904
|
25 |
val max_imperfect : real Unsynchronized.ref
|
blanchet@38904
|
26 |
val max_imperfect_exp : real Unsynchronized.ref
|
blanchet@38899
|
27 |
val threshold_divisor : real Unsynchronized.ref
|
blanchet@38899
|
28 |
val ridiculous_threshold : real Unsynchronized.ref
|
blanchet@38744
|
29 |
val name_thm_pairs_from_ref :
|
blanchet@38696
|
30 |
Proof.context -> unit Symtab.table -> thm list -> Facts.ref
|
blanchet@38752
|
31 |
-> ((string * locality) * thm) list
|
blanchet@37347
|
32 |
val relevant_facts :
|
blanchet@38891
|
33 |
Proof.context -> bool -> real * real -> int -> bool -> relevance_override
|
blanchet@38891
|
34 |
-> thm list -> term list -> term -> ((string * locality) * thm) list
|
paulson@15347
|
35 |
end;
|
paulson@15347
|
36 |
|
blanchet@35826
|
37 |
structure Sledgehammer_Fact_Filter : SLEDGEHAMMER_FACT_FILTER =
|
paulson@15347
|
38 |
struct
|
paulson@15347
|
39 |
|
blanchet@38652
|
40 |
open Sledgehammer_Util
|
blanchet@38652
|
41 |
|
blanchet@37616
|
42 |
val trace = Unsynchronized.ref false
|
blanchet@37616
|
43 |
fun trace_msg msg = if !trace then tracing (msg ()) else ()
|
blanchet@35826
|
44 |
|
blanchet@38828
|
45 |
(* experimental feature *)
|
blanchet@38889
|
46 |
val term_patterns = false
|
blanchet@38827
|
47 |
|
blanchet@37580
|
48 |
val respect_no_atp = true
|
blanchet@37505
|
49 |
|
blanchet@38937
|
50 |
datatype locality = General | Intro | Elim | Simp | Local | Chained
|
blanchet@38752
|
51 |
|
blanchet@35966
|
52 |
type relevance_override =
|
blanchet@35966
|
53 |
{add: Facts.ref list,
|
blanchet@35966
|
54 |
del: Facts.ref list,
|
blanchet@35966
|
55 |
only: bool}
|
paulson@21070
|
56 |
|
blanchet@37616
|
57 |
val sledgehammer_prefix = "Sledgehammer" ^ Long_Name.separator
|
blanchet@37616
|
58 |
|
blanchet@38744
|
59 |
fun repair_name reserved multi j name =
|
blanchet@38744
|
60 |
(name |> Symtab.defined reserved name ? quote) ^
|
blanchet@38744
|
61 |
(if multi then "(" ^ Int.toString j ^ ")" else "")
|
blanchet@38744
|
62 |
|
blanchet@38744
|
63 |
fun name_thm_pairs_from_ref ctxt reserved chained_ths xref =
|
blanchet@38744
|
64 |
let
|
blanchet@38744
|
65 |
val ths = ProofContext.get_fact ctxt xref
|
blanchet@38744
|
66 |
val name = Facts.string_of_ref xref
|
blanchet@38744
|
67 |
val multi = length ths > 1
|
blanchet@38744
|
68 |
in
|
blanchet@38752
|
69 |
(ths, (1, []))
|
blanchet@38752
|
70 |
|-> fold (fn th => fn (j, rest) =>
|
blanchet@38752
|
71 |
(j + 1, ((repair_name reserved multi j name,
|
blanchet@38752
|
72 |
if member Thm.eq_thm chained_ths th then Chained
|
blanchet@38752
|
73 |
else General), th) :: rest))
|
blanchet@38744
|
74 |
|> snd
|
blanchet@38699
|
75 |
end
|
blanchet@37616
|
76 |
|
wenzelm@28477
|
77 |
(***************************************************************)
|
wenzelm@28477
|
78 |
(* Relevance Filtering *)
|
wenzelm@28477
|
79 |
(***************************************************************)
|
mengj@19194
|
80 |
|
paulson@24287
|
81 |
(*** constants with types ***)
|
paulson@24287
|
82 |
|
blanchet@38823
|
83 |
(* An abstraction of Isabelle types and first-order terms *)
|
blanchet@38823
|
84 |
datatype pattern = PVar | PApp of string * pattern list
|
blanchet@38744
|
85 |
|
blanchet@38823
|
86 |
fun string_for_pattern PVar = "_"
|
blanchet@38823
|
87 |
| string_for_pattern (PApp (s, ps)) =
|
blanchet@38823
|
88 |
if null ps then s else s ^ string_for_patterns ps
|
blanchet@38823
|
89 |
and string_for_patterns ps = "(" ^ commas (map string_for_pattern ps) ^ ")"
|
paulson@24287
|
90 |
|
paulson@24287
|
91 |
(*Is the second type an instance of the first one?*)
|
blanchet@38824
|
92 |
fun match_pattern (PVar, _) = true
|
blanchet@38824
|
93 |
| match_pattern (PApp _, PVar) = false
|
blanchet@38824
|
94 |
| match_pattern (PApp (s, ps), PApp (t, qs)) =
|
blanchet@38824
|
95 |
s = t andalso match_patterns (ps, qs)
|
blanchet@38824
|
96 |
and match_patterns (_, []) = true
|
blanchet@38824
|
97 |
| match_patterns ([], _) = false
|
blanchet@38824
|
98 |
| match_patterns (p :: ps, q :: qs) =
|
blanchet@38824
|
99 |
match_pattern (p, q) andalso match_patterns (ps, qs)
|
paulson@24287
|
100 |
|
blanchet@38823
|
101 |
(* Is there a unifiable constant? *)
|
blanchet@38827
|
102 |
fun pconst_mem f consts (s, ps) =
|
blanchet@38827
|
103 |
exists (curry (match_patterns o f) ps)
|
blanchet@38827
|
104 |
(map snd (filter (curry (op =) s o fst) consts))
|
blanchet@38827
|
105 |
fun pconst_hyper_mem f const_tab (s, ps) =
|
blanchet@38823
|
106 |
exists (curry (match_patterns o f) ps) (these (Symtab.lookup const_tab s))
|
blanchet@37505
|
107 |
|
blanchet@38827
|
108 |
fun ptype (Type (s, Ts)) = PApp (s, map ptype Ts)
|
blanchet@38827
|
109 |
| ptype (TFree (s, _)) = PApp (s, [])
|
blanchet@38827
|
110 |
| ptype (TVar _) = PVar
|
blanchet@38827
|
111 |
|
blanchet@38827
|
112 |
fun pterm thy t =
|
blanchet@38827
|
113 |
case strip_comb t of
|
blanchet@38827
|
114 |
(Const x, ts) => PApp (pconst thy true x ts)
|
blanchet@38827
|
115 |
| (Free x, ts) => PApp (pconst thy false x ts)
|
blanchet@38827
|
116 |
| (Var x, []) => PVar
|
blanchet@38827
|
117 |
| _ => PApp ("?", []) (* equivalence class of higher-order constructs *)
|
blanchet@38744
|
118 |
(* Pairs a constant with the list of its type instantiations. *)
|
blanchet@38827
|
119 |
and pconst_args thy const (s, T) ts =
|
blanchet@38828
|
120 |
(if const then map ptype (these (try (Sign.const_typargs thy) (s, T)))
|
blanchet@38828
|
121 |
else []) @
|
blanchet@38889
|
122 |
(if term_patterns then map (pterm thy) ts else [])
|
blanchet@38827
|
123 |
and pconst thy const (s, T) ts = (s, pconst_args thy const (s, T) ts)
|
paulson@24287
|
124 |
|
blanchet@38827
|
125 |
fun string_for_hyper_pconst (s, pss) =
|
blanchet@38823
|
126 |
s ^ "{" ^ commas (map string_for_patterns pss) ^ "}"
|
paulson@24287
|
127 |
|
blanchet@38816
|
128 |
val abs_name = "Sledgehammer.abs"
|
blanchet@38749
|
129 |
val skolem_prefix = "Sledgehammer.sko"
|
blanchet@38747
|
130 |
|
blanchet@38822
|
131 |
(* These are typically simplified away by "Meson.presimplify". Equality is
|
blanchet@38822
|
132 |
handled specially via "fequal". *)
|
blanchet@38682
|
133 |
val boring_consts =
|
blanchet@38822
|
134 |
[@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
|
haftmann@38864
|
135 |
@{const_name HOL.eq}]
|
blanchet@37537
|
136 |
|
blanchet@38823
|
137 |
(* Add a pconstant to the table, but a [] entry means a standard
|
blanchet@38819
|
138 |
connective, which we ignore.*)
|
blanchet@38824
|
139 |
fun add_pconst_to_table also_skolem (c, ps) =
|
blanchet@38819
|
140 |
if member (op =) boring_consts c orelse
|
blanchet@38819
|
141 |
(not also_skolem andalso String.isPrefix skolem_prefix c) then
|
blanchet@38819
|
142 |
I
|
blanchet@38819
|
143 |
else
|
blanchet@38824
|
144 |
Symtab.map_default (c, [ps]) (insert (op =) ps)
|
blanchet@38819
|
145 |
|
blanchet@38819
|
146 |
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT)
|
blanchet@38819
|
147 |
|
blanchet@38906
|
148 |
fun pconsts_in_terms thy also_skolems pos ts =
|
blanchet@37505
|
149 |
let
|
blanchet@38819
|
150 |
val flip = Option.map not
|
blanchet@38587
|
151 |
(* We include free variables, as well as constants, to handle locales. For
|
blanchet@38587
|
152 |
each quantifiers that must necessarily be skolemized by the ATP, we
|
blanchet@38587
|
153 |
introduce a fresh constant to simulate the effect of Skolemization. *)
|
blanchet@38827
|
154 |
fun do_const const (s, T) ts =
|
blanchet@38827
|
155 |
add_pconst_to_table also_skolems (pconst thy const (s, T) ts)
|
blanchet@38827
|
156 |
#> fold do_term ts
|
blanchet@38827
|
157 |
and do_term t =
|
blanchet@38827
|
158 |
case strip_comb t of
|
blanchet@38827
|
159 |
(Const x, ts) => do_const true x ts
|
blanchet@38827
|
160 |
| (Free x, ts) => do_const false x ts
|
blanchet@38827
|
161 |
| (Abs (_, _, t'), ts) =>
|
blanchet@38901
|
162 |
(null ts ? add_pconst_to_table true (abs_name, []))
|
blanchet@38827
|
163 |
#> fold do_term (t' :: ts)
|
blanchet@38827
|
164 |
| (_, ts) => fold do_term ts
|
blanchet@38587
|
165 |
fun do_quantifier will_surely_be_skolemized body_t =
|
blanchet@37537
|
166 |
do_formula pos body_t
|
blanchet@38747
|
167 |
#> (if also_skolems andalso will_surely_be_skolemized then
|
blanchet@38823
|
168 |
add_pconst_to_table true (gensym skolem_prefix, [])
|
blanchet@38587
|
169 |
else
|
blanchet@38587
|
170 |
I)
|
blanchet@38587
|
171 |
and do_term_or_formula T =
|
blanchet@38692
|
172 |
if is_formula_type T then do_formula NONE else do_term
|
blanchet@37537
|
173 |
and do_formula pos t =
|
blanchet@37537
|
174 |
case t of
|
blanchet@37537
|
175 |
Const (@{const_name all}, _) $ Abs (_, _, body_t) =>
|
blanchet@38587
|
176 |
do_quantifier (pos = SOME false) body_t
|
blanchet@37537
|
177 |
| @{const "==>"} $ t1 $ t2 =>
|
blanchet@37537
|
178 |
do_formula (flip pos) t1 #> do_formula pos t2
|
blanchet@37537
|
179 |
| Const (@{const_name "=="}, Type (_, [T, _])) $ t1 $ t2 =>
|
blanchet@38587
|
180 |
fold (do_term_or_formula T) [t1, t2]
|
blanchet@37537
|
181 |
| @{const Trueprop} $ t1 => do_formula pos t1
|
blanchet@37537
|
182 |
| @{const Not} $ t1 => do_formula (flip pos) t1
|
blanchet@37537
|
183 |
| Const (@{const_name All}, _) $ Abs (_, _, body_t) =>
|
blanchet@38587
|
184 |
do_quantifier (pos = SOME false) body_t
|
blanchet@37537
|
185 |
| Const (@{const_name Ex}, _) $ Abs (_, _, body_t) =>
|
blanchet@38587
|
186 |
do_quantifier (pos = SOME true) body_t
|
haftmann@38795
|
187 |
| @{const HOL.conj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
|
haftmann@38795
|
188 |
| @{const HOL.disj} $ t1 $ t2 => fold (do_formula pos) [t1, t2]
|
haftmann@38786
|
189 |
| @{const HOL.implies} $ t1 $ t2 =>
|
blanchet@37537
|
190 |
do_formula (flip pos) t1 #> do_formula pos t2
|
haftmann@38864
|
191 |
| Const (@{const_name HOL.eq}, Type (_, [T, _])) $ t1 $ t2 =>
|
blanchet@38587
|
192 |
fold (do_term_or_formula T) [t1, t2]
|
blanchet@38587
|
193 |
| Const (@{const_name If}, Type (_, [_, Type (_, [T, _])]))
|
blanchet@38587
|
194 |
$ t1 $ t2 $ t3 =>
|
blanchet@38587
|
195 |
do_formula NONE t1 #> fold (do_term_or_formula T) [t2, t3]
|
blanchet@38587
|
196 |
| Const (@{const_name Ex1}, _) $ Abs (_, _, body_t) =>
|
blanchet@38587
|
197 |
do_quantifier (is_some pos) body_t
|
blanchet@38587
|
198 |
| Const (@{const_name Ball}, _) $ t1 $ Abs (_, _, body_t) =>
|
blanchet@38587
|
199 |
do_quantifier (pos = SOME false)
|
blanchet@38587
|
200 |
(HOLogic.mk_imp (incr_boundvars 1 t1 $ Bound 0, body_t))
|
blanchet@38587
|
201 |
| Const (@{const_name Bex}, _) $ t1 $ Abs (_, _, body_t) =>
|
blanchet@38587
|
202 |
do_quantifier (pos = SOME true)
|
blanchet@38587
|
203 |
(HOLogic.mk_conj (incr_boundvars 1 t1 $ Bound 0, body_t))
|
blanchet@37537
|
204 |
| (t0 as Const (_, @{typ bool})) $ t1 =>
|
blanchet@37537
|
205 |
do_term t0 #> do_formula pos t1 (* theory constant *)
|
blanchet@37537
|
206 |
| _ => do_term t
|
blanchet@38819
|
207 |
in Symtab.empty |> fold (do_formula pos) ts end
|
paulson@24287
|
208 |
|
paulson@24287
|
209 |
(*Inserts a dummy "constant" referring to the theory name, so that relevance
|
paulson@24287
|
210 |
takes the given theory into account.*)
|
blanchet@37616
|
211 |
fun theory_const_prop_of theory_relevant th =
|
blanchet@37505
|
212 |
if theory_relevant then
|
blanchet@37505
|
213 |
let
|
blanchet@37505
|
214 |
val name = Context.theory_name (theory_of_thm th)
|
blanchet@37505
|
215 |
val t = Const (name ^ ". 1", @{typ bool})
|
blanchet@37505
|
216 |
in t $ prop_of th end
|
blanchet@37505
|
217 |
else
|
blanchet@37505
|
218 |
prop_of th
|
blanchet@37505
|
219 |
|
paulson@24287
|
220 |
(**** Constant / Type Frequencies ****)
|
paulson@24287
|
221 |
|
blanchet@38743
|
222 |
(* A two-dimensional symbol table counts frequencies of constants. It's keyed
|
blanchet@38743
|
223 |
first by constant name and second by its list of type instantiations. For the
|
blanchet@38823
|
224 |
latter, we need a linear ordering on "pattern list". *)
|
paulson@24287
|
225 |
|
blanchet@38823
|
226 |
fun pattern_ord p =
|
blanchet@38743
|
227 |
case p of
|
blanchet@38744
|
228 |
(PVar, PVar) => EQUAL
|
blanchet@38823
|
229 |
| (PVar, PApp _) => LESS
|
blanchet@38823
|
230 |
| (PApp _, PVar) => GREATER
|
blanchet@38823
|
231 |
| (PApp q1, PApp q2) =>
|
blanchet@38823
|
232 |
prod_ord fast_string_ord (dict_ord pattern_ord) (q1, q2)
|
paulson@24287
|
233 |
|
blanchet@38743
|
234 |
structure CTtab =
|
blanchet@38823
|
235 |
Table(type key = pattern list val ord = dict_ord pattern_ord)
|
paulson@24287
|
236 |
|
blanchet@38827
|
237 |
fun count_axiom_consts theory_relevant thy =
|
blanchet@37503
|
238 |
let
|
blanchet@38827
|
239 |
fun do_const const (s, T) ts =
|
blanchet@38827
|
240 |
(* Two-dimensional table update. Constant maps to types maps to count. *)
|
blanchet@38827
|
241 |
CTtab.map_default (pconst_args thy const (s, T) ts, 0) (Integer.add 1)
|
blanchet@38827
|
242 |
|> Symtab.map_default (s, CTtab.empty)
|
blanchet@38827
|
243 |
#> fold do_term ts
|
blanchet@38827
|
244 |
and do_term t =
|
blanchet@38827
|
245 |
case strip_comb t of
|
blanchet@38827
|
246 |
(Const x, ts) => do_const true x ts
|
blanchet@38827
|
247 |
| (Free x, ts) => do_const false x ts
|
blanchet@38827
|
248 |
| (Abs (_, _, t'), ts) => fold do_term (t' :: ts)
|
blanchet@38827
|
249 |
| (_, ts) => fold do_term ts
|
blanchet@38827
|
250 |
in do_term o theory_const_prop_of theory_relevant o snd end
|
paulson@24287
|
251 |
|
paulson@24287
|
252 |
|
paulson@24287
|
253 |
(**** Actual Filtering Code ****)
|
paulson@24287
|
254 |
|
paulson@24287
|
255 |
(*The frequency of a constant is the sum of those of all instances of its type.*)
|
blanchet@38824
|
256 |
fun pconst_freq match const_tab (c, ps) =
|
blanchet@38824
|
257 |
CTtab.fold (fn (qs, m) => match (ps, qs) ? Integer.add m)
|
blanchet@38686
|
258 |
(the (Symtab.lookup const_tab c)) 0
|
blanchet@38686
|
259 |
|
paulson@24287
|
260 |
|
blanchet@38085
|
261 |
(* A surprising number of theorems contain only a few significant constants.
|
blanchet@38085
|
262 |
These include all induction rules, and other general theorems. *)
|
blanchet@37503
|
263 |
|
blanchet@37503
|
264 |
(* "log" seems best in practice. A constant function of one ignores the constant
|
blanchet@38938
|
265 |
frequencies. Rare constants give more points if they are relevant than less
|
blanchet@38938
|
266 |
rare ones. *)
|
blanchet@38938
|
267 |
fun rel_weight_for_freq n = 1.0 + 2.0 / Math.ln (Real.fromInt n + 1.0)
|
blanchet@38938
|
268 |
|
blanchet@38938
|
269 |
(* FUDGE *)
|
blanchet@38938
|
270 |
val worse_irrel_freq = Unsynchronized.ref 100.0
|
blanchet@38938
|
271 |
|
blanchet@38938
|
272 |
(* Irrelevant constants are treated differently. We associate lower penalties to
|
blanchet@38938
|
273 |
very rare constants and very common ones -- the former because they can't
|
blanchet@38938
|
274 |
lead to the inclusion of too many new facts, and the latter because they are
|
blanchet@38938
|
275 |
so common as to be of little interest. *)
|
blanchet@38938
|
276 |
fun irrel_weight_for_freq n =
|
blanchet@38938
|
277 |
let val (k, x) = !worse_irrel_freq |> `Real.ceil in
|
blanchet@38938
|
278 |
if n < k then Math.ln (Real.fromInt (n + 1)) / Math.ln x
|
blanchet@38938
|
279 |
else rel_weight_for_freq n / rel_weight_for_freq k
|
blanchet@38938
|
280 |
end
|
blanchet@37503
|
281 |
|
blanchet@38752
|
282 |
(* FUDGE *)
|
blanchet@38899
|
283 |
val abs_rel_weight = Unsynchronized.ref 0.5
|
blanchet@38899
|
284 |
val abs_irrel_weight = Unsynchronized.ref 2.0
|
blanchet@38904
|
285 |
val skolem_irrel_weight = Unsynchronized.ref 0.75
|
blanchet@38752
|
286 |
|
blanchet@37503
|
287 |
(* Computes a constant's weight, as determined by its frequency. *)
|
blanchet@38938
|
288 |
fun generic_pconst_weight abs_weight skolem_weight weight_for_freq f const_tab
|
blanchet@38938
|
289 |
(c as (s, _)) =
|
blanchet@38816
|
290 |
if s = abs_name then abs_weight
|
blanchet@38816
|
291 |
else if String.isPrefix skolem_prefix s then skolem_weight
|
blanchet@38938
|
292 |
else weight_for_freq (pconst_freq (match_patterns o f) const_tab c)
|
blanchet@38821
|
293 |
|
blanchet@38938
|
294 |
fun rel_pconst_weight const_tab =
|
blanchet@38938
|
295 |
generic_pconst_weight (!abs_rel_weight) 0.0 rel_weight_for_freq I const_tab
|
blanchet@38938
|
296 |
fun irrel_pconst_weight const_tab =
|
blanchet@38938
|
297 |
generic_pconst_weight (!abs_irrel_weight) (!skolem_irrel_weight)
|
blanchet@38938
|
298 |
irrel_weight_for_freq swap const_tab
|
paulson@24287
|
299 |
|
blanchet@38752
|
300 |
(* FUDGE *)
|
blanchet@38937
|
301 |
val intro_bonus = Unsynchronized.ref 0.15
|
blanchet@38937
|
302 |
val elim_bonus = Unsynchronized.ref 0.15
|
blanchet@38937
|
303 |
val simp_bonus = Unsynchronized.ref 0.15
|
blanchet@38906
|
304 |
val local_bonus = Unsynchronized.ref 0.55
|
blanchet@38901
|
305 |
val chained_bonus = Unsynchronized.ref 1.5
|
blanchet@38899
|
306 |
|
blanchet@38889
|
307 |
fun locality_bonus General = 0.0
|
blanchet@38937
|
308 |
| locality_bonus Intro = !intro_bonus
|
blanchet@38937
|
309 |
| locality_bonus Elim = !elim_bonus
|
blanchet@38937
|
310 |
| locality_bonus Simp = !simp_bonus
|
blanchet@38899
|
311 |
| locality_bonus Local = !local_bonus
|
blanchet@38899
|
312 |
| locality_bonus Chained = !chained_bonus
|
blanchet@38751
|
313 |
|
blanchet@38752
|
314 |
fun axiom_weight loc const_tab relevant_consts axiom_consts =
|
blanchet@38827
|
315 |
case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
|
blanchet@38827
|
316 |
||> filter_out (pconst_hyper_mem swap relevant_consts) of
|
blanchet@38827
|
317 |
([], _) => 0.0
|
blanchet@38744
|
318 |
| (rel, irrel) =>
|
blanchet@38889
|
319 |
let
|
blanchet@38889
|
320 |
val irrel = irrel |> filter_out (pconst_mem swap rel)
|
blanchet@38938
|
321 |
val rel_weight =
|
blanchet@38938
|
322 |
0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
|
blanchet@38889
|
323 |
val irrel_weight =
|
blanchet@38889
|
324 |
~ (locality_bonus loc)
|
blanchet@38938
|
325 |
|> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
|
blanchet@38889
|
326 |
val res = rel_weight / (rel_weight + irrel_weight)
|
blanchet@38889
|
327 |
in if Real.isFinite res then res else 0.0 end
|
blanchet@38747
|
328 |
|
blanchet@38904
|
329 |
(* FIXME: experiment
|
blanchet@38904
|
330 |
fun debug_axiom_weight loc const_tab relevant_consts axiom_consts =
|
blanchet@38904
|
331 |
case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
|
blanchet@38904
|
332 |
||> filter_out (pconst_hyper_mem swap relevant_consts) of
|
blanchet@38904
|
333 |
([], _) => 0.0
|
blanchet@38904
|
334 |
| (rel, irrel) =>
|
blanchet@38904
|
335 |
let
|
blanchet@38904
|
336 |
val irrel = irrel |> filter_out (pconst_mem swap rel)
|
blanchet@38938
|
337 |
val rels_weight =
|
blanchet@38938
|
338 |
0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
|
blanchet@38938
|
339 |
val irrels_weight =
|
blanchet@38904
|
340 |
~ (locality_bonus loc)
|
blanchet@38938
|
341 |
|> fold (curry (op +) o irrel_pconst_weight const_tab) irrel
|
blanchet@38938
|
342 |
val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel))
|
blanchet@38938
|
343 |
val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight const_tab)) irrel))(*###*)
|
blanchet@38938
|
344 |
val res = rels_weight / (rels_weight + irrels_weight)
|
blanchet@38904
|
345 |
in if Real.isFinite res then res else 0.0 end
|
blanchet@38904
|
346 |
*)
|
blanchet@38904
|
347 |
|
blanchet@38823
|
348 |
fun pconsts_in_axiom thy t =
|
blanchet@38825
|
349 |
Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
|
blanchet@38906
|
350 |
(pconsts_in_terms thy true (SOME true) [t]) []
|
blanchet@38687
|
351 |
fun pair_consts_axiom theory_relevant thy axiom =
|
blanchet@38827
|
352 |
case axiom |> snd |> theory_const_prop_of theory_relevant
|
blanchet@38827
|
353 |
|> pconsts_in_axiom thy of
|
blanchet@38827
|
354 |
[] => NONE
|
blanchet@38827
|
355 |
| consts => SOME ((axiom, consts), NONE)
|
paulson@24287
|
356 |
|
blanchet@38699
|
357 |
type annotated_thm =
|
blanchet@38823
|
358 |
(((unit -> string) * locality) * thm) * (string * pattern list) list
|
blanchet@37505
|
359 |
|
blanchet@38904
|
360 |
(* FUDGE *)
|
blanchet@38937
|
361 |
val max_imperfect = Unsynchronized.ref 11.5
|
blanchet@38904
|
362 |
val max_imperfect_exp = Unsynchronized.ref 1.0
|
blanchet@38889
|
363 |
|
blanchet@38904
|
364 |
fun take_most_relevant max_relevant remaining_max
|
blanchet@38747
|
365 |
(candidates : (annotated_thm * real) list) =
|
blanchet@38744
|
366 |
let
|
blanchet@38747
|
367 |
val max_imperfect =
|
blanchet@38904
|
368 |
Real.ceil (Math.pow (!max_imperfect,
|
blanchet@38904
|
369 |
Math.pow (Real.fromInt remaining_max
|
blanchet@38904
|
370 |
/ Real.fromInt max_relevant, !max_imperfect_exp)))
|
blanchet@38747
|
371 |
val (perfect, imperfect) =
|
blanchet@38889
|
372 |
candidates |> sort (Real.compare o swap o pairself snd)
|
blanchet@38889
|
373 |
|> take_prefix (fn (_, w) => w > 0.99999)
|
blanchet@38747
|
374 |
val ((accepts, more_rejects), rejects) =
|
blanchet@38747
|
375 |
chop max_imperfect imperfect |>> append perfect |>> chop remaining_max
|
blanchet@38744
|
376 |
in
|
blanchet@38889
|
377 |
trace_msg (fn () =>
|
blanchet@38889
|
378 |
"Actually passed (" ^ Int.toString (length accepts) ^ " of " ^
|
blanchet@38889
|
379 |
Int.toString (length candidates) ^ "): " ^
|
blanchet@38889
|
380 |
(accepts |> map (fn ((((name, _), _), _), weight) =>
|
blanchet@38752
|
381 |
name () ^ " [" ^ Real.toString weight ^ "]")
|
blanchet@38745
|
382 |
|> commas));
|
blanchet@38747
|
383 |
(accepts, more_rejects @ rejects)
|
blanchet@38744
|
384 |
end
|
paulson@24287
|
385 |
|
blanchet@38819
|
386 |
fun if_empty_replace_with_locality thy axioms loc tab =
|
blanchet@38819
|
387 |
if Symtab.is_empty tab then
|
blanchet@38906
|
388 |
pconsts_in_terms thy false (SOME false)
|
blanchet@38819
|
389 |
(map_filter (fn ((_, loc'), th) =>
|
blanchet@38819
|
390 |
if loc' = loc then SOME (prop_of th) else NONE) axioms)
|
blanchet@38819
|
391 |
else
|
blanchet@38819
|
392 |
tab
|
blanchet@38819
|
393 |
|
blanchet@38752
|
394 |
(* FUDGE *)
|
blanchet@38899
|
395 |
val threshold_divisor = Unsynchronized.ref 2.0
|
blanchet@38899
|
396 |
val ridiculous_threshold = Unsynchronized.ref 0.1
|
blanchet@38683
|
397 |
|
blanchet@38745
|
398 |
fun relevance_filter ctxt threshold0 decay max_relevant theory_relevant
|
blanchet@38745
|
399 |
({add, del, ...} : relevance_override) axioms goal_ts =
|
blanchet@38739
|
400 |
let
|
blanchet@38739
|
401 |
val thy = ProofContext.theory_of ctxt
|
blanchet@38819
|
402 |
val const_tab =
|
blanchet@38819
|
403 |
fold (count_axiom_consts theory_relevant thy) axioms Symtab.empty
|
blanchet@38819
|
404 |
val goal_const_tab =
|
blanchet@38906
|
405 |
pconsts_in_terms thy false (SOME false) goal_ts
|
blanchet@38937
|
406 |
|> fold (if_empty_replace_with_locality thy axioms) [Chained, Local]
|
blanchet@38739
|
407 |
val add_thms = maps (ProofContext.get_fact ctxt) add
|
blanchet@38739
|
408 |
val del_thms = maps (ProofContext.get_fact ctxt) del
|
blanchet@38747
|
409 |
fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
|
blanchet@38739
|
410 |
let
|
blanchet@38744
|
411 |
fun game_over rejects =
|
blanchet@38747
|
412 |
(* Add "add:" facts. *)
|
blanchet@38747
|
413 |
if null add_thms then
|
blanchet@38747
|
414 |
[]
|
blanchet@38744
|
415 |
else
|
blanchet@38747
|
416 |
map_filter (fn ((p as (_, th), _), _) =>
|
blanchet@38747
|
417 |
if member Thm.eq_thm add_thms th then SOME p
|
blanchet@38747
|
418 |
else NONE) rejects
|
blanchet@38889
|
419 |
fun relevant [] rejects [] =
|
blanchet@38747
|
420 |
(* Nothing has been added this iteration. *)
|
blanchet@38899
|
421 |
if j = 0 andalso threshold >= !ridiculous_threshold then
|
blanchet@38747
|
422 |
(* First iteration? Try again. *)
|
blanchet@38899
|
423 |
iter 0 max_relevant (threshold / !threshold_divisor) rel_const_tab
|
blanchet@38747
|
424 |
hopeless hopeful
|
blanchet@38744
|
425 |
else
|
blanchet@38747
|
426 |
game_over (rejects @ hopeless)
|
blanchet@38889
|
427 |
| relevant candidates rejects [] =
|
blanchet@38739
|
428 |
let
|
blanchet@38747
|
429 |
val (accepts, more_rejects) =
|
blanchet@38904
|
430 |
take_most_relevant max_relevant remaining_max candidates
|
blanchet@38739
|
431 |
val rel_const_tab' =
|
blanchet@38745
|
432 |
rel_const_tab
|
blanchet@38901
|
433 |
|> fold (add_pconst_to_table false) (maps (snd o fst) accepts)
|
blanchet@38744
|
434 |
fun is_dirty (c, _) =
|
blanchet@38744
|
435 |
Symtab.lookup rel_const_tab' c <> Symtab.lookup rel_const_tab c
|
blanchet@38745
|
436 |
val (hopeful_rejects, hopeless_rejects) =
|
blanchet@38745
|
437 |
(rejects @ hopeless, ([], []))
|
blanchet@38745
|
438 |
|-> fold (fn (ax as (_, consts), old_weight) =>
|
blanchet@38745
|
439 |
if exists is_dirty consts then
|
blanchet@38745
|
440 |
apfst (cons (ax, NONE))
|
blanchet@38745
|
441 |
else
|
blanchet@38745
|
442 |
apsnd (cons (ax, old_weight)))
|
blanchet@38745
|
443 |
|>> append (more_rejects
|
blanchet@38745
|
444 |
|> map (fn (ax as (_, consts), old_weight) =>
|
blanchet@38745
|
445 |
(ax, if exists is_dirty consts then NONE
|
blanchet@38745
|
446 |
else SOME old_weight)))
|
blanchet@38747
|
447 |
val threshold =
|
blanchet@38822
|
448 |
1.0 - (1.0 - threshold)
|
blanchet@38822
|
449 |
* Math.pow (decay, Real.fromInt (length accepts))
|
blanchet@38747
|
450 |
val remaining_max = remaining_max - length accepts
|
blanchet@38739
|
451 |
in
|
blanchet@38744
|
452 |
trace_msg (fn () => "New or updated constants: " ^
|
blanchet@38744
|
453 |
commas (rel_const_tab' |> Symtab.dest
|
blanchet@38822
|
454 |
|> subtract (op =) (rel_const_tab |> Symtab.dest)
|
blanchet@38827
|
455 |
|> map string_for_hyper_pconst));
|
blanchet@38745
|
456 |
map (fst o fst) accepts @
|
blanchet@38747
|
457 |
(if remaining_max = 0 then
|
blanchet@38745
|
458 |
game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
|
blanchet@38745
|
459 |
else
|
blanchet@38747
|
460 |
iter (j + 1) remaining_max threshold rel_const_tab'
|
blanchet@38747
|
461 |
hopeless_rejects hopeful_rejects)
|
blanchet@38739
|
462 |
end
|
blanchet@38889
|
463 |
| relevant candidates rejects
|
blanchet@38752
|
464 |
(((ax as (((_, loc), th), axiom_consts)), cached_weight)
|
blanchet@38747
|
465 |
:: hopeful) =
|
blanchet@38739
|
466 |
let
|
blanchet@38739
|
467 |
val weight =
|
blanchet@38739
|
468 |
case cached_weight of
|
blanchet@38739
|
469 |
SOME w => w
|
blanchet@38752
|
470 |
| NONE => axiom_weight loc const_tab rel_const_tab axiom_consts
|
blanchet@38904
|
471 |
(* FIXME: experiment
|
blanchet@38752
|
472 |
val name = fst (fst (fst ax)) ()
|
blanchet@38938
|
473 |
val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then
|
blanchet@38901
|
474 |
tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight loc const_tab rel_const_tab axiom_consts))
|
blanchet@38747
|
475 |
else
|
blanchet@38747
|
476 |
()
|
blanchet@38747
|
477 |
*)
|
blanchet@38739
|
478 |
in
|
blanchet@38741
|
479 |
if weight >= threshold then
|
blanchet@38889
|
480 |
relevant ((ax, weight) :: candidates) rejects hopeful
|
blanchet@38739
|
481 |
else
|
blanchet@38889
|
482 |
relevant candidates ((ax, weight) :: rejects) hopeful
|
blanchet@38739
|
483 |
end
|
blanchet@38739
|
484 |
in
|
blanchet@38744
|
485 |
trace_msg (fn () =>
|
blanchet@38744
|
486 |
"ITERATION " ^ string_of_int j ^ ": current threshold: " ^
|
blanchet@38744
|
487 |
Real.toString threshold ^ ", constants: " ^
|
blanchet@38744
|
488 |
commas (rel_const_tab |> Symtab.dest
|
blanchet@38744
|
489 |
|> filter (curry (op <>) [] o snd)
|
blanchet@38827
|
490 |
|> map string_for_hyper_pconst));
|
blanchet@38889
|
491 |
relevant [] [] hopeful
|
blanchet@38739
|
492 |
end
|
blanchet@38739
|
493 |
in
|
blanchet@38739
|
494 |
axioms |> filter_out (member Thm.eq_thm del_thms o snd)
|
blanchet@38827
|
495 |
|> map_filter (pair_consts_axiom theory_relevant thy)
|
blanchet@38819
|
496 |
|> iter 0 max_relevant threshold0 goal_const_tab []
|
blanchet@38739
|
497 |
|> tap (fn res => trace_msg (fn () =>
|
blanchet@38686
|
498 |
"Total relevant: " ^ Int.toString (length res)))
|
blanchet@38739
|
499 |
end
|
paulson@24287
|
500 |
|
blanchet@38744
|
501 |
|
paulson@24287
|
502 |
(***************************************************************)
|
mengj@19768
|
503 |
(* Retrieving and filtering lemmas *)
|
mengj@19768
|
504 |
(***************************************************************)
|
mengj@19768
|
505 |
|
paulson@33022
|
506 |
(*** retrieve lemmas and filter them ***)
|
mengj@19768
|
507 |
|
paulson@20757
|
508 |
(*Reject theorems with names like "List.filter.filter_list_def" or
|
paulson@21690
|
509 |
"Accessible_Part.acc.defs", as these are definitions arising from packages.*)
|
paulson@20757
|
510 |
fun is_package_def a =
|
wenzelm@30364
|
511 |
let val names = Long_Name.explode a
|
paulson@21690
|
512 |
in
|
paulson@21690
|
513 |
length names > 2 andalso
|
paulson@21690
|
514 |
not (hd names = "local") andalso
|
paulson@21690
|
515 |
String.isSuffix "_def" a orelse String.isSuffix "_defs" a
|
paulson@21690
|
516 |
end;
|
paulson@20757
|
517 |
|
blanchet@38937
|
518 |
fun mk_fact_table f xs =
|
blanchet@38937
|
519 |
fold (Termtab.update o `(prop_of o f)) xs Termtab.empty
|
blanchet@38937
|
520 |
fun uniquify xs = Termtab.fold (cons o snd) (mk_fact_table snd xs) []
|
mengj@19768
|
521 |
|
blanchet@37626
|
522 |
(* FIXME: put other record thms here, or declare as "no_atp" *)
|
blanchet@37626
|
523 |
val multi_base_blacklist =
|
blanchet@37626
|
524 |
["defs", "select_defs", "update_defs", "induct", "inducts", "split", "splits",
|
blanchet@38682
|
525 |
"split_asm", "cases", "ext_cases", "eq.simps", "eq.refl", "nchotomy",
|
blanchet@38682
|
526 |
"case_cong", "weak_case_cong"]
|
blanchet@38682
|
527 |
|> map (prefix ".")
|
blanchet@37626
|
528 |
|
blanchet@37626
|
529 |
val max_lambda_nesting = 3
|
blanchet@37626
|
530 |
|
blanchet@37626
|
531 |
fun term_has_too_many_lambdas max (t1 $ t2) =
|
blanchet@37626
|
532 |
exists (term_has_too_many_lambdas max) [t1, t2]
|
blanchet@37626
|
533 |
| term_has_too_many_lambdas max (Abs (_, _, t)) =
|
blanchet@37626
|
534 |
max = 0 orelse term_has_too_many_lambdas (max - 1) t
|
blanchet@37626
|
535 |
| term_has_too_many_lambdas _ _ = false
|
blanchet@37626
|
536 |
|
blanchet@37626
|
537 |
(* Don't count nested lambdas at the level of formulas, since they are
|
blanchet@37626
|
538 |
quantifiers. *)
|
blanchet@37626
|
539 |
fun formula_has_too_many_lambdas Ts (Abs (_, T, t)) =
|
blanchet@37626
|
540 |
formula_has_too_many_lambdas (T :: Ts) t
|
blanchet@37626
|
541 |
| formula_has_too_many_lambdas Ts t =
|
blanchet@37626
|
542 |
if is_formula_type (fastype_of1 (Ts, t)) then
|
blanchet@37626
|
543 |
exists (formula_has_too_many_lambdas Ts) (#2 (strip_comb t))
|
blanchet@37626
|
544 |
else
|
blanchet@37626
|
545 |
term_has_too_many_lambdas max_lambda_nesting t
|
blanchet@37626
|
546 |
|
blanchet@38692
|
547 |
(* The max apply depth of any "metis" call in "Metis_Examples" (on 2007-10-31)
|
blanchet@37626
|
548 |
was 11. *)
|
blanchet@37626
|
549 |
val max_apply_depth = 15
|
blanchet@37626
|
550 |
|
blanchet@37626
|
551 |
fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
|
blanchet@37626
|
552 |
| apply_depth (Abs (_, _, t)) = apply_depth t
|
blanchet@37626
|
553 |
| apply_depth _ = 0
|
blanchet@37626
|
554 |
|
blanchet@37626
|
555 |
fun is_formula_too_complex t =
|
blanchet@38085
|
556 |
apply_depth t > max_apply_depth orelse formula_has_too_many_lambdas [] t
|
blanchet@37626
|
557 |
|
blanchet@37543
|
558 |
val exists_sledgehammer_const =
|
blanchet@37626
|
559 |
exists_Const (fn (s, _) => String.isPrefix sledgehammer_prefix s)
|
blanchet@37626
|
560 |
|
blanchet@38904
|
561 |
(* FIXME: make more reliable *)
|
blanchet@38904
|
562 |
val exists_low_level_class_const =
|
blanchet@38904
|
563 |
exists_Const (fn (s, _) =>
|
blanchet@38904
|
564 |
String.isSubstring (Long_Name.separator ^ "class" ^ Long_Name.separator) s)
|
blanchet@38904
|
565 |
|
blanchet@38821
|
566 |
fun is_metastrange_theorem th =
|
blanchet@37626
|
567 |
case head_of (concl_of th) of
|
blanchet@37626
|
568 |
Const (a, _) => (a <> @{const_name Trueprop} andalso
|
blanchet@37626
|
569 |
a <> @{const_name "=="})
|
blanchet@37626
|
570 |
| _ => false
|
blanchet@37626
|
571 |
|
blanchet@38821
|
572 |
fun is_that_fact th =
|
blanchet@38821
|
573 |
String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
|
blanchet@38821
|
574 |
andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
|
blanchet@38821
|
575 |
| _ => false) (prop_of th)
|
blanchet@38821
|
576 |
|
blanchet@37626
|
577 |
val type_has_top_sort =
|
blanchet@37626
|
578 |
exists_subtype (fn TFree (_, []) => true | TVar (_, []) => true | _ => false)
|
blanchet@37626
|
579 |
|
blanchet@38085
|
580 |
(**** Predicates to detect unwanted facts (prolific or likely to cause
|
blanchet@37347
|
581 |
unsoundness) ****)
|
paulson@21470
|
582 |
|
blanchet@38289
|
583 |
(* Too general means, positive equality literal with a variable X as one
|
blanchet@38289
|
584 |
operand, when X does not occur properly in the other operand. This rules out
|
blanchet@38289
|
585 |
clearly inconsistent facts such as X = a | X = b, though it by no means
|
blanchet@38289
|
586 |
guarantees soundness. *)
|
paulson@21470
|
587 |
|
blanchet@38289
|
588 |
(* Unwanted equalities are those between a (bound or schematic) variable that
|
blanchet@38289
|
589 |
does not properly occur in the second operand. *)
|
blanchet@38607
|
590 |
val is_exhaustive_finite =
|
blanchet@38607
|
591 |
let
|
blanchet@38629
|
592 |
fun is_bad_equal (Var z) t =
|
blanchet@38629
|
593 |
not (exists_subterm (fn Var z' => z = z' | _ => false) t)
|
blanchet@38629
|
594 |
| is_bad_equal (Bound j) t = not (loose_bvar1 (t, j))
|
blanchet@38629
|
595 |
| is_bad_equal _ _ = false
|
blanchet@38629
|
596 |
fun do_equals t1 t2 = is_bad_equal t1 t2 orelse is_bad_equal t2 t1
|
blanchet@38607
|
597 |
fun do_formula pos t =
|
blanchet@38607
|
598 |
case (pos, t) of
|
blanchet@38615
|
599 |
(_, @{const Trueprop} $ t1) => do_formula pos t1
|
blanchet@38607
|
600 |
| (true, Const (@{const_name all}, _) $ Abs (_, _, t')) =>
|
blanchet@38607
|
601 |
do_formula pos t'
|
blanchet@38607
|
602 |
| (true, Const (@{const_name All}, _) $ Abs (_, _, t')) =>
|
blanchet@38607
|
603 |
do_formula pos t'
|
blanchet@38607
|
604 |
| (false, Const (@{const_name Ex}, _) $ Abs (_, _, t')) =>
|
blanchet@38607
|
605 |
do_formula pos t'
|
blanchet@38607
|
606 |
| (_, @{const "==>"} $ t1 $ t2) =>
|
blanchet@38629
|
607 |
do_formula (not pos) t1 andalso
|
blanchet@38629
|
608 |
(t2 = @{prop False} orelse do_formula pos t2)
|
haftmann@38786
|
609 |
| (_, @{const HOL.implies} $ t1 $ t2) =>
|
blanchet@38629
|
610 |
do_formula (not pos) t1 andalso
|
blanchet@38629
|
611 |
(t2 = @{const False} orelse do_formula pos t2)
|
blanchet@38607
|
612 |
| (_, @{const Not} $ t1) => do_formula (not pos) t1
|
haftmann@38795
|
613 |
| (true, @{const HOL.disj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
|
haftmann@38795
|
614 |
| (false, @{const HOL.conj} $ t1 $ t2) => forall (do_formula pos) [t1, t2]
|
haftmann@38864
|
615 |
| (true, Const (@{const_name HOL.eq}, _) $ t1 $ t2) => do_equals t1 t2
|
blanchet@38607
|
616 |
| (true, Const (@{const_name "=="}, _) $ t1 $ t2) => do_equals t1 t2
|
blanchet@38607
|
617 |
| _ => false
|
blanchet@38607
|
618 |
in do_formula true end
|
blanchet@38607
|
619 |
|
blanchet@38592
|
620 |
fun has_bound_or_var_of_type tycons =
|
blanchet@38592
|
621 |
exists_subterm (fn Var (_, Type (s, _)) => member (op =) tycons s
|
blanchet@38592
|
622 |
| Abs (_, Type (s, _), _) => member (op =) tycons s
|
blanchet@38592
|
623 |
| _ => false)
|
paulson@21431
|
624 |
|
blanchet@38085
|
625 |
(* Facts are forbidden to contain variables of these types. The typical reason
|
blanchet@37347
|
626 |
is that they lead to unsoundness. Note that "unit" satisfies numerous
|
blanchet@38085
|
627 |
equations like "?x = ()". The resulting clauses will have no type constraint,
|
blanchet@37347
|
628 |
yielding false proofs. Even "bool" leads to many unsound proofs, though only
|
blanchet@37347
|
629 |
for higher-order problems. *)
|
blanchet@38592
|
630 |
val dangerous_types = [@{type_name unit}, @{type_name bool}, @{type_name prop}];
|
paulson@22217
|
631 |
|
blanchet@38085
|
632 |
(* Facts containing variables of type "unit" or "bool" or of the form
|
blanchet@38290
|
633 |
"ALL x. x = A | x = B | x = C" are likely to lead to unsound proofs if types
|
blanchet@38290
|
634 |
are omitted. *)
|
blanchet@38593
|
635 |
fun is_dangerous_term full_types t =
|
blanchet@38609
|
636 |
not full_types andalso
|
blanchet@38679
|
637 |
let val t = transform_elim_term t in
|
blanchet@38679
|
638 |
has_bound_or_var_of_type dangerous_types t orelse
|
blanchet@38679
|
639 |
is_exhaustive_finite t
|
blanchet@38679
|
640 |
end
|
paulson@21470
|
641 |
|
blanchet@38627
|
642 |
fun is_theorem_bad_for_atps full_types thm =
|
blanchet@38627
|
643 |
let val t = prop_of thm in
|
blanchet@38627
|
644 |
is_formula_too_complex t orelse exists_type type_has_top_sort t orelse
|
blanchet@38627
|
645 |
is_dangerous_term full_types t orelse exists_sledgehammer_const t orelse
|
blanchet@38904
|
646 |
exists_low_level_class_const t orelse is_metastrange_theorem thm orelse
|
blanchet@38904
|
647 |
is_that_fact thm
|
blanchet@38627
|
648 |
end
|
blanchet@38627
|
649 |
|
blanchet@38937
|
650 |
fun clasimpset_rules_of ctxt =
|
blanchet@38937
|
651 |
let
|
blanchet@38937
|
652 |
val {safeIs, safeEs, hazIs, hazEs, ...} = ctxt |> claset_of |> rep_cs
|
blanchet@38937
|
653 |
val intros = safeIs @ hazIs
|
blanchet@38937
|
654 |
val elims = map Classical.classical_rule (safeEs @ hazEs)
|
blanchet@38937
|
655 |
val simps = ctxt |> simpset_of |> dest_ss |> #simps |> map snd
|
blanchet@38937
|
656 |
in (mk_fact_table I intros, mk_fact_table I elims, mk_fact_table I simps) end
|
blanchet@38937
|
657 |
|
blanchet@38696
|
658 |
fun all_name_thms_pairs ctxt reserved full_types add_thms chained_ths =
|
blanchet@38627
|
659 |
let
|
blanchet@38752
|
660 |
val thy = ProofContext.theory_of ctxt
|
blanchet@38752
|
661 |
val global_facts = PureThy.facts_of thy
|
blanchet@38644
|
662 |
val local_facts = ProofContext.facts_of ctxt
|
blanchet@38644
|
663 |
val named_locals = local_facts |> Facts.dest_static []
|
blanchet@38752
|
664 |
val is_chained = member Thm.eq_thm chained_ths
|
blanchet@38937
|
665 |
val (intros, elims, simps) =
|
blanchet@38937
|
666 |
if exists (curry (op <) 0.0) [!intro_bonus, !elim_bonus, !simp_bonus] then
|
blanchet@38937
|
667 |
clasimpset_rules_of ctxt
|
blanchet@38937
|
668 |
else
|
blanchet@38937
|
669 |
(Termtab.empty, Termtab.empty, Termtab.empty)
|
blanchet@38818
|
670 |
(* Unnamed nonchained formulas with schematic variables are omitted, because
|
blanchet@38818
|
671 |
they are rejected by the backticks (`...`) parser for some reason. *)
|
blanchet@38738
|
672 |
fun is_good_unnamed_local th =
|
blanchet@38820
|
673 |
not (Thm.has_name_hint th) andalso
|
blanchet@38820
|
674 |
(not (exists_subterm is_Var (prop_of th)) orelse (is_chained th)) andalso
|
blanchet@38738
|
675 |
forall (fn (_, ths) => not (member Thm.eq_thm ths th)) named_locals
|
blanchet@38644
|
676 |
val unnamed_locals =
|
blanchet@38820
|
677 |
union Thm.eq_thm (Facts.props local_facts) chained_ths
|
blanchet@38820
|
678 |
|> filter is_good_unnamed_local |> map (pair "" o single)
|
blanchet@38627
|
679 |
val full_space =
|
blanchet@38738
|
680 |
Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
|
blanchet@38752
|
681 |
fun add_facts global foldx facts =
|
blanchet@38699
|
682 |
foldx (fn (name0, ths) =>
|
blanchet@38699
|
683 |
if name0 <> "" andalso
|
blanchet@38699
|
684 |
forall (not o member Thm.eq_thm add_thms) ths andalso
|
blanchet@38699
|
685 |
(Facts.is_concealed facts name0 orelse
|
blanchet@38699
|
686 |
(respect_no_atp andalso is_package_def name0) orelse
|
blanchet@38699
|
687 |
exists (fn s => String.isSuffix s name0) multi_base_blacklist orelse
|
blanchet@38699
|
688 |
String.isSuffix "_def_raw" (* FIXME: crude hack *) name0) then
|
blanchet@38627
|
689 |
I
|
blanchet@38627
|
690 |
else
|
blanchet@38627
|
691 |
let
|
blanchet@38699
|
692 |
val multi = length ths > 1
|
blanchet@38696
|
693 |
fun backquotify th =
|
blanchet@38696
|
694 |
"`" ^ Print_Mode.setmp [Print_Mode.input]
|
blanchet@38696
|
695 |
(Syntax.string_of_term ctxt) (prop_of th) ^ "`"
|
blanchet@38738
|
696 |
|> String.translate (fn c => if Char.isPrint c then str c else "")
|
blanchet@38738
|
697 |
|> simplify_spaces
|
blanchet@38699
|
698 |
fun check_thms a =
|
blanchet@38699
|
699 |
case try (ProofContext.get_thms ctxt) a of
|
blanchet@38699
|
700 |
NONE => false
|
blanchet@38699
|
701 |
| SOME ths' => Thm.eq_thms (ths, ths')
|
blanchet@38627
|
702 |
in
|
blanchet@38699
|
703 |
pair 1
|
blanchet@38699
|
704 |
#> fold (fn th => fn (j, rest) =>
|
blanchet@38699
|
705 |
(j + 1,
|
blanchet@38699
|
706 |
if is_theorem_bad_for_atps full_types th andalso
|
blanchet@38699
|
707 |
not (member Thm.eq_thm add_thms th) then
|
blanchet@38699
|
708 |
rest
|
blanchet@38699
|
709 |
else
|
blanchet@38752
|
710 |
(((fn () =>
|
blanchet@38752
|
711 |
if name0 = "" then
|
blanchet@38752
|
712 |
th |> backquotify
|
blanchet@38752
|
713 |
else
|
blanchet@38752
|
714 |
let
|
blanchet@38752
|
715 |
val name1 = Facts.extern facts name0
|
blanchet@38752
|
716 |
val name2 = Name_Space.extern full_space name0
|
blanchet@38752
|
717 |
in
|
blanchet@38752
|
718 |
case find_first check_thms [name1, name2, name0] of
|
blanchet@38752
|
719 |
SOME name => repair_name reserved multi j name
|
blanchet@38752
|
720 |
| NONE => ""
|
blanchet@38937
|
721 |
end),
|
blanchet@38937
|
722 |
let val t = prop_of th in
|
blanchet@38937
|
723 |
if is_chained th then Chained
|
blanchet@38937
|
724 |
else if not global then Local
|
blanchet@38937
|
725 |
else if Termtab.defined intros t then Intro
|
blanchet@38937
|
726 |
else if Termtab.defined elims t then Elim
|
blanchet@38937
|
727 |
else if Termtab.defined simps t then Simp
|
blanchet@38937
|
728 |
else General
|
blanchet@38937
|
729 |
end),
|
blanchet@38752
|
730 |
(multi, th)) :: rest)) ths
|
blanchet@38699
|
731 |
#> snd
|
blanchet@38627
|
732 |
end)
|
blanchet@38644
|
733 |
in
|
blanchet@38752
|
734 |
[] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
|
blanchet@38752
|
735 |
|> add_facts true Facts.fold_static global_facts global_facts
|
blanchet@38644
|
736 |
end
|
blanchet@38627
|
737 |
|
blanchet@38627
|
738 |
(* The single-name theorems go after the multiple-name ones, so that single
|
blanchet@38627
|
739 |
names are preferred when both are available. *)
|
blanchet@38699
|
740 |
fun name_thm_pairs ctxt respect_no_atp =
|
blanchet@38744
|
741 |
List.partition (fst o snd) #> op @ #> map (apsnd snd)
|
blanchet@38699
|
742 |
#> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
|
blanchet@38627
|
743 |
|
blanchet@38627
|
744 |
(***************************************************************)
|
blanchet@38627
|
745 |
(* ATP invocation methods setup *)
|
blanchet@38627
|
746 |
(***************************************************************)
|
blanchet@38627
|
747 |
|
blanchet@38891
|
748 |
fun relevant_facts ctxt full_types (threshold0, threshold1) max_relevant
|
blanchet@38744
|
749 |
theory_relevant (relevance_override as {add, del, only})
|
blanchet@38891
|
750 |
chained_ths hyp_ts concl_t =
|
blanchet@37538
|
751 |
let
|
blanchet@38822
|
752 |
val decay = Math.pow ((1.0 - threshold1) / (1.0 - threshold0),
|
blanchet@38822
|
753 |
1.0 / Real.fromInt (max_relevant + 1))
|
blanchet@37538
|
754 |
val add_thms = maps (ProofContext.get_fact ctxt) add
|
blanchet@38696
|
755 |
val reserved = reserved_isar_keyword_table ()
|
blanchet@37538
|
756 |
val axioms =
|
blanchet@38699
|
757 |
(if only then
|
blanchet@38752
|
758 |
maps (map (fn ((name, loc), th) => ((K name, loc), (true, th)))
|
blanchet@38752
|
759 |
o name_thm_pairs_from_ref ctxt reserved chained_ths) add
|
blanchet@38699
|
760 |
else
|
blanchet@38699
|
761 |
all_name_thms_pairs ctxt reserved full_types add_thms chained_ths)
|
blanchet@38688
|
762 |
|> name_thm_pairs ctxt (respect_no_atp andalso not only)
|
blanchet@38937
|
763 |
|> uniquify
|
blanchet@37538
|
764 |
in
|
blanchet@38688
|
765 |
trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
|
blanchet@38688
|
766 |
" theorems");
|
blanchet@38745
|
767 |
(if threshold0 > 1.0 orelse threshold0 > threshold1 then
|
blanchet@38739
|
768 |
[]
|
blanchet@38745
|
769 |
else if threshold0 < 0.0 then
|
blanchet@38739
|
770 |
axioms
|
blanchet@38739
|
771 |
else
|
blanchet@38745
|
772 |
relevance_filter ctxt threshold0 decay max_relevant theory_relevant
|
blanchet@38745
|
773 |
relevance_override axioms (concl_t :: hyp_ts))
|
blanchet@38822
|
774 |
|> map (apfst (apfst (fn f => f ())))
|
blanchet@37538
|
775 |
end
|
immler@30536
|
776 |
|
paulson@15347
|
777 |
end;
|