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