15608
|
1 |
(* Author: Jia Meng, Cambridge University Computer Laboratory
|
|
2 |
ID: $Id$
|
|
3 |
Copyright 2004 University of Cambridge
|
15347
|
4 |
|
|
5 |
ATPs with TPTP format input.
|
|
6 |
*)
|
15452
|
7 |
|
|
8 |
(*Jia: changed: isar_atp now processes entire proof context. fetch thms from delta simpset/claset*)
|
15644
|
9 |
(*Claire: changed: added actual watcher calls *)
|
15452
|
10 |
|
15347
|
11 |
signature RES_ATP =
|
|
12 |
sig
|
15608
|
13 |
val trace_res : bool ref
|
|
14 |
val axiom_file : Path.T
|
|
15 |
val hyps_file : Path.T
|
|
16 |
val isar_atp : ProofContext.context * Thm.thm -> unit
|
15644
|
17 |
(*val prob_file : Path.T*)
|
|
18 |
(*val atp_ax_tac : Thm.thm list -> int -> Tactical.tactic*)
|
|
19 |
(*val atp_tac : int -> Tactical.tactic*)
|
15608
|
20 |
val debug: bool ref
|
15347
|
21 |
|
|
22 |
end;
|
|
23 |
|
|
24 |
structure ResAtp : RES_ATP =
|
15608
|
25 |
|
15347
|
26 |
struct
|
|
27 |
|
15644
|
28 |
|
|
29 |
|
15608
|
30 |
(* used for debug *)
|
|
31 |
val debug = ref false;
|
15452
|
32 |
|
15608
|
33 |
fun debug_tac tac = (warning "testing";tac);
|
|
34 |
(* default value is false *)
|
15347
|
35 |
|
15608
|
36 |
val trace_res = ref false;
|
15347
|
37 |
|
15608
|
38 |
val skolem_tac = skolemize_tac;
|
|
39 |
|
15644
|
40 |
val num_of_clauses = ref 0;
|
|
41 |
val clause_arr = Array.array(3500, ("empty", 0));
|
|
42 |
|
15347
|
43 |
|
15608
|
44 |
val atomize_tac =
|
|
45 |
SUBGOAL
|
|
46 |
(fn (prop,_) =>
|
|
47 |
let val ts = Logic.strip_assums_hyp prop
|
|
48 |
in EVERY1
|
|
49 |
[METAHYPS
|
|
50 |
(fn hyps => (cut_facts_tac (map (ObjectLogic.atomize_thm o forall_intr_vars) hyps) 1)),
|
|
51 |
REPEAT_DETERM_N (length ts) o (etac thin_rl)]
|
|
52 |
end);
|
15347
|
53 |
|
15608
|
54 |
(* temporarily use these files, which will be loaded by Vampire *)
|
15644
|
55 |
val file_id_num =ref 0;
|
15608
|
56 |
|
15644
|
57 |
fun new_prob_file () = (file_id_num := (!file_id_num) + 1;"prob"^(string_of_int (!file_id_num)));
|
15608
|
58 |
|
15347
|
59 |
|
15644
|
60 |
val axiom_file = File.tmp_path (Path.basic "axioms");
|
|
61 |
val clasimp_file = File.tmp_path (Path.basic "clasimp");
|
|
62 |
val hyps_file = File.tmp_path (Path.basic "hyps");
|
|
63 |
val prob_file = File.tmp_path (Path.basic "prob");
|
|
64 |
val dummy_tac = PRIMITIVE(fn thm => thm );
|
|
65 |
|
|
66 |
|
|
67 |
fun concat_with_and [] str = str
|
|
68 |
| concat_with_and (x::[]) str = str^" ("^x^")"
|
|
69 |
| concat_with_and (x::xs) str = (concat_with_and xs (str^"("^x^")"^" & "))
|
|
70 |
|
15347
|
71 |
|
|
72 |
|
|
73 |
(**** for Isabelle/ML interface ****)
|
|
74 |
|
15644
|
75 |
fun is_proof_char ch = ((33 <= (ord ch)) andalso ((ord ch ) <= 126) andalso (not ((ord ch ) = 63))) orelse (ch = " ")
|
15608
|
76 |
|
15644
|
77 |
fun proofstring x = let val exp = explode x
|
|
78 |
in
|
|
79 |
List.filter (is_proof_char ) exp
|
|
80 |
end
|
15608
|
81 |
|
|
82 |
|
15452
|
83 |
|
15644
|
84 |
(*
|
|
85 |
fun call_atp_tac thms n = (tptp_inputs thms ; dummy_tac);
|
15347
|
86 |
|
15644
|
87 |
*)
|
15347
|
88 |
|
|
89 |
(**** For running in Isar ****)
|
|
90 |
|
15608
|
91 |
(* same function as that in res_axioms.ML *)
|
|
92 |
fun repeat_RS thm1 thm2 =
|
|
93 |
let val thm1' = thm1 RS thm2 handle THM _ => thm1
|
|
94 |
in
|
|
95 |
if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2)
|
|
96 |
end;
|
|
97 |
|
|
98 |
(* a special version of repeat_RS *)
|
|
99 |
fun repeat_someI_ex thm = repeat_RS thm someI_ex;
|
|
100 |
|
15644
|
101 |
(*********************************************************************)
|
15608
|
102 |
(* convert clauses from "assume" to conjecture. write to file "hyps" *)
|
15644
|
103 |
(* hypotheses of the goal currently being proved *)
|
|
104 |
(*********************************************************************)
|
|
105 |
|
15608
|
106 |
fun isar_atp_h thms =
|
15644
|
107 |
|
15608
|
108 |
let val prems = map (skolemize o make_nnf o ObjectLogic.atomize_thm) thms
|
15644
|
109 |
val prems' = map repeat_someI_ex prems
|
|
110 |
val prems'' = make_clauses prems'
|
|
111 |
val prems''' = ResAxioms.rm_Eps [] prems''
|
|
112 |
val clss = map ResClause.make_conjecture_clause prems'''
|
15608
|
113 |
val (tptp_clss,tfree_litss) = ResLib.unzip (map ResClause.clause2tptp clss)
|
|
114 |
val tfree_lits = ResLib.flat_noDup tfree_litss
|
|
115 |
val tfree_clss = map ResClause.tfree_clause tfree_lits
|
|
116 |
val hypsfile = File.sysify_path hyps_file
|
|
117 |
val out = TextIO.openOut(hypsfile)
|
|
118 |
in
|
|
119 |
((ResLib.writeln_strs out (tfree_clss @ tptp_clss); TextIO.closeOut out; if !trace_res then (warning hypsfile) else ());tfree_lits)
|
|
120 |
end;
|
15347
|
121 |
|
15644
|
122 |
|
|
123 |
(*********************************************************************)
|
|
124 |
(* write out a subgoal as tptp clauses to the file "probN" *)
|
|
125 |
(* where N is the number of this subgoal *)
|
|
126 |
(*********************************************************************)
|
|
127 |
|
15608
|
128 |
fun tptp_inputs_tfrees thms n tfrees =
|
|
129 |
let val clss = map (ResClause.make_conjecture_clause_thm) thms
|
|
130 |
val (tptp_clss,tfree_litss) = ResLib.unzip (map ResClause.clause2tptp clss)
|
|
131 |
val tfree_clss = map ResClause.tfree_clause ((ResLib.flat_noDup tfree_litss) \\ tfrees)
|
|
132 |
val probfile = (File.sysify_path prob_file) ^ "_" ^ (string_of_int n)
|
|
133 |
val out = TextIO.openOut(probfile)
|
|
134 |
in
|
|
135 |
(ResLib.writeln_strs out (tfree_clss @ tptp_clss); TextIO.closeOut out; (if !trace_res then (warning probfile) else ()))
|
|
136 |
end;
|
15452
|
137 |
|
15608
|
138 |
|
|
139 |
|
15644
|
140 |
(*********************************************************************)
|
|
141 |
(* call SPASS with settings and problem file for the current subgoal *)
|
|
142 |
(* should be modified to allow other provers to be called *)
|
|
143 |
(*********************************************************************)
|
15608
|
144 |
|
15644
|
145 |
fun call_resolve_tac thms sg_term (childin, childout,pid) = let
|
|
146 |
val newprobfile = new_prob_file ()
|
|
147 |
val thmstring = concat_with_and (map string_of_thm thms) ""
|
|
148 |
val _ = warning ("thmstring in call_res is: "^thmstring)
|
|
149 |
val goalstr = Sign.string_of_term Mainsign sg_term
|
|
150 |
val goalproofstring = proofstring goalstr
|
|
151 |
val no_returns =List.filter not_newline ( goalproofstring)
|
|
152 |
val goalstring = implode no_returns
|
|
153 |
val _ = warning ("goalstring in call_res is: "^goalstring)
|
|
154 |
|
|
155 |
val prob_file =File.tmp_path (Path.basic newprobfile);
|
|
156 |
val clauses = make_clauses thms
|
|
157 |
(*val _ = tptp_inputs clauses prob_file*)
|
|
158 |
val thmstring = concat_with_and (map string_of_thm thms) ""
|
|
159 |
|
|
160 |
val goalstr = Sign.string_of_term Mainsign sg_term
|
|
161 |
val goalproofstring = proofstring goalstr
|
|
162 |
val no_returns =List.filter not_newline ( goalproofstring)
|
|
163 |
val goalstring = implode no_returns
|
15608
|
164 |
|
15644
|
165 |
val thmproofstring = proofstring ( thmstring)
|
|
166 |
val no_returns =List.filter not_newline ( thmproofstring)
|
|
167 |
val thmstr = implode no_returns
|
|
168 |
|
|
169 |
val prob_path = File.sysify_path prob_file
|
|
170 |
val outfile = TextIO.openOut("/home/clq20/Jia_Code/hellofile")
|
|
171 |
val _ = TextIO.output(outfile, "prob file path is "^prob_path^" thmstring is "^thmstr^" goalstring is "^goalstring);
|
|
172 |
val _ = TextIO.flushOut outfile;
|
|
173 |
val _ = TextIO.closeOut outfile
|
|
174 |
in
|
|
175 |
(* without paramodulation *)
|
|
176 |
(*(warning ("goalstring in call_res_tac is: "^goalstring));
|
|
177 |
(warning ("prob path in cal_res_tac is: "^prob_path));
|
|
178 |
Watcher.callResProvers(childout,
|
|
179 |
[("spass",thmstr,goalstring,"/homes/clq20/bin/SPASS",
|
|
180 |
"-FullRed=0%-Auto=0%-ISRe%-ISFc%-RTaut%-RFSub%-RBSub%-DocProof",
|
|
181 |
prob_path)]);*)
|
|
182 |
(* with paramodulation *)
|
|
183 |
(*Watcher.callResProvers(childout,
|
|
184 |
[("spass",thmstr,goalstring,"/homes/clq20/bin/SPASS",
|
|
185 |
"-FullRed=0%-ISPm=1%-Split=0%-PObv=0%-DocProof",
|
|
186 |
prob_path)]); *)
|
|
187 |
Watcher.callResProvers(childout,
|
|
188 |
[("spass",thmstr,goalstring,"/homes/clq20/bin/SPASS",
|
|
189 |
"-DocProof", prob_path)]);
|
|
190 |
dummy_tac
|
|
191 |
end
|
15452
|
192 |
|
15644
|
193 |
(************************************************)
|
|
194 |
(* pass in subgoal as a term and watcher info *)
|
|
195 |
(* process subgoal into skolemized, negated form*)
|
|
196 |
(* then call call_resolve_tac to send to ATP *)
|
|
197 |
(************************************************)
|
15608
|
198 |
|
15644
|
199 |
fun resolve_tac sg_term (childin, childout,pid) =
|
|
200 |
let val _ = warning ("in resolve_tac ")
|
|
201 |
in
|
|
202 |
SELECT_GOAL
|
|
203 |
(EVERY1 [rtac ccontr,atomize_tac,skolemize_tac, METAHYPS(fn negs => (warning ("calling call_resolve_tac next ");call_resolve_tac negs sg_term (childin, childout,pid)))])
|
|
204 |
end;
|
15452
|
205 |
|
15603
|
206 |
|
15608
|
207 |
|
15452
|
208 |
|
15644
|
209 |
(* Need to replace call_atp_tac_tfrees with call res_provers as it's the dummy one *)
|
|
210 |
|
|
211 |
(**********************************************************)
|
|
212 |
(* write out the current subgoal as a tptp file, probN, *)
|
|
213 |
(* then call dummy_tac - should be call_res_tac *)
|
|
214 |
(**********************************************************)
|
|
215 |
|
|
216 |
fun call_atp_tac_tfrees thms n tfrees sg_term (childin, childout,pid) =
|
|
217 |
(tptp_inputs_tfrees (make_clauses thms) n tfrees;
|
|
218 |
resolve_tac sg_term (childin, childout, pid);
|
|
219 |
dummy_tac);
|
|
220 |
|
|
221 |
fun atp_tac_tfrees tfrees sg_term (childin, childout,pid) n =
|
|
222 |
let val _ = (warning ("in atp_tac_tfrees "))
|
|
223 |
in
|
|
224 |
SELECT_GOAL
|
|
225 |
(EVERY1 [rtac ccontr,atomize_tac, skolemize_tac,
|
|
226 |
METAHYPS(fn negs => (call_atp_tac_tfrees (negs) n tfrees sg_term (childin, childout,pid)))]) n
|
|
227 |
end;
|
|
228 |
|
|
229 |
|
|
230 |
fun isar_atp_g tfrees sg_term (childin, childout, pid) =
|
|
231 |
|
|
232 |
( (warning("in isar_atp_g"));atp_tac_tfrees tfrees sg_term (childin, childout, pid));
|
|
233 |
|
|
234 |
|
|
235 |
|
|
236 |
(**********************************************)
|
|
237 |
(* recursively call atp_tac_g on all subgoals *)
|
|
238 |
(* sg_term is the nth subgoal as a term - used*)
|
|
239 |
(* in proof reconstruction *)
|
|
240 |
(**********************************************)
|
|
241 |
|
|
242 |
fun isar_atp_goal' thm k n tfree_lits (childin, childout, pid) =
|
|
243 |
if (k > n)
|
|
244 |
then ()
|
|
245 |
else
|
|
246 |
(let val prems = prems_of thm
|
|
247 |
val sg_term = get_nth n prems
|
|
248 |
in
|
|
249 |
|
|
250 |
(warning("in isar_atp_goal'"));
|
|
251 |
isar_atp_g tfree_lits sg_term (childin, childout, pid) k thm;
|
|
252 |
isar_atp_goal' thm (k+1) n tfree_lits (childin, childout, pid)
|
|
253 |
end);
|
|
254 |
|
|
255 |
|
|
256 |
fun isar_atp_goal thm n_subgoals tfree_lits (childin, childout, pid) = (if (!debug) then warning (string_of_thm thm) else (isar_atp_goal' thm 1 n_subgoals tfree_lits (childin, childout, pid) ));
|
|
257 |
|
|
258 |
(**************************************************)
|
|
259 |
(* convert clauses from "assume" to conjecture. *)
|
|
260 |
(* i.e. apply make_clauses and then get tptp for *)
|
|
261 |
(* any hypotheses in the goal *)
|
|
262 |
(* write to file "hyps" *)
|
|
263 |
(**************************************************)
|
|
264 |
|
|
265 |
|
|
266 |
fun isar_atp_aux thms thm n_subgoals (childin, childout, pid) =
|
15608
|
267 |
let val tfree_lits = isar_atp_h thms
|
|
268 |
in
|
15644
|
269 |
(warning("in isar_atp_aux"));isar_atp_goal thm n_subgoals tfree_lits (childin, childout, pid)
|
15608
|
270 |
end;
|
15644
|
271 |
|
|
272 |
(******************************************************************)
|
|
273 |
(* called in Isar automatically *)
|
|
274 |
(* writes out the current clasimpset to a tptp file *)
|
|
275 |
(* passes all subgoals on to isar_atp_aux for further processing *)
|
|
276 |
(* turns off xsymbol at start of function, restoring it at end *)
|
|
277 |
(******************************************************************)
|
15452
|
278 |
|
15608
|
279 |
fun isar_atp' (thms, thm) =
|
15644
|
280 |
let val _= (print_mode := (Library.gen_rems (op =) (! print_mode, ["xsymbols", "symbols"])))
|
|
281 |
val _= (warning ("in isar_atp'"))
|
|
282 |
val prems = prems_of thm
|
|
283 |
val thms_string =concat_with_and (map string_of_thm thms) ""
|
|
284 |
val thmstring = string_of_thm thm
|
|
285 |
val prems_string = concat_with_and (map (Sign.string_of_term Mainsign) prems) ""
|
|
286 |
(* set up variables for writing out the clasimps to a tptp file *)
|
|
287 |
(* val _ = write_out_clasimp (File.sysify_path axiom_file)*)
|
|
288 |
(* cq: add write out clasimps to file *)
|
|
289 |
(* cq:create watcher and pass to isar_atp_aux *)
|
|
290 |
val (childin,childout,pid) = Watcher.createWatcher(thm)
|
|
291 |
val pidstring = string_of_int(Word.toInt (Word.fromLargeWord ( Posix.Process.pidToWord pid )))
|
15608
|
292 |
in
|
15644
|
293 |
case prems of [] => ()
|
|
294 |
| _ => ((warning ("initial thms: "^thms_string));
|
|
295 |
(warning ("initial thm: "^thmstring));
|
|
296 |
(warning ("subgoals: "^prems_string));
|
|
297 |
(warning ("pid: "^ pidstring)));
|
|
298 |
isar_atp_aux thms thm (length prems) (childin, childout, pid) ;
|
|
299 |
|
|
300 |
print_mode := (["xsymbols", "symbols"] @ ! print_mode)
|
15608
|
301 |
end;
|
|
302 |
|
15452
|
303 |
|
15608
|
304 |
|
|
305 |
|
|
306 |
local
|
|
307 |
|
|
308 |
fun get_thms_cs claset =
|
|
309 |
let val clsset = rep_cs claset
|
|
310 |
val safeEs = #safeEs clsset
|
|
311 |
val safeIs = #safeIs clsset
|
|
312 |
val hazEs = #hazEs clsset
|
|
313 |
val hazIs = #hazIs clsset
|
|
314 |
in
|
|
315 |
safeEs @ safeIs @ hazEs @ hazIs
|
|
316 |
end;
|
|
317 |
|
|
318 |
|
15452
|
319 |
|
15608
|
320 |
fun append_name name [] _ = []
|
|
321 |
| append_name name (thm::thms) k = (Thm.name_thm ((name ^ "_" ^ (string_of_int k)),thm)) :: (append_name name thms (k+1));
|
|
322 |
|
|
323 |
fun append_names (name::names) (thms::thmss) =
|
|
324 |
let val thms' = append_name name thms 0
|
|
325 |
in
|
|
326 |
thms'::(append_names names thmss)
|
|
327 |
end;
|
|
328 |
|
15452
|
329 |
|
15608
|
330 |
fun get_thms_ss [] = []
|
|
331 |
| get_thms_ss thms =
|
|
332 |
let val names = map Thm.name_of_thm thms
|
|
333 |
val thms' = map (mksimps mksimps_pairs) thms
|
|
334 |
val thms'' = append_names names thms'
|
|
335 |
in
|
|
336 |
ResLib.flat_noDup thms''
|
|
337 |
end;
|
|
338 |
|
15452
|
339 |
|
15608
|
340 |
|
|
341 |
|
|
342 |
in
|
|
343 |
|
15452
|
344 |
|
15608
|
345 |
(* convert locally declared rules to axiom clauses *)
|
|
346 |
(* write axiom clauses to ax_file *)
|
15644
|
347 |
(* what about clasimpset - it should already be in the ax file - perhaps append to ax file rather than just *)
|
|
348 |
(* write out ? Or keep as a separate file and then cat them all together in the watcher, like we do with the *)
|
|
349 |
(*claset file and prob file*)
|
15608
|
350 |
fun isar_local_thms (delta_cs, delta_ss_thms) =
|
|
351 |
let val thms_cs = get_thms_cs delta_cs
|
|
352 |
val thms_ss = get_thms_ss delta_ss_thms
|
|
353 |
val thms_clauses = ResLib.flat_noDup (map ResAxioms.clausify_axiom (thms_cs @ thms_ss))
|
|
354 |
val clauses_strs = ResLib.flat_noDup (map ResClause.tptp_clause thms_clauses) (*string list*)
|
|
355 |
val ax_file = File.sysify_path axiom_file
|
|
356 |
val out = TextIO.openOut ax_file
|
|
357 |
in
|
15644
|
358 |
(ResLib.writeln_strs out clauses_strs; (warning ("axiom file is: "^ax_file));TextIO.closeOut out)
|
15608
|
359 |
end;
|
15347
|
360 |
|
15608
|
361 |
|
|
362 |
|
|
363 |
|
15347
|
364 |
|
15608
|
365 |
(* called in Isar automatically *)
|
|
366 |
fun isar_atp (ctxt,thm) =
|
|
367 |
let val prems = ProofContext.prems_of ctxt
|
15644
|
368 |
val d_cs = Classical.get_delta_claset ctxt
|
|
369 |
val d_ss_thms = Simplifier.get_delta_simpset ctxt
|
|
370 |
val thmstring = string_of_thm thm
|
|
371 |
val prem_no = length prems
|
|
372 |
val prems_string = concat_with_and (map string_of_thm prems) ""
|
15608
|
373 |
in
|
15644
|
374 |
|
|
375 |
(warning ("initial thm in isar_atp: "^thmstring));
|
|
376 |
(warning ("subgoals in isar_atp: "^prems_string));
|
|
377 |
(warning ("number of subgoals in isar_atp: "^(string_of_int prem_no)));
|
|
378 |
(isar_local_thms (d_cs,d_ss_thms); (warning("about to call isar_atp'"));
|
|
379 |
isar_atp' (prems, thm))
|
15608
|
380 |
end;
|
15347
|
381 |
|
15608
|
382 |
end
|
|
383 |
|
|
384 |
|
|
385 |
|
15452
|
386 |
|
15347
|
387 |
end;
|
|
388 |
|
|
389 |
Proof.atp_hook := ResAtp.isar_atp;
|