12 type failure = ATP_Proof.failure |
12 type failure = ATP_Proof.failure |
13 |
13 |
14 type atp_config = |
14 type atp_config = |
15 {exec : string * string, |
15 {exec : string * string, |
16 required_execs : (string * string) list, |
16 required_execs : (string * string) list, |
17 arguments : int -> Time.time -> (unit -> (string * real) list) -> string, |
17 arguments : |
|
18 Proof.context -> int -> Time.time -> (unit -> (string * real) list) |
|
19 -> string, |
18 proof_delims : (string * string) list, |
20 proof_delims : (string * string) list, |
19 known_failures : (failure * string) list, |
21 known_failures : (failure * string) list, |
20 hypothesis_kind : formula_kind, |
22 hypothesis_kind : formula_kind, |
21 formats : format list, |
23 formats : format list, |
22 best_slices : unit -> (real * (bool * int)) list, |
24 best_slices : Proof.context -> (real * (bool * int)) list, |
23 best_type_systems : string list} |
25 best_type_systems : string list} |
24 |
26 |
25 datatype e_weight_method = |
27 val e_weight_method : string Config.T |
26 E_Slices | E_Auto | E_Fun_Weight | E_Sym_Offset_Weight |
28 val e_default_fun_weight : real Config.T |
27 |
29 val e_fun_weight_base : real Config.T |
28 (* for experimentation purposes -- do not use in production code *) |
30 val e_fun_weight_span : real Config.T |
29 val e_weight_method : e_weight_method Unsynchronized.ref |
31 val e_default_sym_offs_weight : real Config.T |
30 val e_default_fun_weight : real Unsynchronized.ref |
32 val e_sym_offs_weight_base : real Config.T |
31 val e_fun_weight_base : real Unsynchronized.ref |
33 val e_sym_offs_weight_span : real Config.T |
32 val e_fun_weight_span : real Unsynchronized.ref |
|
33 val e_default_sym_offs_weight : real Unsynchronized.ref |
|
34 val e_sym_offs_weight_base : real Unsynchronized.ref |
|
35 val e_sym_offs_weight_span : real Unsynchronized.ref |
|
36 (* end *) |
|
37 val eN : string |
34 val eN : string |
38 val spassN : string |
35 val spassN : string |
39 val vampireN : string |
36 val vampireN : string |
40 val tofof_eN : string |
37 val tofof_eN : string |
41 val sine_eN : string |
38 val sine_eN : string |
42 val snarkN : string |
39 val snarkN : string |
43 val z3_atpN : string |
40 val z3_atpN : string |
44 val remote_prefix : string |
41 val remote_prefix : string |
45 val remote_atp : |
42 val remote_atp : |
46 string -> string -> string list -> (string * string) list |
43 string -> string -> string list -> (string * string) list |
47 -> (failure * string) list -> formula_kind -> format list -> (unit -> int) |
44 -> (failure * string) list -> formula_kind -> format list |
48 -> string list -> string * atp_config |
45 -> (Proof.context -> int) -> string list -> string * atp_config |
49 val add_atp : string * atp_config -> theory -> theory |
46 val add_atp : string * atp_config -> theory -> theory |
50 val get_atp : theory -> string -> atp_config |
47 val get_atp : theory -> string -> atp_config |
51 val supported_atps : theory -> string list |
48 val supported_atps : theory -> string list |
52 val is_atp_installed : theory -> string -> bool |
49 val is_atp_installed : theory -> string -> bool |
53 val refresh_systems_on_tptp : unit -> unit |
50 val refresh_systems_on_tptp : unit -> unit |
63 (* ATP configuration *) |
60 (* ATP configuration *) |
64 |
61 |
65 type atp_config = |
62 type atp_config = |
66 {exec : string * string, |
63 {exec : string * string, |
67 required_execs : (string * string) list, |
64 required_execs : (string * string) list, |
68 arguments : int -> Time.time -> (unit -> (string * real) list) -> string, |
65 arguments : |
|
66 Proof.context -> int -> Time.time -> (unit -> (string * real) list) |
|
67 -> string, |
69 proof_delims : (string * string) list, |
68 proof_delims : (string * string) list, |
70 known_failures : (failure * string) list, |
69 known_failures : (failure * string) list, |
71 hypothesis_kind : formula_kind, |
70 hypothesis_kind : formula_kind, |
72 formats : format list, |
71 formats : format list, |
73 best_slices : unit -> (real * (bool * int)) list, |
72 best_slices : Proof.context -> (real * (bool * int)) list, |
74 best_type_systems : string list} |
73 best_type_systems : string list} |
75 |
74 |
76 val known_perl_failures = |
75 val known_perl_failures = |
77 [(CantConnect, "HTTP error"), |
76 [(CantConnect, "HTTP error"), |
78 (NoPerl, "env: perl"), |
77 (NoPerl, "env: perl"), |
111 if string_ord (getenv "E_VERSION", "1.1") = LESS then 2000 else 1000 |
110 if string_ord (getenv "E_VERSION", "1.1") = LESS then 2000 else 1000 |
112 |
111 |
113 val tstp_proof_delims = |
112 val tstp_proof_delims = |
114 ("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation") |
113 ("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation") |
115 |
114 |
116 datatype e_weight_method = |
115 val e_slicesN = "slices" |
117 E_Slices | E_Auto | E_Fun_Weight | E_Sym_Offset_Weight |
116 val e_autoN = "auto" |
118 |
117 val e_fun_weightN = "fun_weight" |
119 val e_weight_method = Unsynchronized.ref E_Slices |
118 val e_sym_offset_weightN = "sym_offset_weight" |
|
119 |
|
120 val e_weight_method = |
|
121 Attrib.setup_config_string @{binding atp_e_weight_method} (K e_slicesN) |
120 (* FUDGE *) |
122 (* FUDGE *) |
121 val e_default_fun_weight = Unsynchronized.ref 20.0 |
123 val e_default_fun_weight = |
122 val e_fun_weight_base = Unsynchronized.ref 0.0 |
124 Attrib.setup_config_real @{binding atp_e_default_fun_weight} (K 20.0) |
123 val e_fun_weight_span = Unsynchronized.ref 40.0 |
125 val e_fun_weight_base = |
124 val e_default_sym_offs_weight = Unsynchronized.ref 1.0 |
126 Attrib.setup_config_real @{binding atp_e_fun_weight_base} (K 0.0) |
125 val e_sym_offs_weight_base = Unsynchronized.ref ~20.0 |
127 val e_fun_weight_span = |
126 val e_sym_offs_weight_span = Unsynchronized.ref 60.0 |
128 Attrib.setup_config_real @{binding atp_e_fun_weight_span} (K 40.0) |
|
129 val e_default_sym_offs_weight = |
|
130 Attrib.setup_config_real @{binding atp_e_default_sym_offs_weight} (K 1.0) |
|
131 val e_sym_offs_weight_base = |
|
132 Attrib.setup_config_real @{binding atp_e_sym_offs_weight_base} (K ~20.0) |
|
133 val e_sym_offs_weight_span = |
|
134 Attrib.setup_config_real @{binding atp_e_sym_offs_weight_span} (K 60.0) |
127 |
135 |
128 fun e_weight_method_case method fw sow = |
136 fun e_weight_method_case method fw sow = |
129 case method of |
137 if method = e_fun_weightN then fw |
130 E_Auto => raise Fail "Unexpected \"E_Auto\"" |
138 else if method = e_sym_offset_weightN then sow |
131 | E_Fun_Weight => fw |
139 else raise Fail ("unexpected" ^ quote method) |
132 | E_Sym_Offset_Weight => sow |
140 |
133 |
141 fun scaled_e_weight ctxt method w = |
134 fun scaled_e_weight method w = |
142 w * Config.get ctxt |
135 w * e_weight_method_case method (!e_fun_weight_span) (!e_sym_offs_weight_span) |
143 (e_weight_method_case method e_fun_weight_span e_sym_offs_weight_span) |
136 + e_weight_method_case method (!e_fun_weight_base) (!e_sym_offs_weight_base) |
144 + Config.get ctxt |
|
145 (e_weight_method_case method e_fun_weight_base e_sym_offs_weight_base) |
137 |> Real.ceil |> signed_string_of_int |
146 |> Real.ceil |> signed_string_of_int |
138 |
147 |
139 fun e_weight_arguments method weights = |
148 fun e_weight_arguments ctxt method weights = |
140 if method = E_Auto then |
149 if method = e_autoN then |
141 "-xAutoDev" |
150 "-xAutoDev" |
142 else |
151 else |
143 "--split-clauses=4 --split-reuse-defs --simul-paramod --forward-context-sr \ |
152 "--split-clauses=4 --split-reuse-defs --simul-paramod --forward-context-sr \ |
144 \--destructive-er-aggressive --destructive-er --presat-simplify \ |
153 \--destructive-er-aggressive --destructive-er --presat-simplify \ |
145 \--prefer-initial-clauses -tKBO6 -winvfreqrank -c1 -Ginvfreqconjmax -F1 \ |
154 \--prefer-initial-clauses -tKBO6 -winvfreqrank -c1 -Ginvfreqconjmax -F1 \ |
146 \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred \ |
155 \--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred \ |
147 \-H'(4*" ^ e_weight_method_case method "FunWeight" "SymOffsetWeight" ^ |
156 \-H'(4*" ^ e_weight_method_case method "FunWeight" "SymOffsetWeight" ^ |
148 "(SimulateSOS, " ^ |
157 "(SimulateSOS, " ^ |
149 (e_weight_method_case method (!e_default_fun_weight) |
158 (e_weight_method_case method e_default_fun_weight e_default_sym_offs_weight |
150 (!e_default_sym_offs_weight) |
159 |> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^ |
151 |> Real.ceil |> signed_string_of_int) ^ |
|
152 ",20,1.5,1.5,1" ^ |
160 ",20,1.5,1.5,1" ^ |
153 (weights () |> map (fn (s, w) => "," ^ s ^ ":" ^ scaled_e_weight method w) |
161 (weights () |
154 |> implode) ^ |
162 |> map (fn (s, w) => "," ^ s ^ ":" ^ scaled_e_weight ctxt method w) |
|
163 |> implode) ^ |
155 "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\ |
164 "),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\ |
156 \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\ |
165 \1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\ |
157 \FIFOWeight(PreferProcessed))'" |
166 \FIFOWeight(PreferProcessed))'" |
158 |
167 |
159 fun is_old_e_version () = (string_ord (getenv "E_VERSION", "1.2w") = LESS) |
168 fun is_old_e_version () = (string_ord (getenv "E_VERSION", "1.2w") = LESS) |
160 |
169 |
161 fun effective_e_weight_method () = |
170 fun effective_e_weight_method ctxt = |
162 if is_old_e_version () then E_Auto else !e_weight_method |
171 if is_old_e_version () then e_autoN else Config.get ctxt e_weight_method |
163 |
172 |
164 (* The order here must correspond to the order in "e_config" below. *) |
173 (* The order here must correspond to the order in "e_config" below. *) |
165 fun method_for_slice slice = |
174 fun method_for_slice ctxt slice = |
166 case effective_e_weight_method () of |
175 let val method = effective_e_weight_method ctxt in |
167 E_Slices => (case slice of |
176 if method = e_slicesN then |
168 0 => E_Sym_Offset_Weight |
177 case slice of |
169 | 1 => E_Auto |
178 0 => e_sym_offset_weightN |
170 | 2 => E_Fun_Weight |
179 | 1 => e_autoN |
171 | _ => raise Fail "no such slice") |
180 | 2 => e_fun_weightN |
172 | method => method |
181 | _ => raise Fail "no such slice" |
|
182 else |
|
183 method |
|
184 end |
173 |
185 |
174 val e_config : atp_config = |
186 val e_config : atp_config = |
175 {exec = ("E_HOME", "eproof"), |
187 {exec = ("E_HOME", "eproof"), |
176 required_execs = [], |
188 required_execs = [], |
177 arguments = fn slice => fn timeout => fn weights => |
189 arguments = fn ctxt => fn slice => fn timeout => fn weights => |
178 "--tstp-in --tstp-out -l5 " ^ |
190 "--tstp-in --tstp-out -l5 " ^ |
179 e_weight_arguments (method_for_slice slice) weights ^ |
191 e_weight_arguments ctxt (method_for_slice ctxt slice) weights ^ |
180 " -tAutoDev --silent --cpu-limit=" ^ |
192 " -tAutoDev --silent --cpu-limit=" ^ |
181 string_of_int (to_secs (e_bonus ()) timeout), |
193 string_of_int (to_secs (e_bonus ()) timeout), |
182 proof_delims = [tstp_proof_delims], |
194 proof_delims = [tstp_proof_delims], |
183 known_failures = |
195 known_failures = |
184 [(Unprovable, "SZS status: CounterSatisfiable"), |
196 [(Unprovable, "SZS status: CounterSatisfiable"), |
189 "# Cannot determine problem status within resource limit"), |
201 "# Cannot determine problem status within resource limit"), |
190 (OutOfResources, "SZS status: ResourceOut"), |
202 (OutOfResources, "SZS status: ResourceOut"), |
191 (OutOfResources, "SZS status ResourceOut")], |
203 (OutOfResources, "SZS status ResourceOut")], |
192 hypothesis_kind = Conjecture, |
204 hypothesis_kind = Conjecture, |
193 formats = [Fof], |
205 formats = [Fof], |
194 best_slices = fn () => |
206 best_slices = fn ctxt => |
195 if effective_e_weight_method () = E_Slices then |
207 if effective_e_weight_method ctxt = e_slicesN then |
196 [(0.33333, (true, 100 (* FUDGE *))) (* E_Sym_Offset_Weight *), |
208 [(0.33333, (true, 100 (* FUDGE *))) (* sym_offset_weight *), |
197 (0.33333, (true, 1000 (* FUDGE *))) (* E_Auto *), |
209 (0.33333, (true, 1000 (* FUDGE *))) (* auto *), |
198 (0.33334, (true, 200 (* FUDGE *))) (* E_Fun_Weight *)] |
210 (0.33334, (true, 200 (* FUDGE *))) (* fun_weight *)] |
199 else |
211 else |
200 [(1.0, (true, 250 (* FUDGE *)))], |
212 [(1.0, (true, 250 (* FUDGE *)))], |
201 best_type_systems = ["const_args", "mangled_preds"]} |
213 best_type_systems = ["const_args", "mangled_preds"]} |
202 |
214 |
203 val e = (eN, e_config) |
215 val e = (eN, e_config) |
208 (* The "-VarWeight=3" option helps the higher-order problems, probably by |
220 (* The "-VarWeight=3" option helps the higher-order problems, probably by |
209 counteracting the presence of "hAPP". *) |
221 counteracting the presence of "hAPP". *) |
210 val spass_config : atp_config = |
222 val spass_config : atp_config = |
211 {exec = ("ISABELLE_ATP", "scripts/spass"), |
223 {exec = ("ISABELLE_ATP", "scripts/spass"), |
212 required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")], |
224 required_execs = [("SPASS_HOME", "SPASS"), ("SPASS_HOME", "tptp2dfg")], |
213 arguments = fn slice => fn timeout => fn _ => |
225 arguments = fn _ => fn slice => fn timeout => fn _ => |
214 ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \ |
226 ("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \ |
215 \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 0 timeout)) |
227 \-VarWeight=3 -TimeLimit=" ^ string_of_int (to_secs 0 timeout)) |
216 |> slice = 0 ? prefix "-SOS=1 ", |
228 |> slice = 0 ? prefix "-SOS=1 ", |
217 proof_delims = [("Here is a proof", "Formulae used in the proof")], |
229 proof_delims = [("Here is a proof", "Formulae used in the proof")], |
218 known_failures = |
230 known_failures = |
327 fun remote_config system_name system_versions proof_delims known_failures |
339 fun remote_config system_name system_versions proof_delims known_failures |
328 hypothesis_kind formats best_max_relevant best_type_systems |
340 hypothesis_kind formats best_max_relevant best_type_systems |
329 : atp_config = |
341 : atp_config = |
330 {exec = ("ISABELLE_ATP", "scripts/remote_atp"), |
342 {exec = ("ISABELLE_ATP", "scripts/remote_atp"), |
331 required_execs = [], |
343 required_execs = [], |
332 arguments = fn _ => fn timeout => fn _ => |
344 arguments = fn _ => fn _ => fn timeout => fn _ => |
333 " -t " ^ string_of_int (Int.min (max_remote_secs, (to_secs 0 timeout))) |
345 " -t " ^ string_of_int (Int.min (max_remote_secs, (to_secs 0 timeout))) |
334 ^ " -s " ^ the_system system_name system_versions, |
346 ^ " -s " ^ the_system system_name system_versions, |
335 proof_delims = insert (op =) tstp_proof_delims proof_delims, |
347 proof_delims = insert (op =) tstp_proof_delims proof_delims, |
336 known_failures = |
348 known_failures = |
337 known_failures @ known_perl_failures @ |
349 known_failures @ known_perl_failures @ |
338 [(IncompleteUnprovable, "says Unknown"), |
350 [(IncompleteUnprovable, "says Unknown"), |
339 (IncompleteUnprovable, "says GaveUp"), |
351 (IncompleteUnprovable, "says GaveUp"), |
340 (TimedOut, "says Timeout")], |
352 (TimedOut, "says Timeout")], |
341 hypothesis_kind = hypothesis_kind, |
353 hypothesis_kind = hypothesis_kind, |
342 formats = formats, |
354 formats = formats, |
343 best_slices = fn () => [(1.0, (false, best_max_relevant ()))], |
355 best_slices = fn ctxt => [(1.0, (false, best_max_relevant ctxt))], |
344 best_type_systems = best_type_systems} |
356 best_type_systems = best_type_systems} |
345 |
357 |
346 fun int_average f xs = fold (Integer.add o f) xs 0 div length xs |
358 fun int_average f xs = fold (Integer.add o f) xs 0 div length xs |
347 |
359 |
348 fun remotify_config system_name system_versions |
360 fun remotify_config system_name system_versions |