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