author | wenzelm |
Thu, 18 Jun 1998 10:52:34 +0200 | |
changeset 5047 | 585fa380df1a |
parent 4663 | 27fa93c22b9b |
child 5317 | 3a9214482762 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/ex/meson |
969 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
969 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
The MESON resolution proof procedure for HOL |
|
7 |
||
8 |
When making clauses, avoids using the rewriter -- instead uses RS recursively |
|
1599
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
9 |
|
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
10 |
NEED TO SORT LITERALS BY # OF VARS, USING ==>I/E. ELIMINATES NEED FOR |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
11 |
FUNCTION nodups -- if done to goal clauses too! |
969 | 12 |
*) |
13 |
||
14 |
writeln"File HOL/ex/meson."; |
|
15 |
||
16 |
(*Prove theorems using fast_tac*) |
|
17 |
fun prove_fun s = |
|
18 |
prove_goal HOL.thy s |
|
1820 | 19 |
(fn prems => [ cut_facts_tac prems 1, Fast_tac 1 ]); |
969 | 20 |
|
21 |
(**** Negation Normal Form ****) |
|
22 |
||
23 |
(*** de Morgan laws ***) |
|
24 |
||
25 |
val not_conjD = prove_fun "~(P&Q) ==> ~P | ~Q"; |
|
26 |
val not_disjD = prove_fun "~(P|Q) ==> ~P & ~Q"; |
|
27 |
val not_notD = prove_fun "~~P ==> P"; |
|
3842 | 28 |
val not_allD = prove_fun "~(! x. P(x)) ==> ? x. ~P(x)"; |
29 |
val not_exD = prove_fun "~(? x. P(x)) ==> ! x. ~P(x)"; |
|
969 | 30 |
|
31 |
||
32 |
(*** Removal of --> and <-> (positive and negative occurrences) ***) |
|
33 |
||
34 |
val imp_to_disjD = prove_fun "P-->Q ==> ~P | Q"; |
|
35 |
val not_impD = prove_fun "~(P-->Q) ==> P & ~Q"; |
|
36 |
||
37 |
val iff_to_disjD = prove_fun "P=Q ==> (~P | Q) & (~Q | P)"; |
|
38 |
||
39 |
(*Much more efficient than (P & ~Q) | (Q & ~P) for computing CNF*) |
|
40 |
val not_iffD = prove_fun "~(P=Q) ==> (P | Q) & (~P | ~Q)"; |
|
41 |
||
42 |
||
43 |
(**** Pulling out the existential quantifiers ****) |
|
44 |
||
45 |
(*** Conjunction ***) |
|
46 |
||
3842 | 47 |
val conj_exD1 = prove_fun "(? x. P(x)) & Q ==> ? x. P(x) & Q"; |
48 |
val conj_exD2 = prove_fun "P & (? x. Q(x)) ==> ? x. P & Q(x)"; |
|
969 | 49 |
|
50 |
(*** Disjunction ***) |
|
51 |
||
52 |
(*DO NOT USE with forall-Skolemization: makes fewer schematic variables!! |
|
53 |
With ex-Skolemization, makes fewer Skolem constants*) |
|
3842 | 54 |
val disj_exD = prove_fun "(? x. P(x)) | (? x. Q(x)) ==> ? x. P(x) | Q(x)"; |
969 | 55 |
|
3842 | 56 |
val disj_exD1 = prove_fun "(? x. P(x)) | Q ==> ? x. P(x) | Q"; |
57 |
val disj_exD2 = prove_fun "P | (? x. Q(x)) ==> ? x. P | Q(x)"; |
|
969 | 58 |
|
59 |
||
60 |
||
61 |
(***** Generating clauses for the Meson Proof Procedure *****) |
|
62 |
||
63 |
(*** Disjunctions ***) |
|
64 |
||
65 |
val disj_assoc = prove_fun "(P|Q)|R ==> P|(Q|R)"; |
|
66 |
||
67 |
val disj_comm = prove_fun "P|Q ==> Q|P"; |
|
68 |
||
69 |
val disj_FalseD1 = prove_fun "False|P ==> P"; |
|
70 |
val disj_FalseD2 = prove_fun "P|False ==> P"; |
|
71 |
||
72 |
(*** Generation of contrapositives ***) |
|
73 |
||
74 |
(*Inserts negated disjunct after removing the negation; P is a literal*) |
|
75 |
val [major,minor] = goal HOL.thy "~P|Q ==> ((~P==>P) ==> Q)"; |
|
76 |
by (rtac (major RS disjE) 1); |
|
77 |
by (rtac notE 1); |
|
78 |
by (etac minor 2); |
|
79 |
by (ALLGOALS assume_tac); |
|
80 |
qed "make_neg_rule"; |
|
81 |
||
82 |
(*For Plaisted's "Postive refinement" of the MESON procedure*) |
|
83 |
val [major,minor] = goal HOL.thy "~P|Q ==> (P ==> Q)"; |
|
84 |
by (rtac (major RS disjE) 1); |
|
85 |
by (rtac notE 1); |
|
86 |
by (rtac minor 2); |
|
87 |
by (ALLGOALS assume_tac); |
|
88 |
qed "make_refined_neg_rule"; |
|
89 |
||
90 |
(*P should be a literal*) |
|
91 |
val [major,minor] = goal HOL.thy "P|Q ==> ((P==>~P) ==> Q)"; |
|
92 |
by (rtac (major RS disjE) 1); |
|
93 |
by (rtac notE 1); |
|
94 |
by (etac minor 1); |
|
95 |
by (ALLGOALS assume_tac); |
|
96 |
qed "make_pos_rule"; |
|
97 |
||
98 |
(*** Generation of a goal clause -- put away the final literal ***) |
|
99 |
||
100 |
val [major,minor] = goal HOL.thy "~P ==> ((~P==>P) ==> False)"; |
|
101 |
by (rtac notE 1); |
|
102 |
by (rtac minor 2); |
|
103 |
by (ALLGOALS (rtac major)); |
|
104 |
qed "make_neg_goal"; |
|
105 |
||
106 |
val [major,minor] = goal HOL.thy "P ==> ((P==>~P) ==> False)"; |
|
107 |
by (rtac notE 1); |
|
108 |
by (rtac minor 1); |
|
109 |
by (ALLGOALS (rtac major)); |
|
110 |
qed "make_pos_goal"; |
|
111 |
||
112 |
||
113 |
(**** Lemmas for forward proof (like congruence rules) ****) |
|
114 |
||
115 |
(*NOTE: could handle conjunctions (faster?) by |
|
116 |
nf(th RS conjunct2) RS (nf(th RS conjunct1) RS conjI) *) |
|
117 |
val major::prems = goal HOL.thy |
|
118 |
"[| P'&Q'; P' ==> P; Q' ==> Q |] ==> P&Q"; |
|
119 |
by (rtac (major RS conjE) 1); |
|
120 |
by (rtac conjI 1); |
|
121 |
by (ALLGOALS (eresolve_tac prems)); |
|
122 |
qed "conj_forward"; |
|
123 |
||
124 |
val major::prems = goal HOL.thy |
|
125 |
"[| P'|Q'; P' ==> P; Q' ==> Q |] ==> P|Q"; |
|
126 |
by (rtac (major RS disjE) 1); |
|
127 |
by (ALLGOALS (dresolve_tac prems)); |
|
128 |
by (ALLGOALS (eresolve_tac [disjI1,disjI2])); |
|
129 |
qed "disj_forward"; |
|
130 |
||
131 |
val major::prems = goal HOL.thy |
|
132 |
"[| ! x. P'(x); !!x. P'(x) ==> P(x) |] ==> ! x. P(x)"; |
|
133 |
by (rtac allI 1); |
|
134 |
by (resolve_tac prems 1); |
|
135 |
by (rtac (major RS spec) 1); |
|
136 |
qed "all_forward"; |
|
137 |
||
138 |
val major::prems = goal HOL.thy |
|
139 |
"[| ? x. P'(x); !!x. P'(x) ==> P(x) |] ==> ? x. P(x)"; |
|
140 |
by (rtac (major RS exE) 1); |
|
141 |
by (rtac exI 1); |
|
142 |
by (eresolve_tac prems 1); |
|
143 |
qed "ex_forward"; |
|
144 |
||
145 |
||
146 |
(**** Operators for forward proof ****) |
|
147 |
||
148 |
(*raises exception if no rules apply -- unlike RL*) |
|
149 |
fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls)) |
|
150 |
| tryres (th, []) = raise THM("tryres", 0, [th]); |
|
151 |
||
152 |
val prop_of = #prop o rep_thm; |
|
153 |
||
154 |
(*Permits forward proof from rules that discharge assumptions*) |
|
1512 | 155 |
fun forward_res nf st = |
4271
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents:
4153
diff
changeset
|
156 |
case Seq.pull (ALLGOALS (METAHYPS (fn [prem] => rtac (nf prem) 1)) st) |
969 | 157 |
of Some(th,_) => th |
1512 | 158 |
| None => raise THM("forward_res", 0, [st]); |
969 | 159 |
|
160 |
||
161 |
(*Negation Normal Form*) |
|
162 |
val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, |
|
1465 | 163 |
not_impD, not_iffD, not_allD, not_exD, not_notD]; |
969 | 164 |
fun make_nnf th = make_nnf (tryres(th, nnf_rls)) |
165 |
handle THM _ => |
|
1465 | 166 |
forward_res make_nnf |
167 |
(tryres(th, [conj_forward,disj_forward,all_forward,ex_forward])) |
|
969 | 168 |
handle THM _ => th; |
169 |
||
170 |
||
171 |
(*Are any of the constants in "bs" present in the term?*) |
|
172 |
fun has_consts bs = |
|
173 |
let fun has (Const(a,_)) = a mem bs |
|
1465 | 174 |
| has (f$u) = has f orelse has u |
175 |
| has (Abs(_,_,t)) = has t |
|
176 |
| has _ = false |
|
969 | 177 |
in has end; |
178 |
||
179 |
(*Pull existential quantifiers (Skolemization)*) |
|
180 |
fun skolemize th = |
|
181 |
if not (has_consts ["Ex"] (prop_of th)) then th |
|
182 |
else skolemize (tryres(th, [choice, conj_exD1, conj_exD2, |
|
2031 | 183 |
disj_exD, disj_exD1, disj_exD2])) |
969 | 184 |
handle THM _ => |
1465 | 185 |
skolemize (forward_res skolemize |
2031 | 186 |
(tryres (th, [conj_forward, disj_forward, all_forward]))) |
969 | 187 |
handle THM _ => forward_res skolemize (th RS ex_forward); |
188 |
||
189 |
||
190 |
(**** Clause handling ****) |
|
191 |
||
192 |
fun literals (Const("Trueprop",_) $ P) = literals P |
|
193 |
| literals (Const("op |",_) $ P $ Q) = literals P @ literals Q |
|
2720 | 194 |
| literals (Const("Not",_) $ P) = [(false,P)] |
969 | 195 |
| literals P = [(true,P)]; |
196 |
||
197 |
(*number of literals in a term*) |
|
198 |
val nliterals = length o literals; |
|
199 |
||
1585 | 200 |
(*to detect, and remove, tautologous clauses*) |
969 | 201 |
fun taut_lits [] = false |
202 |
| taut_lits ((flg,t)::ts) = (not flg,t) mem ts orelse taut_lits ts; |
|
203 |
||
1585 | 204 |
val term_False = term_of (read_cterm (sign_of HOL.thy) |
2031 | 205 |
("False", Type("bool",[]))); |
969 | 206 |
|
1585 | 207 |
(*Include False as a literal: an occurrence of ~False is a tautology*) |
208 |
fun is_taut th = taut_lits ((true,term_False) :: literals (prop_of th)); |
|
969 | 209 |
|
210 |
(*Generation of unique names -- maxidx cannot be relied upon to increase! |
|
211 |
Cannot rely on "variant", since variables might coincide when literals |
|
212 |
are joined to make a clause... |
|
213 |
19 chooses "U" as the first variable name*) |
|
214 |
val name_ref = ref 19; |
|
215 |
||
216 |
(*Replaces universally quantified variables by FREE variables -- because |
|
217 |
assumptions may not contain scheme variables. Later, call "generalize". *) |
|
218 |
fun freeze_spec th = |
|
219 |
let val sth = th RS spec |
|
220 |
val newname = (name_ref := !name_ref + 1; |
|
221 |
radixstring(26, "A", !name_ref)) |
|
222 |
in read_instantiate [("x", newname)] sth end; |
|
223 |
||
224 |
fun resop nf [prem] = resolve_tac (nf prem) 1; |
|
225 |
||
226 |
(*Conjunctive normal form, detecting tautologies early. |
|
227 |
Strips universal quantifiers and breaks up conjunctions. *) |
|
228 |
fun cnf_aux seen (th,ths) = |
|
229 |
if taut_lits (literals(prop_of th) @ seen) then ths |
|
230 |
else if not (has_consts ["All","op &"] (prop_of th)) then th::ths |
|
231 |
else (*conjunction?*) |
|
232 |
cnf_aux seen (th RS conjunct1, |
|
1465 | 233 |
cnf_aux seen (th RS conjunct2, ths)) |
969 | 234 |
handle THM _ => (*universal quant?*) |
1465 | 235 |
cnf_aux seen (freeze_spec th, ths) |
969 | 236 |
handle THM _ => (*disjunction?*) |
237 |
let val tac = |
|
1465 | 238 |
(METAHYPS (resop (cnf_nil seen)) 1) THEN |
3537 | 239 |
(fn st' => st' |> |
240 |
METAHYPS (resop (cnf_nil (literals (concl_of st') @ seen))) 1) |
|
4271
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents:
4153
diff
changeset
|
241 |
in Seq.list_of (tac (th RS disj_forward)) @ ths end |
969 | 242 |
and cnf_nil seen th = cnf_aux seen (th,[]); |
243 |
||
244 |
(*Top-level call to cnf -- it's safe to reset name_ref*) |
|
245 |
fun cnf (th,ths) = |
|
246 |
(name_ref := 19; cnf (th RS conjunct1, cnf (th RS conjunct2, ths)) |
|
247 |
handle THM _ => (*not a conjunction*) cnf_aux [] (th, ths)); |
|
248 |
||
249 |
(**** Removal of duplicate literals ****) |
|
250 |
||
251 |
(*Version for removal of duplicate literals*) |
|
252 |
val major::prems = goal HOL.thy |
|
253 |
"[| P'|Q'; P' ==> P; [| Q'; P==>False |] ==> Q |] ==> P|Q"; |
|
254 |
by (rtac (major RS disjE) 1); |
|
255 |
by (rtac disjI1 1); |
|
256 |
by (rtac (disjCI RS disj_comm) 2); |
|
257 |
by (ALLGOALS (eresolve_tac prems)); |
|
258 |
by (etac notE 1); |
|
259 |
by (assume_tac 1); |
|
260 |
qed "disj_forward2"; |
|
261 |
||
262 |
(*Forward proof, passing extra assumptions as theorems to the tactic*) |
|
1512 | 263 |
fun forward_res2 nf hyps st = |
4271
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents:
4153
diff
changeset
|
264 |
case Seq.pull |
1512 | 265 |
(REPEAT |
2031 | 266 |
(METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1) |
267 |
st) |
|
969 | 268 |
of Some(th,_) => th |
1512 | 269 |
| None => raise THM("forward_res2", 0, [st]); |
969 | 270 |
|
271 |
(*Remove duplicates in P|Q by assuming ~P in Q |
|
272 |
rls (initially []) accumulates assumptions of the form P==>False*) |
|
273 |
fun nodups_aux rls th = nodups_aux rls (th RS disj_assoc) |
|
274 |
handle THM _ => tryres(th,rls) |
|
275 |
handle THM _ => tryres(forward_res2 nodups_aux rls (th RS disj_forward2), |
|
1465 | 276 |
[disj_FalseD1, disj_FalseD2, asm_rl]) |
969 | 277 |
handle THM _ => th; |
278 |
||
279 |
(*Remove duplicate literals, if there are any*) |
|
280 |
fun nodups th = |
|
281 |
if null(findrep(literals(prop_of th))) then th |
|
282 |
else nodups_aux [] th; |
|
283 |
||
284 |
||
285 |
(**** Generation of contrapositives ****) |
|
286 |
||
287 |
(*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) |
|
288 |
fun assoc_right th = assoc_right (th RS disj_assoc) |
|
1465 | 289 |
handle THM _ => th; |
969 | 290 |
|
291 |
(*Must check for negative literal first!*) |
|
292 |
val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; |
|
1585 | 293 |
|
294 |
(*For Plaisted's postive refinement. [currently unused] *) |
|
969 | 295 |
val refined_clause_rules = [disj_assoc, make_refined_neg_rule, make_pos_rule]; |
296 |
||
297 |
(*Create a goal or support clause, conclusing False*) |
|
298 |
fun make_goal th = (*Must check for negative literal first!*) |
|
299 |
make_goal (tryres(th, clause_rules)) |
|
300 |
handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); |
|
301 |
||
302 |
(*Sort clauses by number of literals*) |
|
303 |
fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2); |
|
304 |
||
305 |
(*TAUTOLOGY CHECK SHOULD NOT BE NECESSARY!*) |
|
4448 | 306 |
fun sort_clauses ths = sort (make_ord fewerlits) (filter (not o is_taut) ths); |
969 | 307 |
|
308 |
(*Convert all suitable free variables to schematic variables*) |
|
309 |
fun generalize th = forall_elim_vars 0 (forall_intr_frees th); |
|
310 |
||
1764 | 311 |
(*Make clauses from a list of theorems, previously Skolemized and put into nnf. |
312 |
The resulting clauses are HOL disjunctions.*) |
|
969 | 313 |
fun make_clauses ths = |
314 |
sort_clauses (map (generalize o nodups) (foldr cnf (ths,[]))); |
|
315 |
||
1599
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
316 |
(*Create a meta-level Horn clause*) |
969 | 317 |
fun make_horn crules th = make_horn crules (tryres(th,crules)) |
1465 | 318 |
handle THM _ => th; |
969 | 319 |
|
320 |
(*Generate Horn clauses for all contrapositives of a clause*) |
|
321 |
fun add_contras crules (th,hcs) = |
|
322 |
let fun rots (0,th) = hcs |
|
1465 | 323 |
| rots (k,th) = zero_var_indexes (make_horn crules th) :: |
324 |
rots(k-1, assoc_right (th RS disj_comm)) |
|
969 | 325 |
in case nliterals(prop_of th) of |
1465 | 326 |
1 => th::hcs |
969 | 327 |
| n => rots(n, assoc_right th) |
328 |
end; |
|
329 |
||
1599
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
330 |
(*Use "theorem naming" to label the clauses*) |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
331 |
fun name_thms label = |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
332 |
let fun name1 (th, (k,ths)) = |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
333 |
(k-1, Thm.name_thm (label ^ string_of_int k, th) :: ths) |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
334 |
|
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
335 |
in fn ths => #2 (foldr name1 (ths, (length ths, []))) end; |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
336 |
|
969 | 337 |
(*Convert a list of clauses to (contrapositive) Horn clauses*) |
1585 | 338 |
fun make_horns ths = |
1599
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
339 |
name_thms "Horn#" |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
340 |
(gen_distinct eq_thm (foldr (add_contras clause_rules) (ths,[]))); |
969 | 341 |
|
342 |
(*Find an all-negative support clause*) |
|
343 |
fun is_negative th = forall (not o #1) (literals (prop_of th)); |
|
344 |
||
345 |
val neg_clauses = filter is_negative; |
|
346 |
||
347 |
||
348 |
(***** MESON PROOF PROCEDURE *****) |
|
349 |
||
350 |
fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi, |
|
1465 | 351 |
As) = rhyps(phi, A::As) |
969 | 352 |
| rhyps (_, As) = As; |
353 |
||
354 |
(** Detecting repeated assumptions in a subgoal **) |
|
355 |
||
356 |
(*The stringtree detects repeated assumptions.*) |
|
357 |
fun ins_term (net,t) = Net.insert_term((t,t), net, op aconv); |
|
358 |
||
359 |
(*detects repetitions in a list of terms*) |
|
360 |
fun has_reps [] = false |
|
361 |
| has_reps [_] = false |
|
362 |
| has_reps [t,u] = (t aconv u) |
|
363 |
| has_reps ts = (foldl ins_term (Net.empty, ts); false) |
|
1465 | 364 |
handle INSERT => true; |
969 | 365 |
|
1585 | 366 |
(*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) |
4271
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents:
4153
diff
changeset
|
367 |
fun TRYALL_eq_assume_tac 0 st = Seq.single st |
1585 | 368 |
| TRYALL_eq_assume_tac i st = TRYALL_eq_assume_tac (i-1) (eq_assumption i st) |
369 |
handle THM _ => TRYALL_eq_assume_tac (i-1) st; |
|
370 |
||
969 | 371 |
(*Loop checking: FAIL if trying to prove the same thing twice |
1585 | 372 |
-- if *ANY* subgoal has repeated literals*) |
373 |
fun check_tac st = |
|
374 |
if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st) |
|
4271
3a82492e70c5
changed Pure/Sequence interface -- isatool fixseq;
wenzelm
parents:
4153
diff
changeset
|
375 |
then Seq.empty else Seq.single st; |
1585 | 376 |
|
969 | 377 |
|
378 |
(* net_resolve_tac actually made it slower... *) |
|
379 |
fun prolog_step_tac horns i = |
|
1585 | 380 |
(assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN |
381 |
TRYALL eq_assume_tac; |
|
969 | 382 |
|
383 |
||
384 |
(*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) |
|
385 |
local fun addconcl(prem,sz) = size_of_term (Logic.strip_assums_concl prem) + sz |
|
386 |
in |
|
387 |
fun size_of_subgoals st = foldr addconcl (prems_of st, 0) |
|
388 |
end; |
|
389 |
||
390 |
(*Could simply use nprems_of, which would count remaining subgoals -- no |
|
391 |
discrimination as to their size! With BEST_FIRST, fails for problem 41.*) |
|
392 |
||
393 |
fun best_prolog_tac sizef horns = |
|
394 |
BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1); |
|
395 |
||
396 |
fun depth_prolog_tac horns = |
|
397 |
DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1); |
|
398 |
||
399 |
(*Return all negative clauses, as possible goal clauses*) |
|
1599
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
400 |
fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls)); |
969 | 401 |
|
402 |
||
403 |
fun skolemize_tac prems = |
|
404 |
cut_facts_tac (map (skolemize o make_nnf) prems) THEN' |
|
405 |
REPEAT o (etac exE); |
|
406 |
||
1599
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
407 |
(*Shell of all meson-tactics. Supplies cltac with clauses: HOL disjunctions*) |
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
408 |
fun MESON cltac = SELECT_GOAL |
969 | 409 |
(EVERY1 [rtac ccontr, |
1465 | 410 |
METAHYPS (fn negs => |
411 |
EVERY1 [skolemize_tac negs, |
|
1599
b11ac7072422
Now labels the Horn and goal clauses to make the proof
paulson
parents:
1585
diff
changeset
|
412 |
METAHYPS (cltac o make_clauses)])]); |
969 | 413 |
|
1585 | 414 |
(** Best-first search versions **) |
415 |
||
969 | 416 |
fun best_meson_tac sizef = |
417 |
MESON (fn cls => |
|
1585 | 418 |
THEN_BEST_FIRST (resolve_tac (gocls cls) 1) |
419 |
(has_fewer_prems 1, sizef) |
|
2031 | 420 |
(prolog_step_tac (make_horns cls) 1)); |
969 | 421 |
|
422 |
(*First, breaks the goal into independent units*) |
|
1585 | 423 |
val safe_best_meson_tac = |
4153 | 424 |
SELECT_GOAL (TRY Safe_tac THEN |
1465 | 425 |
TRYALL (best_meson_tac size_of_subgoals)); |
969 | 426 |
|
1585 | 427 |
(** Depth-first search version **) |
428 |
||
969 | 429 |
val depth_meson_tac = |
430 |
MESON (fn cls => EVERY [resolve_tac (gocls cls) 1, |
|
1465 | 431 |
depth_prolog_tac (make_horns cls)]); |
969 | 432 |
|
1585 | 433 |
|
434 |
||
435 |
(** Iterative deepening version **) |
|
436 |
||
437 |
(*This version does only one inference per call; |
|
438 |
having only one eq_assume_tac speeds it up!*) |
|
439 |
fun prolog_step_tac' horns = |
|
440 |
let val (horn0s, hornps) = (*0 subgoals vs 1 or more*) |
|
441 |
take_prefix (fn rl => nprems_of rl=0) horns |
|
442 |
val nrtac = net_resolve_tac horns |
|
443 |
in fn i => eq_assume_tac i ORELSE |
|
444 |
match_tac horn0s i ORELSE (*no backtracking if unit MATCHES*) |
|
2031 | 445 |
((assume_tac i APPEND nrtac i) THEN check_tac) |
1585 | 446 |
end; |
447 |
||
448 |
fun iter_deepen_prolog_tac horns = |
|
449 |
ITER_DEEPEN (has_fewer_prems 1) (prolog_step_tac' horns); |
|
450 |
||
451 |
val iter_deepen_meson_tac = |
|
452 |
MESON (fn cls => |
|
453 |
(THEN_ITER_DEEPEN (resolve_tac (gocls cls) 1) |
|
2031 | 454 |
(has_fewer_prems 1) |
455 |
(prolog_step_tac' (make_horns cls)))); |
|
1585 | 456 |
|
457 |
val safe_meson_tac = |
|
4153 | 458 |
SELECT_GOAL (TRY Safe_tac THEN |
1585 | 459 |
TRYALL (iter_deepen_meson_tac)); |
460 |
||
461 |
||
969 | 462 |
writeln"Reached end of file."; |