| author | nipkow | 
| Wed, 26 Jun 2002 11:07:14 +0200 | |
| changeset 13249 | 4b3de6370184 | 
| parent 13105 | 3d1e7a199bdc | 
| child 14733 | 3eda95792083 | 
| permissions | -rw-r--r-- | 
| 9869 | 1 | (* Title: HOL/Tools/meson.ML | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 4 | Copyright 1992 University of Cambridge | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 5 | |
| 9869 | 6 | The MESON resolution proof procedure for HOL. | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 7 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 8 | When making clauses, avoids using the rewriter -- instead uses RS recursively | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 9 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 10 | NEED TO SORT LITERALS BY # OF VARS, USING ==>I/E. ELIMINATES NEED FOR | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 11 | FUNCTION nodups -- if done to goal clauses too! | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 12 | *) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 13 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 14 | local | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 15 | |
| 12299 | 16 | val not_conjD = thm "meson_not_conjD"; | 
| 17 | val not_disjD = thm "meson_not_disjD"; | |
| 18 | val not_notD = thm "meson_not_notD"; | |
| 19 | val not_allD = thm "meson_not_allD"; | |
| 20 | val not_exD = thm "meson_not_exD"; | |
| 21 | val imp_to_disjD = thm "meson_imp_to_disjD"; | |
| 22 | val not_impD = thm "meson_not_impD"; | |
| 23 | val iff_to_disjD = thm "meson_iff_to_disjD"; | |
| 24 | val not_iffD = thm "meson_not_iffD"; | |
| 25 | val conj_exD1 = thm "meson_conj_exD1"; | |
| 26 | val conj_exD2 = thm "meson_conj_exD2"; | |
| 27 | val disj_exD = thm "meson_disj_exD"; | |
| 28 | val disj_exD1 = thm "meson_disj_exD1"; | |
| 29 | val disj_exD2 = thm "meson_disj_exD2"; | |
| 30 | val disj_assoc = thm "meson_disj_assoc"; | |
| 31 | val disj_comm = thm "meson_disj_comm"; | |
| 32 | val disj_FalseD1 = thm "meson_disj_FalseD1"; | |
| 33 | val disj_FalseD2 = thm "meson_disj_FalseD2"; | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 34 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 35 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 36 | (**** Operators for forward proof ****) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 37 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 38 | (*raises exception if no rules apply -- unlike RL*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 39 | fun tryres (th, rl::rls) = (th RS rl handle THM _ => tryres(th,rls)) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 40 |    | tryres (th, []) = raise THM("tryres", 0, [th]);
 | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 41 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 42 | val prop_of = #prop o rep_thm; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 43 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 44 | (*Permits forward proof from rules that discharge assumptions*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 45 | fun forward_res nf st = | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 46 | case Seq.pull (ALLGOALS (METAHYPS (fn [prem] => rtac (nf prem) 1)) st) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 47 | of Some(th,_) => th | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 48 |     | None => raise THM("forward_res", 0, [st]);
 | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 49 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 50 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 51 | (*Are any of the constants in "bs" present in the term?*) | 
| 9869 | 52 | fun has_consts bs = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 53 | let fun has (Const(a,_)) = a mem bs | 
| 9869 | 54 | | has (f$u) = has f orelse has u | 
| 55 | | has (Abs(_,_,t)) = has t | |
| 56 | | has _ = false | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 57 | in has end; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 58 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 59 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 60 | (**** Clause handling ****) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 61 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 62 |  fun literals (Const("Trueprop",_) $ P) = literals P
 | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 63 |    | literals (Const("op |",_) $ P $ Q) = literals P @ literals Q
 | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 64 |    | literals (Const("Not",_) $ P) = [(false,P)]
 | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 65 | | literals P = [(true,P)]; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 66 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 67 | (*number of literals in a term*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 68 | val nliterals = length o literals; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 69 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 70 | (*to detect, and remove, tautologous clauses*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 71 | fun taut_lits [] = false | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 72 | | taut_lits ((flg,t)::ts) = (not flg,t) mem ts orelse taut_lits ts; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 73 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 74 | (*Include False as a literal: an occurrence of ~False is a tautology*) | 
| 9869 | 75 | fun is_taut th = taut_lits ((true, HOLogic.false_const) :: | 
| 76 | literals (prop_of th)); | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 77 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 78 | (*Generation of unique names -- maxidx cannot be relied upon to increase! | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 79 | Cannot rely on "variant", since variables might coincide when literals | 
| 9869 | 80 | are joined to make a clause... | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 81 | 19 chooses "U" as the first variable name*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 82 | val name_ref = ref 19; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 83 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 84 | (*Replaces universally quantified variables by FREE variables -- because | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 85 | assumptions may not contain scheme variables. Later, call "generalize". *) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 86 | fun freeze_spec th = | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 87 | let val sth = th RS spec | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 88 | val newname = (name_ref := !name_ref + 1; | 
| 9869 | 89 | radixstring(26, "A", !name_ref)) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 90 |    in  read_instantiate [("x", newname)] sth  end;
 | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 91 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 92 | fun resop nf [prem] = resolve_tac (nf prem) 1; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 93 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 94 | (*Conjunctive normal form, detecting tautologies early. | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 95 | Strips universal quantifiers and breaks up conjunctions. *) | 
| 9869 | 96 | fun cnf_aux seen (th,ths) = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 97 | if taut_lits (literals(prop_of th) @ seen) then ths | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 98 | else if not (has_consts ["All","op &"] (prop_of th)) then th::ths | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 99 | else (*conjunction?*) | 
| 9869 | 100 | cnf_aux seen (th RS conjunct1, | 
| 101 | cnf_aux seen (th RS conjunct2, ths)) | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 102 | handle THM _ => (*universal quant?*) | 
| 9869 | 103 | cnf_aux seen (freeze_spec th, ths) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 104 | handle THM _ => (*disjunction?*) | 
| 9869 | 105 | let val tac = | 
| 106 | (METAHYPS (resop (cnf_nil seen)) 1) THEN | |
| 107 | (fn st' => st' |> | |
| 108 | METAHYPS (resop (cnf_nil (literals (concl_of st') @ seen))) 1) | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 109 | in Seq.list_of (tac (th RS disj_forward)) @ ths end | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 110 | and cnf_nil seen th = cnf_aux seen (th,[]); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 111 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 112 | (*Top-level call to cnf -- it's safe to reset name_ref*) | 
| 9869 | 113 | fun cnf (th,ths) = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 114 | (name_ref := 19; cnf (th RS conjunct1, cnf (th RS conjunct2, ths)) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 115 | handle THM _ => (*not a conjunction*) cnf_aux [] (th, ths)); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 116 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 117 | (**** Removal of duplicate literals ****) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 118 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 119 | (*Forward proof, passing extra assumptions as theorems to the tactic*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 120 | fun forward_res2 nf hyps st = | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 121 | case Seq.pull | 
| 9869 | 122 | (REPEAT | 
| 123 | (METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1) | |
| 124 | st) | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 125 | of Some(th,_) => th | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 126 |     | None => raise THM("forward_res2", 0, [st]);
 | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 127 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 128 | (*Remove duplicates in P|Q by assuming ~P in Q | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 129 | rls (initially []) accumulates assumptions of the form P==>False*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 130 | fun nodups_aux rls th = nodups_aux rls (th RS disj_assoc) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 131 | handle THM _ => tryres(th,rls) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 132 | handle THM _ => tryres(forward_res2 nodups_aux rls (th RS disj_forward2), | 
| 9869 | 133 | [disj_FalseD1, disj_FalseD2, asm_rl]) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 134 | handle THM _ => th; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 135 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 136 | (*Remove duplicate literals, if there are any*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 137 | fun nodups th = | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 138 | if null(findrep(literals(prop_of th))) then th | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 139 | else nodups_aux [] th; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 140 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 141 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 142 | (**** Generation of contrapositives ****) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 143 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 144 | (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 145 | fun assoc_right th = assoc_right (th RS disj_assoc) | 
| 9869 | 146 | handle THM _ => th; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 147 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 148 | (*Must check for negative literal first!*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 149 | val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 150 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 151 | (*For Plaisted's postive refinement. [currently unused] *) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 152 | val refined_clause_rules = [disj_assoc, make_refined_neg_rule, make_pos_rule]; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 153 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 154 | (*Create a goal or support clause, conclusing False*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 155 | fun make_goal th = (*Must check for negative literal first!*) | 
| 9869 | 156 | make_goal (tryres(th, clause_rules)) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 157 | handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 158 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 159 | (*Sort clauses by number of literals*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 160 | fun fewerlits(th1,th2) = nliterals(prop_of th1) < nliterals(prop_of th2); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 161 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 162 | (*TAUTOLOGY CHECK SHOULD NOT BE NECESSARY!*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 163 | fun sort_clauses ths = sort (make_ord fewerlits) (filter (not o is_taut) ths); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 164 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 165 | (*Convert all suitable free variables to schematic variables*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 166 | fun generalize th = forall_elim_vars 0 (forall_intr_frees th); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 167 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 168 | (*Create a meta-level Horn clause*) | 
| 9869 | 169 | fun make_horn crules th = make_horn crules (tryres(th,crules)) | 
| 170 | handle THM _ => th; | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 171 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 172 | (*Generate Horn clauses for all contrapositives of a clause*) | 
| 9869 | 173 | fun add_contras crules (th,hcs) = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 174 | let fun rots (0,th) = hcs | 
| 9869 | 175 | | rots (k,th) = zero_var_indexes (make_horn crules th) :: | 
| 176 | rots(k-1, assoc_right (th RS disj_comm)) | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 177 | in case nliterals(prop_of th) of | 
| 9869 | 178 | 1 => th::hcs | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 179 | | n => rots(n, assoc_right th) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 180 | end; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 181 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 182 | (*Use "theorem naming" to label the clauses*) | 
| 9869 | 183 | fun name_thms label = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 184 | let fun name1 (th, (k,ths)) = | 
| 9869 | 185 | (k-1, Thm.name_thm (label ^ string_of_int k, th) :: ths) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 186 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 187 | in fn ths => #2 (foldr name1 (ths, (length ths, []))) end; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 188 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 189 | (*Find an all-negative support clause*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 190 | fun is_negative th = forall (not o #1) (literals (prop_of th)); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 191 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 192 | val neg_clauses = filter is_negative; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 193 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 194 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 195 | (***** MESON PROOF PROCEDURE *****) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 196 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 197 |  fun rhyps (Const("==>",_) $ (Const("Trueprop",_) $ A) $ phi,
 | 
| 9869 | 198 | As) = rhyps(phi, A::As) | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 199 | | rhyps (_, As) = As; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 200 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 201 | (** Detecting repeated assumptions in a subgoal **) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 202 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 203 | (*The stringtree detects repeated assumptions.*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 204 | fun ins_term (net,t) = Net.insert_term((t,t), net, op aconv); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 205 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 206 | (*detects repetitions in a list of terms*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 207 | fun has_reps [] = false | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 208 | | has_reps [_] = false | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 209 | | has_reps [t,u] = (t aconv u) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 210 | | has_reps ts = (foldl ins_term (Net.empty, ts); false) | 
| 9869 | 211 | handle INSERT => true; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 212 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 213 | (*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 214 | fun TRYALL_eq_assume_tac 0 st = Seq.single st | 
| 9869 | 215 | | TRYALL_eq_assume_tac i st = | 
| 216 | TRYALL_eq_assume_tac (i-1) (eq_assumption i st) | |
| 217 | handle THM _ => TRYALL_eq_assume_tac (i-1) st; | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 218 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 219 | (*Loop checking: FAIL if trying to prove the same thing twice | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 220 | -- if *ANY* subgoal has repeated literals*) | 
| 9869 | 221 | fun check_tac st = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 222 | if exists (fn prem => has_reps (rhyps(prem,[]))) (prems_of st) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 223 | then Seq.empty else Seq.single st; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 224 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 225 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 226 | (* net_resolve_tac actually made it slower... *) | 
| 9869 | 227 | fun prolog_step_tac horns i = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 228 | (assume_tac i APPEND resolve_tac horns i) THEN check_tac THEN | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 229 | TRYALL eq_assume_tac; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 230 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 231 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 232 | in | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 233 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 234 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 235 | (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 236 | local fun addconcl(prem,sz) = size_of_term(Logic.strip_assums_concl prem) + sz | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 237 | in | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 238 | fun size_of_subgoals st = foldr addconcl (prems_of st, 0) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 239 | end; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 240 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 241 | (*Negation Normal Form*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 242 | val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, | 
| 9869 | 243 | not_impD, not_iffD, not_allD, not_exD, not_notD]; | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 244 | fun make_nnf th = make_nnf (tryres(th, nnf_rls)) | 
| 9869 | 245 | handle THM _ => | 
| 246 | forward_res make_nnf | |
| 247 | (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward])) | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 248 | handle THM _ => th; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 249 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 250 | (*Pull existential quantifiers (Skolemization)*) | 
| 9869 | 251 | fun skolemize th = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 252 | if not (has_consts ["Ex"] (prop_of th)) then th | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 253 | else skolemize (tryres(th, [choice, conj_exD1, conj_exD2, | 
| 9869 | 254 | disj_exD, disj_exD1, disj_exD2])) | 
| 255 | handle THM _ => | |
| 256 | skolemize (forward_res skolemize | |
| 257 | (tryres (th, [conj_forward, disj_forward, all_forward]))) | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 258 | handle THM _ => forward_res skolemize (th RS ex_forward); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 259 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 260 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 261 | (*Make clauses from a list of theorems, previously Skolemized and put into nnf. | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 262 | The resulting clauses are HOL disjunctions.*) | 
| 9869 | 263 | fun make_clauses ths = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 264 | sort_clauses (map (generalize o nodups) (foldr cnf (ths,[]))); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 265 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 266 | (*Convert a list of clauses to (contrapositive) Horn clauses*) | 
| 9869 | 267 | fun make_horns ths = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 268 | name_thms "Horn#" | 
| 13105 
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
 wenzelm parents: 
12299diff
changeset | 269 | (gen_distinct Drule.eq_thm_prop (foldr (add_contras clause_rules) (ths,[]))); | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 270 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 271 | (*Could simply use nprems_of, which would count remaining subgoals -- no | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 272 | discrimination as to their size! With BEST_FIRST, fails for problem 41.*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 273 | |
| 9869 | 274 | fun best_prolog_tac sizef horns = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 275 | BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac horns 1); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 276 | |
| 9869 | 277 | fun depth_prolog_tac horns = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 278 | DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac horns 1); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 279 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 280 | (*Return all negative clauses, as possible goal clauses*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 281 | fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls)); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 282 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 283 | |
| 9869 | 284 | fun skolemize_tac prems = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 285 | cut_facts_tac (map (skolemize o make_nnf) prems) THEN' | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 286 | REPEAT o (etac exE); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 287 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 288 | (*Shell of all meson-tactics. Supplies cltac with clauses: HOL disjunctions*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 289 | fun MESON cltac = SELECT_GOAL | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 290 | (EVERY1 [rtac ccontr, | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 291 | METAHYPS (fn negs => | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 292 | EVERY1 [skolemize_tac negs, | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 293 | METAHYPS (cltac o make_clauses)])]); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 294 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 295 | (** Best-first search versions **) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 296 | |
| 9869 | 297 | fun best_meson_tac sizef = | 
| 298 | MESON (fn cls => | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 299 | THEN_BEST_FIRST (resolve_tac (gocls cls) 1) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 300 | (has_fewer_prems 1, sizef) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 301 | (prolog_step_tac (make_horns cls) 1)); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 302 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 303 | (*First, breaks the goal into independent units*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 304 | val safe_best_meson_tac = | 
| 9869 | 305 | SELECT_GOAL (TRY Safe_tac THEN | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 306 | TRYALL (best_meson_tac size_of_subgoals)); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 307 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 308 | (** Depth-first search version **) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 309 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 310 | val depth_meson_tac = | 
| 9869 | 311 | MESON (fn cls => EVERY [resolve_tac (gocls cls) 1, | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 312 | depth_prolog_tac (make_horns cls)]); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 313 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 314 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 315 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 316 | (** Iterative deepening version **) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 317 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 318 | (*This version does only one inference per call; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 319 | having only one eq_assume_tac speeds it up!*) | 
| 9869 | 320 | fun prolog_step_tac' horns = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 321 | let val (horn0s, hornps) = (*0 subgoals vs 1 or more*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 322 | take_prefix Thm.no_prems horns | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 323 | val nrtac = net_resolve_tac horns | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 324 | in fn i => eq_assume_tac i ORELSE | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 325 | match_tac horn0s i ORELSE (*no backtracking if unit MATCHES*) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 326 | ((assume_tac i APPEND nrtac i) THEN check_tac) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 327 | end; | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 328 | |
| 9869 | 329 | fun iter_deepen_prolog_tac horns = | 
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 330 | ITER_DEEPEN (has_fewer_prems 1) (prolog_step_tac' horns); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 331 | |
| 9869 | 332 | val iter_deepen_meson_tac = | 
| 333 | MESON (fn cls => | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 334 | (THEN_ITER_DEEPEN (resolve_tac (gocls cls) 1) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 335 | (has_fewer_prems 1) | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 336 | (prolog_step_tac' (make_horns cls)))); | 
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 337 | |
| 9869 | 338 | fun meson_claset_tac cs = | 
| 339 | SELECT_GOAL (TRY (safe_tac cs) THEN TRYALL iter_deepen_meson_tac); | |
| 340 | ||
| 341 | val meson_tac = CLASET' meson_claset_tac; | |
| 342 | ||
| 343 | ||
| 344 | (* proof method setup *) | |
| 345 | ||
| 346 | local | |
| 347 | ||
| 348 | fun meson_meth ctxt = | |
| 10821 | 349 | Method.SIMPLE_METHOD' HEADGOAL | 
| 350 | (CHANGED_PROP o meson_claset_tac (Classical.get_local_claset ctxt)); | |
| 9869 | 351 | |
| 352 | in | |
| 353 | ||
| 354 | val meson_setup = | |
| 355 | [Method.add_methods | |
| 356 |   [("meson", Method.ctxt_args meson_meth, "The MESON resolution proof procedure")]];
 | |
| 9840 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 357 | |
| 
9dfcb0224f8c
meson.ML moved from HOL/ex to HOL/Tools: meson_tac installed by default
 paulson parents: diff
changeset | 358 | end; | 
| 9869 | 359 | |
| 360 | end; |