| author | wenzelm | 
| Wed, 10 May 2006 16:23:21 +0200 | |
| changeset 19608 | 81fe44909dd5 | 
| parent 17776 | 4a518eec4a20 | 
| child 19643 | 213e12ad2c03 | 
| permissions | -rw-r--r-- | 
| 15789 
4cb16144c81b
added hearder lines and deleted some redundant material
 paulson parents: 
15774diff
changeset | 1 | (* ID: $Id$ | 
| 
4cb16144c81b
added hearder lines and deleted some redundant material
 paulson parents: 
15774diff
changeset | 2 | Author: Claire Quigley | 
| 
4cb16144c81b
added hearder lines and deleted some redundant material
 paulson parents: 
15774diff
changeset | 3 | Copyright 2004 University of Cambridge | 
| 
4cb16144c81b
added hearder lines and deleted some redundant material
 paulson parents: 
15774diff
changeset | 4 | *) | 
| 15642 | 5 | |
| 16803 | 6 | structure ReconTranslateProof = | 
| 7 | struct | |
| 8 | ||
| 15642 | 9 | fun add_in_order (x:string) [] = [x] | 
| 10 | | add_in_order (x:string) ((y:string)::ys) = if (x < y) | |
| 11 | then | |
| 12 | (x::(y::ys)) | |
| 13 | else | |
| 14 | (y::(add_in_order x ys)) | |
| 15 | fun add_once x [] = [x] | |
| 16157 | 16 | | add_once x (y::ys) = if x mem (y::ys) then (y::ys) | 
| 17 | else add_in_order x (y::ys) | |
| 15642 | 18 | |
| 15684 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 19 | fun thm_vars thm = map (fn ((name,ix),_) => name) (Drule.vars_of thm); | 
| 15642 | 20 | |
| 21 | exception Reflex_equal; | |
| 22 | ||
| 23 | (********************************) | |
| 24 | (* Proofstep datatype *) | |
| 25 | (********************************) | |
| 26 | (* Assume rewrite steps have only two clauses in them just now, but lcl109-5 has Rew[5,3,5,3] *) | |
| 27 | ||
| 28 | datatype Side = Left |Right | |
| 29 | ||
| 16803 | 30 | datatype Proofstep = ExtraAxiom | 
| 16418 | 31 | | OrigAxiom | 
| 32 | | VampInput | |
| 15642 | 33 | | Axiom | 
| 34 | | Binary of (int * int) * (int * int) (* (clause1,lit1), (clause2, lit2) *) | |
| 35 | | MRR of (int * int) * (int * int) | |
| 36 | | Factor of (int * int * int) (* (clause,lit1, lit2) *) | |
| 37 | | Para of (int * int) * (int * int) | |
| 16357 | 38 | | Super_l of (int * int) * (int * int) | 
| 39 | | Super_r of (int * int) * (int * int) | |
| 16548 
aa36ae6b955e
Temporarily removed Rewrite from the translation code so that parsing with work on lists of numbers.
 quigley parents: 
16433diff
changeset | 40 | (*| Rewrite of (int * int) * (int * int) *) | 
| 
aa36ae6b955e
Temporarily removed Rewrite from the translation code so that parsing with work on lists of numbers.
 quigley parents: 
16433diff
changeset | 41 | | Rewrite of (int * int) list | 
| 16357 | 42 | | SortSimp of (int * int) * (int * int) | 
| 43 | | UnitConf of (int * int) * (int * int) | |
| 15642 | 44 | | Obvious of (int * int) | 
| 16357 | 45 | | AED of (int*int) | 
| 46 | | EqualRes of (int*int) | |
| 47 | | Condense of (int*int) | |
| 15642 | 48 | (*| Hyper of int list*) | 
| 49 | | Unusedstep of unit | |
| 50 | ||
| 51 | ||
| 52 | datatype Clausesimp = Binary_s of int * int | |
| 53 | | Factor_s of unit | |
| 54 | | Demod_s of (int * int) list | |
| 55 | | Hyper_s of int list | |
| 56 | | Unusedstep_s of unit | |
| 57 | ||
| 58 | ||
| 59 | ||
| 60 | datatype Step_label = T_info | |
| 61 | |P_info | |
| 62 | ||
| 63 | ||
| 16061 | 64 | fun is_alpha_space_digit_or_neg ch = | 
| 65 | (ch = "~") orelse (ReconOrderClauses.is_alpha ch) orelse | |
| 66 | (ReconOrderClauses.is_digit ch) orelse ( ch = " "); | |
| 15642 | 67 | |
| 17776 | 68 | fun lit_string_with_nums t = let val termstr = Term.str_of_term t | 
| 15642 | 69 | val exp_term = explode termstr | 
| 70 | in | |
| 71 | implode(List.filter is_alpha_space_digit_or_neg exp_term) | |
| 72 | end | |
| 73 | ||
| 15774 | 74 | fun reconstruction_failed fname = error (fname ^ ": reconstruction failed") | 
| 15642 | 75 | |
| 15919 
b30a35432f5a
Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
 quigley parents: 
