author | blanchet |
Tue, 31 Aug 2010 23:46:23 +0200 | |
changeset 38986 | e34c1b09bb5e |
parent 38985 | src/HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML@162bbbea4e4d |
child 38988 | 483879af0643 |
permissions | -rw-r--r-- |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_fact_minimize.ML |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
2 |
Author: Philipp Meyer, TU Muenchen |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
4 |
|
35867 | 5 |
Minimization of theorem list for Metis using automatic theorem provers. |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
6 |
*) |
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
7 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
8 |
signature SLEDGEHAMMER_FACT_MINIMIZE = |
32525 | 9 |
sig |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
10 |
type locality = Sledgehammer_Fact_Filter.locality |
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38015
diff
changeset
|
11 |
type params = Sledgehammer.params |
35867 | 12 |
|
38050 | 13 |
val minimize_theorems : |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
14 |
params -> int -> int -> Proof.state -> ((string * locality) * thm list) list |
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
15 |
-> ((string * locality) * thm list) list option * string |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
16 |
val run_minimize : params -> int -> Facts.ref list -> Proof.state -> unit |
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
17 |
end; |
32525 | 18 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
19 |
structure Sledgehammer_Fact_Minimize : SLEDGEHAMMER_FACT_MINIMIZE = |
31037
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
20 |
struct |
ac8669134e7a
added Philipp Meyer's implementation of AtpMinimal
immler@in.tum.de
parents:
diff
changeset
|
21 |
|
38586 | 22 |
open ATP_Systems |
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36063
diff
changeset
|
23 |
open Sledgehammer_Util |
38045 | 24 |
open Sledgehammer_Fact_Filter |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
35969
diff
changeset
|
25 |
open Sledgehammer_Proof_Reconstruct |
38021
e024504943d1
rename "ATP_Manager" ML module to "Sledgehammer";
blanchet
parents:
38015
diff
changeset
|
26 |
open Sledgehammer |
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
27 |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
28 |
(* wrapper for calling external prover *) |
31236 | 29 |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
30 |
fun string_for_failure Unprovable = "Unprovable." |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
31 |
| string_for_failure TimedOut = "Timed out." |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
32 |
| string_for_failure _ = "Unknown error." |
31236 | 33 |
|
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
34 |
fun n_theorems names = |
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
35 |
let val n = length names in |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
36 |
string_of_int n ^ " theorem" ^ plural_s n ^ |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
37 |
(if n > 0 then |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
38 |
": " ^ (names |> map fst |
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
39 |
|> sort_distinct string_ord |> space_implode " ") |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
40 |
else |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
41 |
"") |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
42 |
end |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
43 |
|
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
44 |
fun test_theorems ({debug, verbose, overlord, atps, full_types, isar_proof, |
38741 | 45 |
isar_shrink_factor, ...} : params) |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
46 |
(prover : prover) explicit_apply timeout subgoal state |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
47 |
axioms = |
31236 | 48 |
let |
38015 | 49 |
val _ = |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
50 |
priority ("Testing " ^ n_theorems (map fst axioms) ^ "...") |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
51 |
val params = |
38982
820b8221ed48
added "blocking" option to Sledgehammer to run in synchronous mode;
blanchet
parents:
38752
diff
changeset
|
52 |
{blocking = true, debug = debug, verbose = verbose, overlord = overlord, |
820b8221ed48
added "blocking" option to Sledgehammer to run in synchronous mode;
blanchet
parents:
38752
diff
changeset
|
53 |
atps = atps, full_types = full_types, explicit_apply = explicit_apply, |
38745
ad577fd62ee4
reorganize options regarding to the relevance threshold and decay
blanchet
parents:
38744
diff
changeset
|
54 |
relevance_thresholds = (1.01, 1.01), max_relevant = NONE, |
38741 | 55 |
theory_relevant = NONE, isar_proof = isar_proof, |
38985
162bbbea4e4d
added "expect" feature of Nitpick to Sledgehammer, for regression testing
blanchet
parents:
38982
diff
changeset
|
56 |
isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""} |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
57 |
val axioms = maps (fn (n, ths) => map (pair n) ths) axioms |
36263
56bf63d70120
use "Proof.goal" in Sledgehammer's minimizer (just like everywhere else in Sledgehammer), not "Proof.raw_goal"
blanchet
parents:
36232
diff
changeset
|
58 |
val {context = ctxt, facts, goal} = Proof.goal state |
32941
72d48e333b77
eliminated extraneous wrapping of public records;
wenzelm
parents:
32937
diff
changeset
|
59 |
val problem = |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
35969
diff
changeset
|
60 |
{subgoal = subgoal, goal = (ctxt, (facts, goal)), |
35969 | 61 |
relevance_override = {add = [], del = [], only = false}, |
38084
e2aac207d13b
"axiom_clauses" -> "axioms" (these are no longer clauses)
blanchet
parents:
38083
diff
changeset
|
62 |
axioms = SOME axioms} |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
63 |
val result as {outcome, used_thm_names, ...} = prover params (K "") problem |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36143
diff
changeset
|
64 |
in |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
65 |
priority (case outcome of |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
66 |
NONE => |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
67 |
if length used_thm_names = length axioms then |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
68 |
"Found proof." |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
69 |
else |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
70 |
"Found proof with " ^ n_theorems used_thm_names ^ "." |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
71 |
| SOME failure => string_for_failure failure); |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
72 |
result |
36223
217ca1273786
make Sledgehammer's minimizer also minimize Isar proofs
blanchet
parents:
36143
diff
changeset
|
73 |
end |
31236 | 74 |
|
75 |
(* minimalization of thms *) |
|
76 |
||
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
77 |
fun filter_used_facts used = filter (member (op =) used o fst) |
38015 | 78 |
|
79 |
fun sublinear_minimize _ [] p = p |
|
80 |
| sublinear_minimize test (x :: xs) (seen, result) = |
|
81 |
case test (xs @ seen) of |
|
38488 | 82 |
result as {outcome = NONE, proof, used_thm_names, ...} : prover_result => |
38015 | 83 |
sublinear_minimize test (filter_used_facts used_thm_names xs) |
84 |
(filter_used_facts used_thm_names seen, result) |
|
85 |
| _ => sublinear_minimize test xs (x :: seen, result) |
|
86 |
||
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
87 |
(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
88 |
preprocessing time is included in the estimate below but isn't part of the |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
89 |
timeout. *) |
38590
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents:
38589
diff
changeset
|
90 |
val fudge_msecs = 1000 |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
91 |
|
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
92 |
fun minimize_theorems {atps = [], ...} _ _ _ _ = error "No ATP is set." |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
93 |
| minimize_theorems (params as {debug, atps = atp :: _, full_types, |
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
94 |
isar_proof, isar_shrink_factor, timeout, ...}) |
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
95 |
i n state axioms = |
31236 | 96 |
let |
36378 | 97 |
val thy = Proof.theory_of state |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
98 |
val prover = get_prover_fun thy atp |
38590
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents:
38589
diff
changeset
|
99 |
val msecs = Time.toMilliseconds timeout |
bd443b426d56
get rid of "minimize_timeout", now that there's an automatic adaptive timeout mechanism in "minimize"
blanchet
parents:
38589
diff
changeset
|
100 |
val _ = priority ("Sledgehammer minimize: ATP " ^ quote atp ^ ".") |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
101 |
val {context = ctxt, goal, ...} = Proof.goal state |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
102 |
val (_, hyp_ts, concl_t) = strip_subgoal goal i |
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
103 |
val explicit_apply = |
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
104 |
not (forall (Meson.is_fol_term thy) |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
105 |
(concl_t :: hyp_ts @ maps (map prop_of o snd) axioms)) |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
106 |
fun do_test timeout = |
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
107 |
test_theorems params prover explicit_apply timeout i state |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
108 |
val timer = Timer.startRealTimer () |
31236 | 109 |
in |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
110 |
(case do_test timeout axioms of |
38083
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents:
38050
diff
changeset
|
111 |
result as {outcome = NONE, pool, used_thm_names, |
c4b57f68ddb3
remove the "extra_clauses" business introduced in 19a5f1c8a844;
blanchet
parents:
38050
diff
changeset
|
112 |
conjecture_shape, ...} => |
38015 | 113 |
let |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
114 |
val time = Timer.checkRealTimer timer |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
115 |
val new_timeout = |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
116 |
Int.min (msecs, Time.toMilliseconds time + fudge_msecs) |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38084
diff
changeset
|
117 |
|> Time.fromMilliseconds |
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
118 |
val (min_thms, {proof, axiom_names, ...}) = |
38100
e458a0dd3dc1
use "explicit_apply" in the minimizer whenever it might make a difference to prevent freak failures;
blanchet
parents:
38094
diff
changeset
|
119 |
sublinear_minimize (do_test new_timeout) |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
120 |
(filter_used_facts used_thm_names axioms) ([], result) |
38094 | 121 |
val n = length min_thms |
38015 | 122 |
val _ = priority (cat_lines |
38094 | 123 |
["Minimized: " ^ string_of_int n ^ " theorem" ^ plural_s n] ^ |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
124 |
(case length (filter (curry (op =) Chained o snd o fst) min_thms) of |
38698
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
125 |
0 => "" |
d19c3a7ce38b
clean handling of whether a fact is chained or not;
blanchet
parents:
38696
diff
changeset
|
126 |
| n => " (including " ^ Int.toString n ^ " chained)") ^ ".") |
38015 | 127 |
in |
128 |
(SOME min_thms, |
|
129 |
proof_text isar_proof |
|
130 |
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape) |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
131 |
(full_types, K "", proof, axiom_names, goal, i) |> fst) |
38015 | 132 |
end |
133 |
| {outcome = SOME TimedOut, ...} => |
|
134 |
(NONE, "Timeout: You can increase the time limit using the \"timeout\" \ |
|
135 |
\option (e.g., \"timeout = " ^ |
|
136 |
string_of_int (10 + msecs div 1000) ^ " s\").") |
|
137 |
| {outcome = SOME UnknownError, ...} => |
|
138 |
(* Failure sometimes mean timeout, unfortunately. *) |
|
139 |
(NONE, "Failure: No proof was found with the current time limit. You \ |
|
140 |
\can increase the time limit using the \"timeout\" \ |
|
141 |
\option (e.g., \"timeout = " ^ |
|
142 |
string_of_int (10 + msecs div 1000) ^ " s\").") |
|
143 |
| {message, ...} => (NONE, "ATP error: " ^ message)) |
|
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37926
diff
changeset
|
144 |
handle ERROR msg => (NONE, "Error: " ^ msg) |
31236 | 145 |
end |
146 |
||
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38100
diff
changeset
|
147 |
fun run_minimize params i refs state = |
38045 | 148 |
let |
149 |
val ctxt = Proof.context_of state |
|
38696
4c6b65d6a135
quote facts whose names collide with a keyword or command name (cf. "subclass" in "Jinja/J/TypeSafe.thy")
blanchet
parents:
38617
diff
changeset
|
150 |
val reserved = reserved_isar_keyword_table () |
38045 | 151 |
val chained_ths = #facts (Proof.goal state) |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
152 |
val axioms = |
38752
6628adcae4a7
consider "locality" when assigning weights to facts
blanchet
parents:
38745
diff
changeset
|
153 |
maps (map (apsnd single) |
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
154 |
o name_thm_pairs_from_ref ctxt reserved chained_ths) refs |
38045 | 155 |
in |
156 |
case subgoal_count state of |
|
157 |
0 => priority "No subgoal!" |
|
158 |
| n => |
|
38744
2b6333f78a9e
make relevance filter work in term of a "max_relevant" option + use Vampire SOS;
blanchet
parents:
38741
diff
changeset
|
159 |
(kill_atps (); priority (#2 (minimize_theorems params i n state axioms))) |
38045 | 160 |
end |
161 |
||
35866
513074557e06
move the Sledgehammer Isar commands together into one file;
blanchet
parents:
35865
diff
changeset
|
162 |
end; |