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