15794diff
changeset | 76 | (************************************************) | 
| 
b30a35432f5a
Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
 quigley parents: 
15794diff
changeset | 77 | (* Reconstruct an axiom resolution step *) | 
| 
b30a35432f5a
Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
 quigley parents: 
15794diff
changeset | 78 | (* clauses is a list of (clausenum,clause) pairs*) | 
| 
b30a35432f5a
Replaced reference to SPASS with general one - set SPASS_HOME in settings file.
 quigley parents: 
15794diff
changeset | 79 | (************************************************) | 
| 15642 | 80 | |
| 15774 | 81 | fun follow_axiom clauses allvars (clausenum:int) thml clause_strs = | 
| 17374 | 82 | let val this_axiom = (the o AList.lookup (op =) clauses) clausenum | 
| 16061 | 83 | val (res, numlist, symlist, nsymlist) = (ReconOrderClauses.rearrange_clause this_axiom clause_strs allvars) | 
| 84 | val thmvars = thm_vars res | |
| 85 | in | |
| 86 | (res, (Axiom,clause_strs,thmvars ) ) | |
| 87 | end | |
| 88 | handle Option => reconstruction_failed "follow_axiom" | |
| 15642 | 89 | |
| 90 | (****************************************) | |
| 91 | (* Reconstruct a binary resolution step *) | |
| 92 | (****************************************) | |
| 93 | ||
| 94 | (* numbers of clauses and literals*) (*vars*) (*list of current thms*) (* literal strings as parsed from spass *) | |
| 16061 | 95 | fun follow_binary ((clause1, lit1), (clause2 , lit2)) allvars thml clause_strs = | 
| 17374 | 96 | let val thm1 = select_literal (lit1 + 1) ((the o AList.lookup (op =) thml) clause1) | 
| 97 | val thm2 = (the o AList.lookup (op =) thml) clause2 | |
| 16061 | 98 | val res = thm1 RSN ((lit2 +1), thm2) | 
| 99 | val (res', numlist, symlist, nsymlist) = (ReconOrderClauses.rearrange_clause res clause_strs allvars) | |
| 100 | val thmvars = thm_vars res | |
| 101 | in | |
| 102 | (res', ((Binary ((clause1, lit1), (clause2 , lit2))),clause_strs,thmvars) ) | |
| 103 | end | |
| 104 | handle Option => reconstruction_failed "follow_binary" | |
| 15642 | 105 | |
| 106 | ||
| 107 | ||
| 108 | (******************************************************) | |
| 109 | (* Reconstruct a matching replacement resolution step *) | |
| 110 | (******************************************************) | |
| 111 | ||
| 112 | ||
| 16061 | 113 | fun follow_mrr ((clause1, lit1), (clause2 , lit2)) allvars thml clause_strs = | 
| 17374 | 114 | let val thm1 = select_literal (lit1 + 1) ((the o AList.lookup (op =) thml) clause1) | 
| 115 | val thm2 = (the o AList.lookup (op =) thml) clause2 | |
| 16061 | 116 | val res = thm1 RSN ((lit2 +1), thm2) | 
| 117 | val (res', numlist, symlist, nsymlist) = | |
| 118 | (ReconOrderClauses.rearrange_clause res clause_strs allvars) | |
| 119 | val thmvars = thm_vars res | |
| 120 | in | |
| 121 | (res', ((MRR ((clause1, lit1), (clause2 , lit2))) ,clause_strs,thmvars)) | |
| 122 | end | |
| 123 | handle Option => reconstruction_failed "follow_mrr" | |
| 15642 | 124 | |
| 125 | ||
| 126 | (*******************************************) | |
| 127 | (* Reconstruct a factoring resolution step *) | |
| 128 | (*******************************************) | |
| 129 | ||
| 130 | fun mksubstlist [] sublist = sublist | |
| 16061 | 131 | |mksubstlist ((a, (_, b)) :: rest) sublist = | 
| 132 | let val vartype = type_of b | |
| 133 | val avar = Var(a,vartype) | |
| 134 | val newlist = ((avar,b)::sublist) | |
| 135 | in | |
| 136 | mksubstlist rest newlist | |
| 137 | end; | |
| 15642 | 138 | |
| 139 | (* based on Tactic.distinct_subgoals_tac *) | |
| 140 | ||
| 141 | fun delete_assump_tac state lit = | |
| 142 | let val (frozth,thawfn) = freeze_thaw state | |
| 143 | val froz_prems = cprems_of frozth | |
| 144 | val assumed = implies_elim_list frozth (map assume froz_prems) | |
| 16061 | 145 | val new_prems = ReconOrderClauses.remove_nth lit froz_prems | 
| 15642 | 146 | val implied = implies_intr_list new_prems assumed | 
| 147 | in | |
| 148 | Seq.single (thawfn implied) | |
| 149 | end | |
| 150 | ||
| 151 | ||
| 152 | ||
| 153 | ||
| 154 | ||
| 155 | (*************************************) | |
| 156 | (* Reconstruct a factoring step *) | |
| 157 | (*************************************) | |
| 158 | ||
| 17488 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 159 | fun getnewenv seq = fst (fst (the (Seq.pull seq))); | 
| 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 160 | |
| 15739 
bb2acfed8212
yet more tidying up: removal of some references to Main
 paulson parents: 
15736diff
changeset | 161 | (*FIXME: share code with that in Tools/reconstruction.ML*) | 
| 16061 | 162 | fun follow_factor clause lit1 lit2 allvars thml clause_strs = | 
| 163 | let | |
| 17374 | 164 | val th = (the o AList.lookup (op =) thml) clause | 
| 16061 | 165 | val prems = prems_of th | 
| 166 | val sign = sign_of_thm th | |
| 167 | val fac1 = ReconOrderClauses.get_nth (lit1+1) prems | |
| 168 | val fac2 = ReconOrderClauses.get_nth (lit2+1) prems | |
| 169 | val unif_env = Unify.unifiers (sign,Envir.empty 0, [(fac1, fac2)]) | |
| 170 | val newenv = getnewenv unif_env | |
| 171 | val envlist = Envir.alist_of newenv | |
| 172 | ||
| 173 | val (t1,t2)::_ = mksubstlist envlist [] | |
| 174 | in | |
| 175 | if (is_Var t1) | |
| 176 | then | |
| 177 | let | |
| 178 | val str1 = lit_string_with_nums t1; | |
| 179 | val str2 = lit_string_with_nums t2; | |
| 180 | val facthm = read_instantiate [(str1,str2)] th; | |
| 181 | val res = hd (Seq.list_of(delete_assump_tac facthm (lit1 + 1))) | |
| 182 | val (res', numlist, symlist, nsymlist) = | |
| 183 | ReconOrderClauses.rearrange_clause res clause_strs allvars | |
| 184 | val thmvars = thm_vars res' | |
| 185 | in | |
| 186 | (res',((Factor (clause, lit1, lit2)),clause_strs,thmvars)) | |
| 187 | end | |
| 188 | else | |
| 189 | let | |
| 190 | val str2 = lit_string_with_nums t1; | |
| 191 | val str1 = lit_string_with_nums t2; | |
| 192 | val facthm = read_instantiate [(str1,str2)] th; | |
| 193 | val res = hd (Seq.list_of(delete_assump_tac facthm (lit1 + 1))) | |
| 194 | val (res', numlist, symlist, nsymlist) = | |
| 195 | ReconOrderClauses.rearrange_clause res clause_strs allvars | |
| 196 | val thmvars = thm_vars res' | |
| 197 | in | |
| 198 | (res', ((Factor (clause, lit1, lit2)),clause_strs, thmvars)) | |
| 199 | end | |
| 200 | end | |
| 201 | handle Option => reconstruction_failed "follow_factor" | |
| 15642 | 202 | |
| 203 | ||
| 204 | ||
| 15684 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 205 | fun get_unif_comb t eqterm = | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 206 | if ((type_of t) = (type_of eqterm)) | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 207 | then t | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 208 | else | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 209 | let val _ $ rand = t | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 210 | in get_unif_comb rand eqterm end; | 
| 15642 | 211 | |
| 15684 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 212 | fun get_unif_lit t eqterm = | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 213 | if (can HOLogic.dest_eq t) | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 214 | then | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 215 | let val (lhs,rhs) = HOLogic.dest_eq(HOLogic.dest_Trueprop eqterm) | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 216 | in lhs end | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 217 | else | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 218 | get_unif_comb t eqterm; | 
| 
5ec4d21889d6
Reconstruction code, now packaged to avoid name clashes
 paulson parents: 
15658diff
changeset | 219 | |
| 15642 | 220 | |
| 221 | ||
| 222 | (*************************************) | |
| 223 | (* Reconstruct a paramodulation step *) | |
| 224 | (*************************************) | |
| 225 | ||
| 226 | val rev_subst = rotate_prems 1 subst; | |
| 227 | val rev_ssubst = rotate_prems 1 ssubst; | |
| 228 | ||
| 229 | ||
| 16061 | 230 | fun follow_standard_para ((clause1, lit1), (clause2 , lit2)) allvars thml clause_strs = | 
| 231 | let | |
| 17374 | 232 | val th1 = (the o AList.lookup (op =) thml) clause1 | 
| 233 | val th2 = (the o AList.lookup (op =) thml) clause2 | |
| 16061 | 234 | val eq_lit_th = select_literal (lit1+1) th1 | 
| 235 | val mod_lit_th = select_literal (lit2+1) th2 | |
| 236 | val eqsubst = eq_lit_th RSN (2,rev_subst) | |
| 237 | val newth = Seq.hd (biresolution false [(false, mod_lit_th)] 1 eqsubst) | |
| 238 | val newth' =negate_head newth | |
| 239 | val (res, numlist, symlist, nsymlist) = | |
| 240 | (ReconOrderClauses.rearrange_clause newth' clause_strs allvars | |
| 241 | handle Not_in_list => | |
| 242 | let val mod_lit_th' = mod_lit_th RS not_sym | |
| 243 | val newth = Seq.hd (biresolution false [(false, mod_lit_th')] 1 eqsubst) | |
| 244 | val newth' =negate_head newth | |
| 245 | in | |
| 246 | ReconOrderClauses.rearrange_clause newth' clause_strs allvars | |
| 247 | end) | |
| 248 | val thmvars = thm_vars res | |
| 249 | in | |
| 250 | (res, ((Para((clause1, lit1), (clause2 , lit2))),clause_strs,thmvars)) | |
| 251 | end | |
| 252 | handle Option => reconstruction_failed "follow_standard_para" | |
| 15642 | 253 | |
| 254 | ||
| 255 | (********************************) | |
| 256 | (* Reconstruct a rewriting step *) | |
| 257 | (********************************) | |
| 258 | ||
| 259 | (* | |
| 260 | ||
| 261 | val rev_subst = rotate_prems 1 subst; | |
| 262 | ||
| 263 | fun paramod_rule ((cl1, lit1), (cl2 , lit2)) = | |
| 264 | let val eq_lit_th = select_literal (lit1+1) cl1 | |
| 265 | val mod_lit_th = select_literal (lit2+1) cl2 | |
| 266 | val eqsubst = eq_lit_th RSN (2,rev_subst) | |
| 267 | val newth = Seq.hd (biresolution false [(false, mod_lit_th)] 1 | |
| 268 | eqsubst) | |
| 15697 | 269 | in Meson.negated_asm_of_head newth end; | 
| 15642 | 270 | |
| 271 | val newth = Seq.hd (biresolution false [(false, mod_lit_th)] 1 | |
| 272 | eqsubst) | |
| 273 | ||
| 274 | val mod_lit_th' = mod_lit_th RS not_sym | |
| 275 | ||
| 276 | *) | |
| 277 | ||
| 278 | ||
| 279 | fun delete_assumps 0 th = th | |
| 280 | | delete_assumps n th = let val th_prems = length (prems_of th) | |
| 281 | val th' = hd (Seq.list_of(delete_assump_tac th (th_prems -1))) | |
| 282 | in | |
| 283 | delete_assumps (n-1) th' | |
| 284 | end | |
| 285 | ||
| 286 | (* extra conditions from the equality turn up as 2nd to last literals of result. Need to delete them *) | |
| 287 | (* changed negate_nead to negate_head here too*) | |
| 288 | (* clause1, lit1 is thing to rewrite with *) | |
| 16548 
aa36ae6b955e
Temporarily removed Rewrite from the translation code so that parsing with work on lists of numbers.
 quigley parents: 
16433diff
changeset | 289 | (*fun follow_rewrite ((clause1, lit1), (clause2, lit2)) allvars thml clause_strs = | 
| 17374 | 290 | let val th1 = (the o AList.lookup (op =) thml) clause1 | 
| 291 | val th2 = (the o AList.lookup (op =) thml) clause2 | |
| 16061 | 292 | val eq_lit_th = select_literal (lit1+1) th1 | 
| 293 | val mod_lit_th = select_literal (lit2+1) th2 | |
| 294 | val eqsubst = eq_lit_th RSN (2,rev_subst) | |
| 295 | val eq_lit_prem_num = length (prems_of eq_lit_th) | |
| 16433 | 296 | val sign = Theory.merge (sign_of_thm th1, sign_of_thm th2) | 
| 16061 | 297 | val newth = Seq.hd (biresolution false [(false, mod_lit_th)] 1 eqsubst) | 
| 298 | val newthm = negate_head newth | |
| 299 | val newthm' = delete_assumps eq_lit_prem_num newthm | |
| 300 | val (res, numlist, symlist, nsymlist) = | |
| 301 | ReconOrderClauses.rearrange_clause newthm clause_strs allvars | |
| 302 | val thmvars = thm_vars res | |
| 303 | in | |
| 304 | (res, ((Rewrite ((clause1, lit1), (clause2, lit2))),clause_strs,thmvars)) | |
| 305 | end | |
| 306 | handle Option => reconstruction_failed "follow_rewrite" | |
| 15642 | 307 | |
| 16548 
aa36ae6b955e
Temporarily removed Rewrite from the translation code so that parsing with work on lists of numbers.
 quigley parents: 
16433diff
changeset | 308 | *) | 
| 15642 | 309 | |
| 310 | (*****************************************) | |
| 311 | (* Reconstruct an obvious reduction step *) | |
| 312 | (*****************************************) | |
| 313 | ||
| 314 | ||
| 16061 | 315 | fun follow_obvious (clause1, lit1) allvars thml clause_strs = | 
| 17374 | 316 | let val th1 = (the o AList.lookup (op =) thml) clause1 | 
| 16061 | 317 | val prems1 = prems_of th1 | 
| 318 | val newthm = refl RSN ((lit1+ 1), th1) | |
| 319 | handle THM _ => hd (Seq.list_of(delete_assump_tac th1 (lit1 + 1))) | |
| 320 | val (res, numlist, symlist, nsymlist) = | |
| 321 | ReconOrderClauses.rearrange_clause newthm clause_strs allvars | |
| 322 | val thmvars = thm_vars res | |
| 323 | in | |
| 324 | (res, ((Obvious (clause1, lit1)),clause_strs,thmvars)) | |
| 325 | end | |
| 326 | handle Option => reconstruction_failed "follow_obvious" | |
| 15642 | 327 | |
| 328 | (**************************************************************************************) | |
| 329 | (* reconstruct a single line of the res. proof - depending on what sort of proof step it was*) | |
| 330 | (**************************************************************************************) | |
| 331 | ||
| 332 | fun follow_proof clauses allvars clausenum thml (Axiom ) clause_strs | |
| 333 | = follow_axiom clauses allvars clausenum thml clause_strs | |
| 334 | ||
| 335 | | follow_proof clauses allvars clausenum thml (Binary (a, b)) clause_strs | |
| 336 | = follow_binary (a, b) allvars thml clause_strs | |
| 337 | ||
| 338 | | follow_proof clauses allvars clausenum thml (MRR (a, b)) clause_strs | |
| 339 | = follow_mrr (a, b) allvars thml clause_strs | |
| 340 | ||
| 341 | | follow_proof clauses allvars clausenum thml (Factor (a, b, c)) clause_strs | |
| 342 | = follow_factor a b c allvars thml clause_strs | |
| 343 | ||
| 344 | | follow_proof clauses allvars clausenum thml (Para (a, b)) clause_strs | |
| 345 | = follow_standard_para (a, b) allvars thml clause_strs | |
| 346 | ||
| 16548 
aa36ae6b955e
Temporarily removed Rewrite from the translation code so that parsing with work on lists of numbers.
 quigley parents: 
16433diff
changeset | 347 | (* | follow_proof clauses allvars clausenum thml (Rewrite (a,b)) clause_strs | 
| 
aa36ae6b955e
Temporarily removed Rewrite from the translation code so that parsing with work on lists of numbers.
 quigley parents: 
16433diff
changeset | 348 | = follow_rewrite (a,b) allvars thml clause_strs*) | 
| 15642 | 349 | |
| 350 | | follow_proof clauses allvars clausenum thml (Obvious (a,b)) clause_strs | |
| 351 | = follow_obvious (a,b) allvars thml clause_strs | |
| 352 | ||
| 353 | (*| follow_proof clauses clausenum thml (Hyper l) | |
| 354 | = follow_hyper l thml*) | |
| 355 | | follow_proof clauses allvars clausenum thml _ _ = | |
| 17488 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 356 | error "proof step not implemented" | 
| 15642 | 357 | |
| 358 | ||
| 359 | ||
| 360 | ||
| 361 | (**************************************************************************************) | |
| 362 | (* reconstruct a single line of the res. proof - at the moment do only inference steps*) | |
| 363 | (**************************************************************************************) | |
| 364 | ||
| 17488 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 365 | fun follow_line clauses allvars thml (clause_num, proof_step, clause_strs) recons = | 
| 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 366 | let | 
| 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 367 | val (thm, recon_fun) = follow_proof clauses allvars clause_num thml proof_step clause_strs | 
| 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 368 | in | 
| 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 369 | ((clause_num, thm)::thml, (clause_num,recon_fun)::recons) | 
| 
67376a311a2b
further simplification of the Isabelle-ATP linkup
 paulson parents: 
17374diff
changeset | 370 | end | 
| 15642 | 371 | |
| 372 | (***************************************************************) | |
| 373 | (* follow through the res. proof, creating an Isabelle theorem *) | |
| 374 | (***************************************************************) | |
| 375 | ||
| 376 | fun is_proof_char ch = ((33 <= (ord ch)) andalso ((ord ch ) <= 126)) orelse (ch = " ") | |
| 377 | ||
| 378 | fun proofstring x = let val exp = explode x | |
| 379 | in | |
| 380 | List.filter (is_proof_char ) exp | |
| 381 | end | |
| 382 | ||
| 16061 | 383 | fun replace_clause_strs [] recons newrecons = (newrecons) | 
| 384 | | replace_clause_strs ((thmnum, thm)::thml) | |
| 385 | ((clausenum,(step,clause_strs,thmvars))::recons) newrecons = | |
| 386 | let val new_clause_strs = ReconOrderClauses.get_meta_lits_bracket thm | |
| 387 | in | |
| 388 | replace_clause_strs thml recons | |
| 389 | ((clausenum,(step,new_clause_strs,thmvars))::newrecons) | |
| 390 | end | |
| 15642 | 391 | |
| 392 | ||
| 393 | fun follow clauses [] allvars thml recons = | |
| 16061 | 394 | let | 
| 395 | val new_recons = replace_clause_strs thml recons [] | |
| 396 | in | |
| 397 | (#2(hd thml), new_recons) | |
| 398 | end | |
| 399 | | follow clauses (h::t) allvars thml recons = | |
| 400 | let | |
| 401 | val (thml', recons') = follow_line clauses allvars thml h recons | |
| 402 | val (thm, recons_list) = follow clauses t allvars thml' recons' | |
| 403 | in | |
| 404 | (thm,recons_list) | |
| 405 | end | |
| 15642 | 406 | |
| 407 | (* Assume we have the cnf clauses as a list of (clauseno, clause) *) | |
| 408 | (* and the proof as a list of the proper datatype *) | |
| 409 | (* take the cnf clauses of the goal and the proof from the res. prover *) | |
| 410 | (* as a list of type Proofstep and return the thm goal ==> False *) | |
| 411 | ||
| 412 | (* takes original axioms, proof_steps parsed from spass, variables *) | |
| 413 | ||
| 414 | fun translate_proof clauses proof allvars | |
| 415 | = let val (thm, recons) = follow clauses proof allvars [] [] | |
| 416 | in | |
| 417 | (thm, (recons)) | |
| 418 | end | |
| 419 | ||
| 16803 | 420 | end; |