author | blanchet |
Mon, 26 Jul 2010 14:14:24 +0200 | |
changeset 37994 | b04307085a09 |
parent 37926 | e6ff246c0cdb |
child 37996 | 11c076ea92e9 |
permissions | -rw-r--r-- |
32327
0971cc0b6a57
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
wenzelm
parents:
31793
diff
changeset
|
1 |
(* Title: HOL/Tools/ATP_Manager/atp_manager.ML |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
2 |
Author: Fabian Immler, TU Muenchen |
32996
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
wenzelm
parents:
32995
diff
changeset
|
3 |
Author: Makarius |
35969 | 4 |
Author: Jasmin Blanchette, TU Muenchen |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
5 |
|
32996
d2e48879e65a
removed disjunctive group cancellation -- provers run independently;
wenzelm
parents:
32995
diff
changeset
|
6 |
Central manager component for ATP threads. |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
7 |
*) |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
8 |
|
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
9 |
signature ATP_MANAGER = |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
10 |
sig |
35969 | 11 |
type relevance_override = Sledgehammer_Fact_Filter.relevance_override |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
12 |
type minimize_command = Sledgehammer_Proof_Reconstruct.minimize_command |
35969 | 13 |
type params = |
14 |
{debug: bool, |
|
15 |
verbose: bool, |
|
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36064
diff
changeset
|
16 |
overlord: bool, |
35969 | 17 |
atps: string list, |
18 |
full_types: bool, |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36231
diff
changeset
|
19 |
explicit_apply: bool, |
35969 | 20 |
relevance_threshold: real, |
36922 | 21 |
relevance_convergence: real, |
36220
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents:
36184
diff
changeset
|
22 |
theory_relevant: bool option, |
36922 | 23 |
defs_relevant: bool, |
35969 | 24 |
isar_proof: bool, |
36924 | 25 |
isar_shrink_factor: int, |
35969 | 26 |
timeout: Time.time, |
27 |
minimize_timeout: Time.time} |
|
35867 | 28 |
type problem = |
35969 | 29 |
{subgoal: int, |
30 |
goal: Proof.context * (thm list * thm), |
|
31 |
relevance_override: relevance_override, |
|
37621 | 32 |
axiom_clauses: ((string * int) * thm) list option, |
33 |
filtered_clauses: ((string * int) * thm) list option} |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
34 |
datatype failure = |
37627 | 35 |
Unprovable | IncompleteUnprovable | CantConnect | TimedOut | |
36 |
OutOfResources | OldSpass | MalformedInput | MalformedOutput | UnknownError |
|
35867 | 37 |
type prover_result = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
38 |
{outcome: failure option, |
35969 | 39 |
message: string, |
37926
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
blanchet
parents:
37627
diff
changeset
|
40 |
pool: string Symtab.table, |
35969 | 41 |
relevant_thm_names: string list, |
42 |
atp_run_time_in_msecs: int, |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
43 |
output: string, |
35969 | 44 |
proof: string, |
45 |
internal_thm_names: string Vector.vector, |
|
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
46 |
conjecture_shape: int list list, |
37621 | 47 |
filtered_clauses: ((string * int) * thm) list} |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
48 |
type prover = |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
49 |
params -> minimize_command -> Time.time -> problem -> prover_result |
35867 | 50 |
|
35969 | 51 |
val kill_atps: unit -> unit |
52 |
val running_atps: unit -> unit |
|
29112
f2b45eea6dac
added 'atp_messages' command, which displays recent messages synchronously;
wenzelm
parents:
28835
diff
changeset
|
53 |
val messages: int option -> unit |
35867 | 54 |
val add_prover: string * prover -> theory -> theory |
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
55 |
val get_prover: theory -> string -> prover |
35969 | 56 |
val available_atps: theory -> unit |
37584 | 57 |
val start_prover_thread : |
58 |
params -> int -> int -> relevance_override -> (string -> minimize_command) |
|
59 |
-> Proof.state -> string -> unit |
|
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
60 |
end; |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
61 |
|
35865 | 62 |
structure ATP_Manager : ATP_MANAGER = |
28477
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
63 |
struct |
9339d4dcec8b
version of sledgehammer using threads instead of processes, misc cleanup;
wenzelm
parents:
diff
changeset
|
64 |
|
37578
9367cb36b1c4
renamed "Sledgehammer_FOL_Clauses" to "Metis_Clauses", so that Metis doesn't depend on Sledgehammer
blanchet
parents:
37577
diff
changeset
|
65 |
open Metis_Clauses |
36063
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36059
diff
changeset
|
66 |
open Sledgehammer_Fact_Filter |
cdc6855a6387
make Sledgehammer output "by" vs. "apply", "qed" vs. "next", and any necessary "prefer"
blanchet
parents:
36059
diff
changeset
|
67 |
open Sledgehammer_Proof_Reconstruct |
37583
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
68 |
|
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
69 |
(** The Sledgehammer **) |
9ce2451647d5
factored non-ATP specific code from "ATP_Manager" out, so that it can be reused for the LEO-II integration
blanchet
parents:
37581
diff
changeset
|
70 |
|
37585 | 71 |
val das_Tool = "Sledgehammer" |
72 |
||
73 |
fun kill_atps () = Async_Manager.kill_threads das_Tool "ATPs" |
|
74 |
fun running_atps () = Async_Manager.running_threads das_Tool "ATPs" |
|
75 |
val messages = Async_Manager.thread_messages das_Tool "ATP" |
|
35969 | 76 |
|
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
77 |
(** problems, results, provers, etc. **) |
35969 | 78 |
|
79 |
type params = |
|
80 |
{debug: bool, |
|
81 |
verbose: bool, |
|
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36064
diff
changeset
|
82 |
overlord: bool, |
35969 | 83 |
atps: string list, |
84 |
full_types: bool, |
|
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36231
diff
changeset
|
85 |
explicit_apply: bool, |
35969 | 86 |
relevance_threshold: real, |
36922 | 87 |
relevance_convergence: real, |
36220
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents:
36184
diff
changeset
|
88 |
theory_relevant: bool option, |
36922 | 89 |
defs_relevant: bool, |
35969 | 90 |
isar_proof: bool, |
36924 | 91 |
isar_shrink_factor: int, |
35969 | 92 |
timeout: Time.time, |
93 |
minimize_timeout: Time.time} |
|
35867 | 94 |
|
95 |
type problem = |
|
35969 | 96 |
{subgoal: int, |
97 |
goal: Proof.context * (thm list * thm), |
|
98 |
relevance_override: relevance_override, |
|
37621 | 99 |
axiom_clauses: ((string * int) * thm) list option, |
100 |
filtered_clauses: ((string * int) * thm) list option} |
|
35867 | 101 |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
102 |
datatype failure = |
37627 | 103 |
Unprovable | IncompleteUnprovable | CantConnect | TimedOut | OutOfResources | |
104 |
OldSpass | MalformedInput | MalformedOutput | UnknownError |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
105 |
|
35867 | 106 |
type prover_result = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
107 |
{outcome: failure option, |
35969 | 108 |
message: string, |
37926
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
blanchet
parents:
37627
diff
changeset
|
109 |
pool: string Symtab.table, |
35969 | 110 |
relevant_thm_names: string list, |
111 |
atp_run_time_in_msecs: int, |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
112 |
output: string, |
35969 | 113 |
proof: string, |
114 |
internal_thm_names: string Vector.vector, |
|
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
115 |
conjecture_shape: int list list, |
37621 | 116 |
filtered_clauses: ((string * int) * thm) list} |
35867 | 117 |
|
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
118 |
type prover = |
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36235
diff
changeset
|
119 |
params -> minimize_command -> Time.time -> problem -> prover_result |
35867 | 120 |
|
28582 | 121 |
(* named provers *) |
28484 | 122 |
|
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
123 |
structure Data = Theory_Data |
28582 | 124 |
( |
35867 | 125 |
type T = (prover * stamp) Symtab.table; |
32938 | 126 |
val empty = Symtab.empty; |
127 |
val extend = I; |
|
33522 | 128 |
fun merge data : T = Symtab.merge (eq_snd op =) data |
37580
c2c1caff5dea
got rid of "respect_no_atp" option, which even I don't use
blanchet
parents:
37578
diff
changeset
|
129 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") |
28582 | 130 |
); |
28484 | 131 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32824
diff
changeset
|
132 |
fun add_prover (name, prover) thy = |
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
133 |
Data.map (Symtab.update_new (name, (prover, stamp ()))) thy |
37580
c2c1caff5dea
got rid of "respect_no_atp" option, which even I don't use
blanchet
parents:
37578
diff
changeset
|
134 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") |
28582 | 135 |
|
32995 | 136 |
fun get_prover thy name = |
37580
c2c1caff5dea
got rid of "respect_no_atp" option, which even I don't use
blanchet
parents:
37578
diff
changeset
|
137 |
the (Symtab.lookup (Data.get thy) name) |> fst |
c2c1caff5dea
got rid of "respect_no_atp" option, which even I don't use
blanchet
parents:
37578
diff
changeset
|
138 |
handle Option.Option => error ("Unknown ATP: " ^ name ^ ".") |
32995 | 139 |
|
36229
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
blanchet
parents:
36226
diff
changeset
|
140 |
fun available_atps thy = |
c95fab3f9cc5
workaround for Proof General's off-by-a-few sendback display bug, whereby "pr" in "proof" is not highlighted
blanchet
parents:
36226
diff
changeset
|
141 |
priority ("Available ATPs: " ^ |
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
142 |
commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".") |
28571 | 143 |
|
28586
d238b83ba3fc
renamed kill_all to kill, in conformance with atp_kill command;
wenzelm
parents:
28582
diff
changeset
|
144 |
|
28595
67e3945b53f1
add_prover: plain prover function, without thread;
wenzelm
parents:
28589
diff
changeset
|
145 |
(* start prover thread *) |
28484 | 146 |
|
37584 | 147 |
fun start_prover_thread (params as {verbose, full_types, timeout, ...}) i n |
148 |
relevance_override minimize_command proof_state name = |
|
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
149 |
let |
37584 | 150 |
val birth_time = Time.now () |
151 |
val death_time = Time.+ (birth_time, timeout) |
|
36379
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
152 |
val prover = get_prover (Proof.theory_of proof_state) name |
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
153 |
val {context = ctxt, facts, goal} = Proof.goal proof_state; |
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
154 |
val desc = |
20ef039bccff
make "ATP_Manager.get_prover" a total function, since we always want to show the same error text
blanchet
parents:
36373
diff
changeset
|
155 |
"ATP " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^ |
36392 | 156 |
Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)); |
37584 | 157 |
in |
37585 | 158 |
Async_Manager.launch das_Tool verbose birth_time death_time desc |
37584 | 159 |
(fn () => |
160 |
let |
|
161 |
val problem = |
|
162 |
{subgoal = i, goal = (ctxt, (facts, goal)), |
|
163 |
relevance_override = relevance_override, axiom_clauses = NONE, |
|
164 |
filtered_clauses = NONE} |
|
165 |
in |
|
166 |
prover params (minimize_command name) timeout problem |> #message |
|
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37926
diff
changeset
|
167 |
handle ERROR message => "Error: " ^ message ^ "\n" |
37584 | 168 |
end) |
169 |
end |
|
28582 | 170 |
|
171 |
end; |