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